Page MenuHomeIsabelle/Phabricator

No OneTemporary

This file is larger than 256 KB, so syntax highlighting was skipped.
diff --git a/src/HOL/Analysis/Abstract_Euclidean_Space.thy b/src/HOL/Analysis/Abstract_Euclidean_Space.thy
--- a/src/HOL/Analysis/Abstract_Euclidean_Space.thy
+++ b/src/HOL/Analysis/Abstract_Euclidean_Space.thy
@@ -1,405 +1,405 @@
(* Author: LCP, ported from HOL Light
*)
section\<open>Euclidean space and n-spheres, as subtopologies of n-dimensional space\<close>
theory Abstract_Euclidean_Space
imports Homotopy Locally
begin
subsection \<open>Euclidean spaces as abstract topologies\<close>
definition Euclidean_space :: "nat \<Rightarrow> (nat \<Rightarrow> real) topology"
where "Euclidean_space n \<equiv> subtopology (powertop_real UNIV) {x. \<forall>i\<ge>n. x i = 0}"
lemma topspace_Euclidean_space:
"topspace(Euclidean_space n) = {x. \<forall>i\<ge>n. x i = 0}"
by (simp add: Euclidean_space_def)
lemma nonempty_Euclidean_space: "topspace(Euclidean_space n) \<noteq> {}"
by (force simp: topspace_Euclidean_space)
lemma subset_Euclidean_space [simp]:
"topspace(Euclidean_space m) \<subseteq> topspace(Euclidean_space n) \<longleftrightarrow> m \<le> n"
apply (simp add: topspace_Euclidean_space subset_iff, safe)
apply (drule_tac x="(\<lambda>i. if i < m then 1 else 0)" in spec)
apply auto
using not_less by fastforce
lemma topspace_Euclidean_space_alt:
"topspace(Euclidean_space n) = (\<Inter>i \<in> {n..}. {x. x \<in> topspace(powertop_real UNIV) \<and> x i \<in> {0}})"
by (auto simp: topspace_Euclidean_space)
lemma closedin_Euclidean_space:
"closedin (powertop_real UNIV) (topspace(Euclidean_space n))"
proof -
have "closedin (powertop_real UNIV) {x. x i = 0}" if "n \<le> i" for i
proof -
have "closedin (powertop_real UNIV) {x \<in> topspace (powertop_real UNIV). x i \<in> {0}}"
proof (rule closedin_continuous_map_preimage)
show "continuous_map (powertop_real UNIV) euclideanreal (\<lambda>x. x i)"
by (metis UNIV_I continuous_map_product_coordinates)
show "closedin euclideanreal {0}"
by simp
qed
then show ?thesis
by auto
qed
then show ?thesis
unfolding topspace_Euclidean_space_alt
by force
qed
lemma closedin_Euclidean_imp_closed: "closedin (Euclidean_space m) S \<Longrightarrow> closed S"
by (metis Euclidean_space_def closed_closedin closedin_Euclidean_space closedin_closed_subtopology euclidean_product_topology topspace_Euclidean_space)
lemma closedin_Euclidean_space_iff:
"closedin (Euclidean_space m) S \<longleftrightarrow> closed S \<and> S \<subseteq> topspace (Euclidean_space m)"
(is "?lhs \<longleftrightarrow> ?rhs")
proof
show "?lhs \<Longrightarrow> ?rhs"
using closedin_closed_subtopology topspace_Euclidean_space
by (fastforce simp: topspace_Euclidean_space_alt closedin_Euclidean_imp_closed)
show "?rhs \<Longrightarrow> ?lhs"
apply (simp add: closedin_subtopology Euclidean_space_def)
by (metis (no_types) closed_closedin euclidean_product_topology inf.orderE)
qed
lemma continuous_map_componentwise_Euclidean_space:
"continuous_map X (Euclidean_space n) (\<lambda>x i. if i < n then f x i else 0) \<longleftrightarrow>
(\<forall>i < n. continuous_map X euclideanreal (\<lambda>x. f x i))"
proof -
have *: "continuous_map X euclideanreal (\<lambda>x. if k < n then f x k else 0)"
if "\<And>i. i<n \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x i)" for k
by (intro continuous_intros that)
show ?thesis
unfolding Euclidean_space_def continuous_map_in_subtopology
by (fastforce simp: continuous_map_componentwise_UNIV * elim: continuous_map_eq)
qed
lemma continuous_map_Euclidean_space_add [continuous_intros]:
"\<lbrakk>continuous_map X (Euclidean_space n) f; continuous_map X (Euclidean_space n) g\<rbrakk>
\<Longrightarrow> continuous_map X (Euclidean_space n) (\<lambda>x i. f x i + g x i)"
unfolding Euclidean_space_def continuous_map_in_subtopology
by (fastforce simp add: continuous_map_componentwise_UNIV continuous_map_add)
lemma continuous_map_Euclidean_space_diff [continuous_intros]:
"\<lbrakk>continuous_map X (Euclidean_space n) f; continuous_map X (Euclidean_space n) g\<rbrakk>
\<Longrightarrow> continuous_map X (Euclidean_space n) (\<lambda>x i. f x i - g x i)"
unfolding Euclidean_space_def continuous_map_in_subtopology
by (fastforce simp add: continuous_map_componentwise_UNIV continuous_map_diff)
lemma continuous_map_Euclidean_space_iff:
"continuous_map (Euclidean_space m) euclidean g
= continuous_on (topspace (Euclidean_space m)) g"
proof
assume "continuous_map (Euclidean_space m) euclidean g"
then have "continuous_map (top_of_set {f. \<forall>n\<ge>m. f n = 0}) euclidean g"
by (simp add: Euclidean_space_def euclidean_product_topology)
then show "continuous_on (topspace (Euclidean_space m)) g"
by (metis continuous_map_subtopology_eu subtopology_topspace topspace_Euclidean_space)
next
assume "continuous_on (topspace (Euclidean_space m)) g"
then have "continuous_map (top_of_set {f. \<forall>n\<ge>m. f n = 0}) euclidean g"
by (metis (lifting) continuous_map_into_fulltopology continuous_map_subtopology_eu order_refl topspace_Euclidean_space)
then show "continuous_map (Euclidean_space m) euclidean g"
by (simp add: Euclidean_space_def euclidean_product_topology)
qed
lemma cm_Euclidean_space_iff_continuous_on:
"continuous_map (subtopology (Euclidean_space m) S) (Euclidean_space n) f
\<longleftrightarrow> continuous_on (topspace (subtopology (Euclidean_space m) S)) f \<and>
f ` (topspace (subtopology (Euclidean_space m) S)) \<subseteq> topspace (Euclidean_space n)"
(is "?P \<longleftrightarrow> ?Q \<and> ?R")
proof -
have ?Q if ?P
proof -
have "\<And>n. Euclidean_space n = top_of_set {f. \<forall>m\<ge>n. f m = 0}"
by (simp add: Euclidean_space_def euclidean_product_topology)
with that show ?thesis
by (simp add: subtopology_subtopology)
qed
moreover
have ?R if ?P
using that by (simp add: image_subset_iff continuous_map_def)
moreover
have ?P if ?Q ?R
proof -
have "continuous_map (top_of_set (topspace (subtopology (subtopology (powertop_real UNIV) {f. \<forall>n\<ge>m. f n = 0}) S))) (top_of_set (topspace (subtopology (powertop_real UNIV) {f. \<forall>na\<ge>n. f na = 0}))) f"
using Euclidean_space_def that by auto
then show ?thesis
by (simp add: Euclidean_space_def euclidean_product_topology subtopology_subtopology)
qed
ultimately show ?thesis
by auto
qed
lemma homeomorphic_Euclidean_space_product_topology:
"Euclidean_space n homeomorphic_space product_topology (\<lambda>i. euclideanreal) {..<n}"
proof -
have cm: "continuous_map (product_topology (\<lambda>i. euclideanreal) {..<n})
euclideanreal (\<lambda>x. if k < n then x k else 0)" for k
by (auto intro: continuous_map_if continuous_map_product_projection)
show ?thesis
unfolding homeomorphic_space_def homeomorphic_maps_def
apply (rule_tac x="\<lambda>f. restrict f {..<n}" in exI)
apply (rule_tac x="\<lambda>f i. if i < n then f i else 0" in exI)
apply (simp add: Euclidean_space_def continuous_map_in_subtopology)
apply (intro conjI continuous_map_from_subtopology)
apply (force simp: continuous_map_componentwise cm intro: continuous_map_product_projection)+
done
qed
lemma contractible_Euclidean_space [simp]: "contractible_space (Euclidean_space n)"
using homeomorphic_Euclidean_space_product_topology contractible_space_euclideanreal
contractible_space_product_topology homeomorphic_space_contractibility by blast
lemma path_connected_Euclidean_space: "path_connected_space (Euclidean_space n)"
by (simp add: contractible_imp_path_connected_space)
lemma connected_Euclidean_space: "connected_space (Euclidean_space n)"
- by (simp add: contractible_Euclidean_space contractible_imp_connected_space)
+ by (simp add: contractible_imp_connected_space)
lemma locally_path_connected_Euclidean_space:
"locally_path_connected_space (Euclidean_space n)"
apply (simp add: homeomorphic_locally_path_connected_space [OF homeomorphic_Euclidean_space_product_topology [of n]]
locally_path_connected_space_product_topology)
using locally_path_connected_space_euclideanreal by auto
lemma compact_Euclidean_space:
"compact_space (Euclidean_space n) \<longleftrightarrow> n = 0"
by (auto simp: homeomorphic_compact_space [OF homeomorphic_Euclidean_space_product_topology] compact_space_product_topology)
subsection\<open>n-dimensional spheres\<close>
definition nsphere where
"nsphere n \<equiv> subtopology (Euclidean_space (Suc n)) { x. (\<Sum>i\<le>n. x i ^ 2) = 1 }"
lemma nsphere:
"nsphere n = subtopology (powertop_real UNIV)
{x. (\<Sum>i\<le>n. x i ^ 2) = 1 \<and> (\<forall>i>n. x i = 0)}"
by (simp add: nsphere_def Euclidean_space_def subtopology_subtopology Suc_le_eq Collect_conj_eq Int_commute)
lemma continuous_map_nsphere_projection: "continuous_map (nsphere n) euclideanreal (\<lambda>x. x k)"
unfolding nsphere
by (blast intro: continuous_map_from_subtopology [OF continuous_map_product_projection])
lemma in_topspace_nsphere: "(\<lambda>n. if n = 0 then 1 else 0) \<in> topspace (nsphere n)"
by (simp add: nsphere_def topspace_Euclidean_space power2_eq_square if_distrib [where f = "\<lambda>x. x * _"] cong: if_cong)
lemma nonempty_nsphere [simp]: "~ (topspace(nsphere n) = {})"
using in_topspace_nsphere by auto
lemma subtopology_nsphere_equator:
"subtopology (nsphere (Suc n)) {x. x(Suc n) = 0} = nsphere n"
proof -
have "({x. (\<Sum>i\<le>n. (x i)\<^sup>2) + (x (Suc n))\<^sup>2 = 1 \<and> (\<forall>i>Suc n. x i = 0)} \<inter> {x. x (Suc n) = 0})
= {x. (\<Sum>i\<le>n. (x i)\<^sup>2) = 1 \<and> (\<forall>i>n. x i = (0::real))}"
using Suc_lessI [of n] by (fastforce simp: set_eq_iff)
then show ?thesis
by (simp add: nsphere subtopology_subtopology)
qed
lemma topspace_nsphere_minus1:
assumes x: "x \<in> topspace (nsphere n)" and "x n = 0"
shows "x \<in> topspace (nsphere (n - Suc 0))"
proof (cases "n = 0")
case True
then show ?thesis
using x by auto
next
case False
have subt_eq: "nsphere (n - Suc 0) = subtopology (nsphere n) {x. x n = 0}"
by (metis False Suc_pred le_zero_eq not_le subtopology_nsphere_equator)
with x show ?thesis
by (simp add: assms)
qed
lemma continuous_map_nsphere_reflection:
"continuous_map (nsphere n) (nsphere n) (\<lambda>x i. if i = k then -x i else x i)"
proof -
have cm: "continuous_map (powertop_real UNIV) euclideanreal (\<lambda>x. if j = k then - x j else x j)" for j
proof (cases "j=k")
case True
then show ?thesis
by simp (metis UNIV_I continuous_map_product_projection)
next
case False
then show ?thesis
by (auto intro: continuous_map_product_projection)
qed
have eq: "(if i = k then x k * x k else x i * x i) = x i * x i" for i and x :: "nat \<Rightarrow> real"
by simp
show ?thesis
apply (simp add: nsphere continuous_map_in_subtopology continuous_map_componentwise_UNIV
continuous_map_from_subtopology cm)
apply (intro conjI allI impI continuous_intros continuous_map_from_subtopology continuous_map_product_projection)
apply (auto simp: power2_eq_square if_distrib [where f = "\<lambda>x. x * _"] eq cong: if_cong)
done
qed
proposition contractible_space_upper_hemisphere:
assumes "k \<le> n"
shows "contractible_space(subtopology (nsphere n) {x. x k \<ge> 0})"
proof -
define p:: "nat \<Rightarrow> real" where "p \<equiv> \<lambda>i. if i = k then 1 else 0"
have "p \<in> topspace(nsphere n)"
using assms
by (simp add: nsphere p_def power2_eq_square if_distrib [where f = "\<lambda>x. x * _"] cong: if_cong)
let ?g = "\<lambda>x i. x i / sqrt(\<Sum>j\<le>n. x j ^ 2)"
let ?h = "\<lambda>(t,q) i. (1 - t) * q i + t * p i"
let ?Y = "subtopology (Euclidean_space (Suc n)) {x. 0 \<le> x k \<and> (\<exists>i\<le>n. x i \<noteq> 0)}"
have "continuous_map (prod_topology (top_of_set {0..1}) (subtopology (nsphere n) {x. 0 \<le> x k}))
(subtopology (nsphere n) {x. 0 \<le> x k}) (?g \<circ> ?h)"
proof (rule continuous_map_compose)
have *: "\<lbrakk>0 \<le> b k; (\<Sum>i\<le>n. (b i)\<^sup>2) = 1; \<forall>i>n. b i = 0; 0 \<le> a; a \<le> 1\<rbrakk>
\<Longrightarrow> \<exists>i. (i = k \<longrightarrow> (1 - a) * b k + a \<noteq> 0) \<and>
(i \<noteq> k \<longrightarrow> i \<le> n \<and> a \<noteq> 1 \<and> b i \<noteq> 0)" for a::real and b
apply (cases "a \<noteq> 1 \<and> b k = 0"; simp)
apply (metis (no_types, lifting) atMost_iff sum.neutral zero_power2)
by (metis add.commute add_le_same_cancel2 diff_ge_0_iff_ge diff_zero less_eq_real_def mult_eq_0_iff mult_nonneg_nonneg not_le numeral_One zero_neq_numeral)
show "continuous_map (prod_topology (top_of_set {0..1}) (subtopology (nsphere n) {x. 0 \<le> x k})) ?Y ?h"
using assms
apply (auto simp: * nsphere continuous_map_componentwise_UNIV
prod_topology_subtopology subtopology_subtopology case_prod_unfold
continuous_map_in_subtopology Euclidean_space_def p_def if_distrib [where f = "\<lambda>x. _ * x"] cong: if_cong)
apply (intro continuous_map_prod_snd continuous_intros continuous_map_from_subtopology)
apply auto
done
next
have 1: "\<And>x i. \<lbrakk> i \<le> n; x i \<noteq> 0\<rbrakk> \<Longrightarrow> (\<Sum>i\<le>n. (x i / sqrt (\<Sum>j\<le>n. (x j)\<^sup>2))\<^sup>2) = 1"
by (force simp: sum_nonneg sum_nonneg_eq_0_iff field_split_simps simp flip: sum_divide_distrib)
have cm: "continuous_map ?Y (nsphere n) (\<lambda>x i. x i / sqrt (\<Sum>j\<le>n. (x j)\<^sup>2))"
unfolding Euclidean_space_def nsphere subtopology_subtopology continuous_map_in_subtopology
proof (intro continuous_intros conjI)
show "continuous_map
(subtopology (powertop_real UNIV) ({x. \<forall>i\<ge>Suc n. x i = 0} \<inter> {x. 0 \<le> x k \<and> (\<exists>i\<le>n. x i \<noteq> 0)}))
(powertop_real UNIV) (\<lambda>x i. x i / sqrt (\<Sum>j\<le>n. (x j)\<^sup>2))"
unfolding continuous_map_componentwise
by (intro continuous_intros conjI ballI) (auto simp: sum_nonneg_eq_0_iff)
qed (auto simp: 1)
show "continuous_map ?Y (subtopology (nsphere n) {x. 0 \<le> x k}) (\<lambda>x i. x i / sqrt (\<Sum>j\<le>n. (x j)\<^sup>2))"
by (force simp: cm sum_nonneg continuous_map_in_subtopology if_distrib [where f = "\<lambda>x. _ * x"] cong: if_cong)
qed
moreover have "(?g \<circ> ?h) (0, x) = x"
if "x \<in> topspace (subtopology (nsphere n) {x. 0 \<le> x k})" for x
using that
by (simp add: assms nsphere)
moreover
have "(?g \<circ> ?h) (1, x) = p"
if "x \<in> topspace (subtopology (nsphere n) {x. 0 \<le> x k})" for x
by (force simp: assms p_def power2_eq_square if_distrib [where f = "\<lambda>x. x * _"] cong: if_cong)
ultimately
show ?thesis
apply (simp add: contractible_space_def homotopic_with)
apply (rule_tac x=p in exI)
apply (rule_tac x="?g \<circ> ?h" in exI, force)
done
qed
corollary contractible_space_lower_hemisphere:
assumes "k \<le> n"
shows "contractible_space(subtopology (nsphere n) {x. x k \<le> 0})"
proof -
have "contractible_space (subtopology (nsphere n) {x. 0 \<le> x k}) = ?thesis"
proof (rule homeomorphic_space_contractibility)
show "subtopology (nsphere n) {x. 0 \<le> x k} homeomorphic_space subtopology (nsphere n) {x. x k \<le> 0}"
unfolding homeomorphic_space_def homeomorphic_maps_def
apply (rule_tac x="\<lambda>x i. if i = k then -(x i) else x i" in exI)+
apply (auto simp: continuous_map_in_subtopology continuous_map_from_subtopology
continuous_map_nsphere_reflection)
done
qed
then show ?thesis
using contractible_space_upper_hemisphere [OF assms] by metis
qed
proposition nullhomotopic_nonsurjective_sphere_map:
assumes f: "continuous_map (nsphere p) (nsphere p) f"
and fim: "f ` (topspace(nsphere p)) \<noteq> topspace(nsphere p)"
obtains a where "homotopic_with (\<lambda>x. True) (nsphere p) (nsphere p) f (\<lambda>x. a)"
proof -
obtain a where a: "a \<in> topspace(nsphere p)" "a \<notin> f ` (topspace(nsphere p))"
using fim continuous_map_image_subset_topspace f by blast
then have a1: "(\<Sum>i\<le>p. (a i)\<^sup>2) = 1" and a0: "\<And>i. i > p \<Longrightarrow> a i = 0"
by (simp_all add: nsphere)
have f1: "(\<Sum>j\<le>p. (f x j)\<^sup>2) = 1" if "x \<in> topspace (nsphere p)" for x
proof -
have "f x \<in> topspace (nsphere p)"
using continuous_map_image_subset_topspace f that by blast
then show ?thesis
by (simp add: nsphere)
qed
show thesis
proof
let ?g = "\<lambda>x i. x i / sqrt(\<Sum>j\<le>p. x j ^ 2)"
let ?h = "\<lambda>(t,x) i. (1 - t) * f x i - t * a i"
let ?Y = "subtopology (Euclidean_space(Suc p)) (- {\<lambda>i. 0})"
let ?T01 = "top_of_set {0..1::real}"
have 1: "continuous_map (prod_topology ?T01 (nsphere p)) (nsphere p) (?g \<circ> ?h)"
proof (rule continuous_map_compose)
have "continuous_map (prod_topology ?T01 (nsphere p)) euclideanreal ((\<lambda>x. f x k) \<circ> snd)" for k
unfolding nsphere
apply (simp add: continuous_map_of_snd)
apply (rule continuous_map_compose [of _ "nsphere p" f, unfolded o_def])
using f apply (simp add: nsphere)
by (simp add: continuous_map_nsphere_projection)
then have "continuous_map (prod_topology ?T01 (nsphere p)) euclideanreal (\<lambda>r. ?h r k)"
for k
unfolding case_prod_unfold o_def
by (intro continuous_map_into_fulltopology [OF continuous_map_fst] continuous_intros) auto
moreover have "?h ` ({0..1} \<times> topspace (nsphere p)) \<subseteq> {x. \<forall>i\<ge>Suc p. x i = 0}"
using continuous_map_image_subset_topspace [OF f]
by (auto simp: nsphere image_subset_iff a0)
moreover have "(\<lambda>i. 0) \<notin> ?h ` ({0..1} \<times> topspace (nsphere p))"
proof clarify
fix t b
assume eq: "(\<lambda>i. 0) = (\<lambda>i. (1 - t) * f b i - t * a i)" and "t \<in> {0..1}" and b: "b \<in> topspace (nsphere p)"
have "(1 - t)\<^sup>2 = (\<Sum>i\<le>p. ((1 - t) * f b i)^2)"
using f1 [OF b] by (simp add: power_mult_distrib flip: sum_distrib_left)
also have "\<dots> = (\<Sum>i\<le>p. (t * a i)^2)"
using eq by (simp add: fun_eq_iff)
also have "\<dots> = t\<^sup>2"
using a1 by (simp add: power_mult_distrib flip: sum_distrib_left)
finally have "1 - t = t"
by (simp add: power2_eq_iff)
then have *: "t = 1/2"
by simp
have fba: "f b \<noteq> a"
using a(2) b by blast
then show False
using eq unfolding * by (simp add: fun_eq_iff)
qed
ultimately show "continuous_map (prod_topology ?T01 (nsphere p)) ?Y ?h"
by (simp add: Euclidean_space_def continuous_map_in_subtopology continuous_map_componentwise_UNIV)
next
have *: "\<lbrakk>\<forall>i\<ge>Suc p. x i = 0; x \<noteq> (\<lambda>i. 0)\<rbrakk> \<Longrightarrow> (\<Sum>j\<le>p. (x j)\<^sup>2) \<noteq> 0" for x :: "nat \<Rightarrow> real"
by (force simp: fun_eq_iff not_less_eq_eq sum_nonneg_eq_0_iff)
show "continuous_map ?Y (nsphere p) ?g"
apply (simp add: Euclidean_space_def continuous_map_in_subtopology continuous_map_componentwise_UNIV
nsphere continuous_map_componentwise subtopology_subtopology)
apply (intro conjI allI continuous_intros continuous_map_from_subtopology [OF continuous_map_product_projection])
apply (simp_all add: *)
apply (force simp: sum_nonneg fun_eq_iff not_less_eq_eq sum_nonneg_eq_0_iff power_divide simp flip: sum_divide_distrib)
done
qed
have 2: "(?g \<circ> ?h) (0, x) = f x" if "x \<in> topspace (nsphere p)" for x
using that f1 by simp
have 3: "(?g \<circ> ?h) (1, x) = (\<lambda>i. - a i)" for x
using a by (force simp: field_split_simps nsphere)
then show "homotopic_with (\<lambda>x. True) (nsphere p) (nsphere p) f (\<lambda>x. (\<lambda>i. - a i))"
by (force simp: homotopic_with intro: 1 2 3)
qed
qed
lemma Hausdorff_Euclidean_space:
"Hausdorff_space (Euclidean_space n)"
unfolding Euclidean_space_def
by (rule Hausdorff_space_subtopology) (metis Hausdorff_space_euclidean Hausdorff_space_product_topology)
end
diff --git a/src/HOL/Analysis/Abstract_Topology.thy b/src/HOL/Analysis/Abstract_Topology.thy
--- a/src/HOL/Analysis/Abstract_Topology.thy
+++ b/src/HOL/Analysis/Abstract_Topology.thy
@@ -1,4607 +1,4601 @@
(* Author: L C Paulson, University of Cambridge [ported from HOL Light]
*)
section \<open>Operators involving abstract topology\<close>
theory Abstract_Topology
imports
Complex_Main
"HOL-Library.Set_Idioms"
"HOL-Library.FuncSet"
begin
subsection \<open>General notion of a topology as a value\<close>
definition\<^marker>\<open>tag important\<close> istopology :: "('a set \<Rightarrow> bool) \<Rightarrow> bool" where
"istopology L \<equiv> (\<forall>S T. L S \<longrightarrow> L T \<longrightarrow> L (S \<inter> T)) \<and> (\<forall>\<K>. (\<forall>K\<in>\<K>. L K) \<longrightarrow> L (\<Union>\<K>))"
typedef\<^marker>\<open>tag important\<close> 'a topology = "{L::('a set) \<Rightarrow> bool. istopology L}"
morphisms "openin" "topology"
unfolding istopology_def by blast
lemma istopology_openin[intro]: "istopology(openin U)"
using openin[of U] by blast
lemma istopology_open: "istopology open"
by (auto simp: istopology_def)
lemma topology_inverse': "istopology U \<Longrightarrow> openin (topology U) = U"
using topology_inverse[unfolded mem_Collect_eq] .
lemma topology_inverse_iff: "istopology U \<longleftrightarrow> openin (topology U) = U"
using topology_inverse[of U] istopology_openin[of "topology U"] by auto
lemma topology_eq: "T1 = T2 \<longleftrightarrow> (\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S)"
proof
assume "T1 = T2"
then show "\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S" by simp
next
assume H: "\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S"
then have "openin T1 = openin T2" by (simp add: fun_eq_iff)
then have "topology (openin T1) = topology (openin T2)" by simp
then show "T1 = T2" unfolding openin_inverse .
qed
text\<open>The "universe": the union of all sets in the topology.\<close>
definition "topspace T = \<Union>{S. openin T S}"
subsubsection \<open>Main properties of open sets\<close>
proposition openin_clauses:
fixes U :: "'a topology"
shows
"openin U {}"
"\<And>S T. openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S\<inter>T)"
"\<And>K. (\<forall>S \<in> K. openin U S) \<Longrightarrow> openin U (\<Union>K)"
using openin[of U] unfolding istopology_def by auto
lemma openin_subset[intro]: "openin U S \<Longrightarrow> S \<subseteq> topspace U"
unfolding topspace_def by blast
lemma openin_empty[simp]: "openin U {}"
by (rule openin_clauses)
lemma openin_Int[intro]: "openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S \<inter> T)"
by (rule openin_clauses)
lemma openin_Union[intro]: "(\<And>S. S \<in> K \<Longrightarrow> openin U S) \<Longrightarrow> openin U (\<Union>K)"
using openin_clauses by blast
lemma openin_Un[intro]: "openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S \<union> T)"
using openin_Union[of "{S,T}" U] by auto
lemma openin_topspace[intro, simp]: "openin U (topspace U)"
by (force simp: openin_Union topspace_def)
lemma openin_subopen: "openin U S \<longleftrightarrow> (\<forall>x \<in> S. \<exists>T. openin U T \<and> x \<in> T \<and> T \<subseteq> S)"
(is "?lhs \<longleftrightarrow> ?rhs")
proof
assume ?lhs
then show ?rhs by auto
next
assume H: ?rhs
let ?t = "\<Union>{T. openin U T \<and> T \<subseteq> S}"
have "openin U ?t" by (force simp: openin_Union)
also have "?t = S" using H by auto
finally show "openin U S" .
qed
lemma openin_INT [intro]:
assumes "finite I"
"\<And>i. i \<in> I \<Longrightarrow> openin T (U i)"
shows "openin T ((\<Inter>i \<in> I. U i) \<inter> topspace T)"
using assms by (induct, auto simp: inf_sup_aci(2) openin_Int)
lemma openin_INT2 [intro]:
assumes "finite I" "I \<noteq> {}"
"\<And>i. i \<in> I \<Longrightarrow> openin T (U i)"
shows "openin T (\<Inter>i \<in> I. U i)"
proof -
have "(\<Inter>i \<in> I. U i) \<subseteq> topspace T"
using \<open>I \<noteq> {}\<close> openin_subset[OF assms(3)] by auto
then show ?thesis
using openin_INT[of _ _ U, OF assms(1) assms(3)] by (simp add: inf.absorb2 inf_commute)
qed
lemma openin_Inter [intro]:
assumes "finite \<F>" "\<F> \<noteq> {}" "\<And>X. X \<in> \<F> \<Longrightarrow> openin T X" shows "openin T (\<Inter>\<F>)"
by (metis (full_types) assms openin_INT2 image_ident)
lemma openin_Int_Inter:
assumes "finite \<F>" "openin T U" "\<And>X. X \<in> \<F> \<Longrightarrow> openin T X" shows "openin T (U \<inter> \<Inter>\<F>)"
using openin_Inter [of "insert U \<F>"] assms by auto
subsubsection \<open>Closed sets\<close>
definition\<^marker>\<open>tag important\<close> closedin :: "'a topology \<Rightarrow> 'a set \<Rightarrow> bool" where
"closedin U S \<longleftrightarrow> S \<subseteq> topspace U \<and> openin U (topspace U - S)"
lemma closedin_subset: "closedin U S \<Longrightarrow> S \<subseteq> topspace U"
by (metis closedin_def)
lemma closedin_empty[simp]: "closedin U {}"
by (simp add: closedin_def)
lemma closedin_topspace[intro, simp]: "closedin U (topspace U)"
by (simp add: closedin_def)
lemma closedin_Un[intro]: "closedin U S \<Longrightarrow> closedin U T \<Longrightarrow> closedin U (S \<union> T)"
by (auto simp: Diff_Un closedin_def)
lemma Diff_Inter[intro]: "A - \<Inter>S = \<Union>{A - s|s. s\<in>S}"
by auto
lemma closedin_Union:
assumes "finite S" "\<And>T. T \<in> S \<Longrightarrow> closedin U T"
shows "closedin U (\<Union>S)"
using assms by induction auto
lemma closedin_Inter[intro]:
assumes Ke: "K \<noteq> {}"
and Kc: "\<And>S. S \<in>K \<Longrightarrow> closedin U S"
shows "closedin U (\<Inter>K)"
using Ke Kc unfolding closedin_def Diff_Inter by auto
lemma closedin_INT[intro]:
assumes "A \<noteq> {}" "\<And>x. x \<in> A \<Longrightarrow> closedin U (B x)"
shows "closedin U (\<Inter>x\<in>A. B x)"
apply (rule closedin_Inter)
using assms
apply auto
done
lemma closedin_Int[intro]: "closedin U S \<Longrightarrow> closedin U T \<Longrightarrow> closedin U (S \<inter> T)"
using closedin_Inter[of "{S,T}" U] by auto
lemma openin_closedin_eq: "openin U S \<longleftrightarrow> S \<subseteq> topspace U \<and> closedin U (topspace U - S)"
apply (auto simp: closedin_def Diff_Diff_Int inf_absorb2)
apply (metis openin_subset subset_eq)
done
lemma topology_finer_closedin:
"topspace X = topspace Y \<Longrightarrow> (\<forall>S. openin Y S \<longrightarrow> openin X S) \<longleftrightarrow> (\<forall>S. closedin Y S \<longrightarrow> closedin X S)"
apply safe
apply (simp add: closedin_def)
by (simp add: openin_closedin_eq)
lemma openin_closedin: "S \<subseteq> topspace U \<Longrightarrow> (openin U S \<longleftrightarrow> closedin U (topspace U - S))"
by (simp add: openin_closedin_eq)
lemma openin_diff[intro]:
assumes oS: "openin U S"
and cT: "closedin U T"
shows "openin U (S - T)"
proof -
have "S - T = S \<inter> (topspace U - T)" using openin_subset[of U S] oS cT
by (auto simp: topspace_def openin_subset)
then show ?thesis using oS cT
by (auto simp: closedin_def)
qed
lemma closedin_diff[intro]:
assumes oS: "closedin U S"
and cT: "openin U T"
shows "closedin U (S - T)"
proof -
have "S - T = S \<inter> (topspace U - T)"
using closedin_subset[of U S] oS cT by (auto simp: topspace_def)
then show ?thesis
using oS cT by (auto simp: openin_closedin_eq)
qed
subsection\<open>The discrete topology\<close>
definition discrete_topology where "discrete_topology U \<equiv> topology (\<lambda>S. S \<subseteq> U)"
lemma openin_discrete_topology [simp]: "openin (discrete_topology U) S \<longleftrightarrow> S \<subseteq> U"
proof -
have "istopology (\<lambda>S. S \<subseteq> U)"
by (auto simp: istopology_def)
then show ?thesis
by (simp add: discrete_topology_def topology_inverse')
qed
lemma topspace_discrete_topology [simp]: "topspace(discrete_topology U) = U"
by (meson openin_discrete_topology openin_subset openin_topspace order_refl subset_antisym)
lemma closedin_discrete_topology [simp]: "closedin (discrete_topology U) S \<longleftrightarrow> S \<subseteq> U"
by (simp add: closedin_def)
lemma discrete_topology_unique:
"discrete_topology U = X \<longleftrightarrow> topspace X = U \<and> (\<forall>x \<in> U. openin X {x})" (is "?lhs = ?rhs")
proof
assume R: ?rhs
then have "openin X S" if "S \<subseteq> U" for S
using openin_subopen subsetD that by fastforce
moreover have "x \<in> topspace X" if "openin X S" and "x \<in> S" for x S
using openin_subset that by blast
ultimately
show ?lhs
using R by (auto simp: topology_eq)
qed auto
lemma discrete_topology_unique_alt:
"discrete_topology U = X \<longleftrightarrow> topspace X \<subseteq> U \<and> (\<forall>x \<in> U. openin X {x})"
using openin_subset
by (auto simp: discrete_topology_unique)
lemma subtopology_eq_discrete_topology_empty:
"X = discrete_topology {} \<longleftrightarrow> topspace X = {}"
using discrete_topology_unique [of "{}" X] by auto
lemma subtopology_eq_discrete_topology_sing:
"X = discrete_topology {a} \<longleftrightarrow> topspace X = {a}"
by (metis discrete_topology_unique openin_topspace singletonD)
subsection \<open>Subspace topology\<close>
definition\<^marker>\<open>tag important\<close> subtopology :: "'a topology \<Rightarrow> 'a set \<Rightarrow> 'a topology" where
"subtopology U V = topology (\<lambda>T. \<exists>S. T = S \<inter> V \<and> openin U S)"
lemma istopology_subtopology: "istopology (\<lambda>T. \<exists>S. T = S \<inter> V \<and> openin U S)"
(is "istopology ?L")
proof -
have "?L {}" by blast
{
fix A B
assume A: "?L A" and B: "?L B"
from A B obtain Sa and Sb where Sa: "openin U Sa" "A = Sa \<inter> V" and Sb: "openin U Sb" "B = Sb \<inter> V"
by blast
have "A \<inter> B = (Sa \<inter> Sb) \<inter> V" "openin U (Sa \<inter> Sb)"
using Sa Sb by blast+
then have "?L (A \<inter> B)" by blast
}
moreover
{
fix K
assume K: "K \<subseteq> Collect ?L"
have th0: "Collect ?L = (\<lambda>S. S \<inter> V) ` Collect (openin U)"
by blast
from K[unfolded th0 subset_image_iff]
obtain Sk where Sk: "Sk \<subseteq> Collect (openin U)" "K = (\<lambda>S. S \<inter> V) ` Sk"
by blast
have "\<Union>K = (\<Union>Sk) \<inter> V"
using Sk by auto
moreover have "openin U (\<Union>Sk)"
using Sk by (auto simp: subset_eq)
ultimately have "?L (\<Union>K)" by blast
}
ultimately show ?thesis
unfolding subset_eq mem_Collect_eq istopology_def by auto
qed
lemma openin_subtopology: "openin (subtopology U V) S \<longleftrightarrow> (\<exists>T. openin U T \<and> S = T \<inter> V)"
unfolding subtopology_def topology_inverse'[OF istopology_subtopology]
by auto
lemma openin_subtopology_Int:
"openin X S \<Longrightarrow> openin (subtopology X T) (S \<inter> T)"
using openin_subtopology by auto
lemma openin_subtopology_Int2:
"openin X T \<Longrightarrow> openin (subtopology X S) (S \<inter> T)"
using openin_subtopology by auto
lemma openin_subtopology_diff_closed:
"\<lbrakk>S \<subseteq> topspace X; closedin X T\<rbrakk> \<Longrightarrow> openin (subtopology X S) (S - T)"
unfolding closedin_def openin_subtopology
by (rule_tac x="topspace X - T" in exI) auto
lemma openin_relative_to: "(openin X relative_to S) = openin (subtopology X S)"
by (force simp: relative_to_def openin_subtopology)
lemma topspace_subtopology [simp]: "topspace (subtopology U V) = topspace U \<inter> V"
by (auto simp: topspace_def openin_subtopology)
lemma topspace_subtopology_subset:
"S \<subseteq> topspace X \<Longrightarrow> topspace(subtopology X S) = S"
by (simp add: inf.absorb_iff2)
lemma closedin_subtopology: "closedin (subtopology U V) S \<longleftrightarrow> (\<exists>T. closedin U T \<and> S = T \<inter> V)"
unfolding closedin_def topspace_subtopology
by (auto simp: openin_subtopology)
lemma openin_subtopology_refl: "openin (subtopology U V) V \<longleftrightarrow> V \<subseteq> topspace U"
unfolding openin_subtopology
by auto (metis IntD1 in_mono openin_subset)
lemma subtopology_subtopology:
"subtopology (subtopology X S) T = subtopology X (S \<inter> T)"
proof -
have eq: "\<And>T'. (\<exists>S'. T' = S' \<inter> T \<and> (\<exists>T. openin X T \<and> S' = T \<inter> S)) = (\<exists>Sa. T' = Sa \<inter> (S \<inter> T) \<and> openin X Sa)"
by (metis inf_assoc)
have "subtopology (subtopology X S) T = topology (\<lambda>Ta. \<exists>Sa. Ta = Sa \<inter> T \<and> openin (subtopology X S) Sa)"
by (simp add: subtopology_def)
also have "\<dots> = subtopology X (S \<inter> T)"
by (simp add: openin_subtopology eq) (simp add: subtopology_def)
finally show ?thesis .
qed
lemma openin_subtopology_alt:
"openin (subtopology X U) S \<longleftrightarrow> S \<in> (\<lambda>T. U \<inter> T) ` Collect (openin X)"
by (simp add: image_iff inf_commute openin_subtopology)
lemma closedin_subtopology_alt:
"closedin (subtopology X U) S \<longleftrightarrow> S \<in> (\<lambda>T. U \<inter> T) ` Collect (closedin X)"
by (simp add: image_iff inf_commute closedin_subtopology)
lemma subtopology_superset:
assumes UV: "topspace U \<subseteq> V"
shows "subtopology U V = U"
proof -
{
fix S
{
fix T
assume T: "openin U T" "S = T \<inter> V"
from T openin_subset[OF T(1)] UV have eq: "S = T"
by blast
have "openin U S"
unfolding eq using T by blast
}
moreover
{
assume S: "openin U S"
then have "\<exists>T. openin U T \<and> S = T \<inter> V"
using openin_subset[OF S] UV by auto
}
ultimately have "(\<exists>T. openin U T \<and> S = T \<inter> V) \<longleftrightarrow> openin U S"
by blast
}
then show ?thesis
unfolding topology_eq openin_subtopology by blast
qed
lemma subtopology_topspace[simp]: "subtopology U (topspace U) = U"
by (simp add: subtopology_superset)
lemma subtopology_UNIV[simp]: "subtopology U UNIV = U"
by (simp add: subtopology_superset)
lemma subtopology_restrict:
"subtopology X (topspace X \<inter> S) = subtopology X S"
by (metis subtopology_subtopology subtopology_topspace)
lemma openin_subtopology_empty:
"openin (subtopology U {}) S \<longleftrightarrow> S = {}"
by (metis Int_empty_right openin_empty openin_subtopology)
lemma closedin_subtopology_empty:
"closedin (subtopology U {}) S \<longleftrightarrow> S = {}"
by (metis Int_empty_right closedin_empty closedin_subtopology)
lemma closedin_subtopology_refl [simp]:
"closedin (subtopology U X) X \<longleftrightarrow> X \<subseteq> topspace U"
by (metis closedin_def closedin_topspace inf.absorb_iff2 le_inf_iff topspace_subtopology)
lemma closedin_topspace_empty: "topspace T = {} \<Longrightarrow> (closedin T S \<longleftrightarrow> S = {})"
by (simp add: closedin_def)
lemma open_in_topspace_empty:
"topspace X = {} \<Longrightarrow> openin X S \<longleftrightarrow> S = {}"
by (simp add: openin_closedin_eq)
lemma openin_imp_subset:
"openin (subtopology U S) T \<Longrightarrow> T \<subseteq> S"
by (metis Int_iff openin_subtopology subsetI)
lemma closedin_imp_subset:
"closedin (subtopology U S) T \<Longrightarrow> T \<subseteq> S"
by (simp add: closedin_def)
lemma openin_open_subtopology:
"openin X S \<Longrightarrow> openin (subtopology X S) T \<longleftrightarrow> openin X T \<and> T \<subseteq> S"
by (metis inf.orderE openin_Int openin_imp_subset openin_subtopology)
lemma closedin_closed_subtopology:
"closedin X S \<Longrightarrow> (closedin (subtopology X S) T \<longleftrightarrow> closedin X T \<and> T \<subseteq> S)"
by (metis closedin_Int closedin_imp_subset closedin_subtopology inf.orderE)
lemma openin_subtopology_Un:
"\<lbrakk>openin (subtopology X T) S; openin (subtopology X U) S\<rbrakk>
\<Longrightarrow> openin (subtopology X (T \<union> U)) S"
by (simp add: openin_subtopology) blast
lemma closedin_subtopology_Un:
"\<lbrakk>closedin (subtopology X T) S; closedin (subtopology X U) S\<rbrakk>
\<Longrightarrow> closedin (subtopology X (T \<union> U)) S"
by (simp add: closedin_subtopology) blast
lemma openin_trans_full:
"\<lbrakk>openin (subtopology X U) S; openin X U\<rbrakk> \<Longrightarrow> openin X S"
by (simp add: openin_open_subtopology)
subsection \<open>The canonical topology from the underlying type class\<close>
abbreviation\<^marker>\<open>tag important\<close> euclidean :: "'a::topological_space topology"
where "euclidean \<equiv> topology open"
abbreviation top_of_set :: "'a::topological_space set \<Rightarrow> 'a topology"
where "top_of_set \<equiv> subtopology (topology open)"
lemma open_openin: "open S \<longleftrightarrow> openin euclidean S"
apply (rule cong[where x=S and y=S])
apply (rule topology_inverse[symmetric])
apply (auto simp: istopology_def)
done
declare open_openin [symmetric, simp]
lemma topspace_euclidean [simp]: "topspace euclidean = UNIV"
by (force simp: topspace_def)
lemma topspace_euclidean_subtopology[simp]: "topspace (top_of_set S) = S"
by (simp)
lemma closed_closedin: "closed S \<longleftrightarrow> closedin euclidean S"
by (simp add: closed_def closedin_def Compl_eq_Diff_UNIV)
declare closed_closedin [symmetric, simp]
lemma openin_subtopology_self [simp]: "openin (top_of_set S) S"
by (metis openin_topspace topspace_euclidean_subtopology)
subsubsection\<open>The most basic facts about the usual topology and metric on R\<close>
abbreviation euclideanreal :: "real topology"
where "euclideanreal \<equiv> topology open"
subsection \<open>Basic "localization" results are handy for connectedness.\<close>
lemma openin_open: "openin (top_of_set U) S \<longleftrightarrow> (\<exists>T. open T \<and> (S = U \<inter> T))"
by (auto simp: openin_subtopology)
lemma openin_Int_open:
"\<lbrakk>openin (top_of_set U) S; open T\<rbrakk>
\<Longrightarrow> openin (top_of_set U) (S \<inter> T)"
by (metis open_Int Int_assoc openin_open)
lemma openin_open_Int[intro]: "open S \<Longrightarrow> openin (top_of_set U) (U \<inter> S)"
by (auto simp: openin_open)
lemma open_openin_trans[trans]:
"open S \<Longrightarrow> open T \<Longrightarrow> T \<subseteq> S \<Longrightarrow> openin (top_of_set S) T"
by (metis Int_absorb1 openin_open_Int)
lemma open_subset: "S \<subseteq> T \<Longrightarrow> open S \<Longrightarrow> openin (top_of_set T) S"
by (auto simp: openin_open)
lemma closedin_closed: "closedin (top_of_set U) S \<longleftrightarrow> (\<exists>T. closed T \<and> S = U \<inter> T)"
by (simp add: closedin_subtopology Int_ac)
lemma closedin_closed_Int: "closed S \<Longrightarrow> closedin (top_of_set U) (U \<inter> S)"
by (metis closedin_closed)
lemma closed_subset: "S \<subseteq> T \<Longrightarrow> closed S \<Longrightarrow> closedin (top_of_set T) S"
by (auto simp: closedin_closed)
lemma closedin_closed_subset:
"\<lbrakk>closedin (top_of_set U) V; T \<subseteq> U; S = V \<inter> T\<rbrakk>
\<Longrightarrow> closedin (top_of_set T) S"
by (metis (no_types, lifting) Int_assoc Int_commute closedin_closed inf.orderE)
lemma finite_imp_closedin:
fixes S :: "'a::t1_space set"
shows "\<lbrakk>finite S; S \<subseteq> T\<rbrakk> \<Longrightarrow> closedin (top_of_set T) S"
by (simp add: finite_imp_closed closed_subset)
lemma closedin_singleton [simp]:
fixes a :: "'a::t1_space"
shows "closedin (top_of_set U) {a} \<longleftrightarrow> a \<in> U"
using closedin_subset by (force intro: closed_subset)
lemma openin_euclidean_subtopology_iff:
fixes S U :: "'a::metric_space set"
shows "openin (top_of_set U) S \<longleftrightarrow>
S \<subseteq> U \<and> (\<forall>x\<in>S. \<exists>e>0. \<forall>x'\<in>U. dist x' x < e \<longrightarrow> x'\<in> S)"
(is "?lhs \<longleftrightarrow> ?rhs")
proof
assume ?lhs
then show ?rhs
unfolding openin_open open_dist by blast
next
define T where "T = {x. \<exists>a\<in>S. \<exists>d>0. (\<forall>y\<in>U. dist y a < d \<longrightarrow> y \<in> S) \<and> dist x a < d}"
have 1: "\<forall>x\<in>T. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> T"
unfolding T_def
apply clarsimp
apply (rule_tac x="d - dist x a" in exI)
apply (clarsimp simp add: less_diff_eq)
by (metis dist_commute dist_triangle_lt)
assume ?rhs then have 2: "S = U \<inter> T"
unfolding T_def
by auto (metis dist_self)
from 1 2 show ?lhs
unfolding openin_open open_dist by fast
qed
lemma connected_openin:
"connected S \<longleftrightarrow>
\<not>(\<exists>E1 E2. openin (top_of_set S) E1 \<and>
openin (top_of_set S) E2 \<and>
S \<subseteq> E1 \<union> E2 \<and> E1 \<inter> E2 = {} \<and> E1 \<noteq> {} \<and> E2 \<noteq> {})"
apply (simp add: connected_def openin_open disjoint_iff_not_equal, safe)
by (simp_all, blast+) (* SLOW *)
lemma connected_openin_eq:
"connected S \<longleftrightarrow>
\<not>(\<exists>E1 E2. openin (top_of_set S) E1 \<and>
openin (top_of_set S) E2 \<and>
E1 \<union> E2 = S \<and> E1 \<inter> E2 = {} \<and>
E1 \<noteq> {} \<and> E2 \<noteq> {})"
apply (simp add: connected_openin, safe, blast)
by (metis Int_lower1 Un_subset_iff openin_open subset_antisym)
lemma connected_closedin:
"connected S \<longleftrightarrow>
(\<nexists>E1 E2.
closedin (top_of_set S) E1 \<and>
closedin (top_of_set S) E2 \<and>
S \<subseteq> E1 \<union> E2 \<and> E1 \<inter> E2 = {} \<and> E1 \<noteq> {} \<and> E2 \<noteq> {})"
(is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
by (auto simp add: connected_closed closedin_closed)
next
assume R: ?rhs
then show ?lhs
proof (clarsimp simp add: connected_closed closedin_closed)
fix A B
assume s_sub: "S \<subseteq> A \<union> B" "B \<inter> S \<noteq> {}"
and disj: "A \<inter> B \<inter> S = {}"
and cl: "closed A" "closed B"
have "S \<inter> (A \<union> B) = S"
using s_sub(1) by auto
have "S - A = B \<inter> S"
using Diff_subset_conv Un_Diff_Int disj s_sub(1) by auto
then have "S \<inter> A = {}"
by (metis Diff_Diff_Int Diff_disjoint Un_Diff_Int R cl closedin_closed_Int inf_commute order_refl s_sub(2))
then show "A \<inter> S = {}"
by blast
qed
qed
lemma connected_closedin_eq:
"connected S \<longleftrightarrow>
\<not>(\<exists>E1 E2.
closedin (top_of_set S) E1 \<and>
closedin (top_of_set S) E2 \<and>
E1 \<union> E2 = S \<and> E1 \<inter> E2 = {} \<and>
E1 \<noteq> {} \<and> E2 \<noteq> {})"
apply (simp add: connected_closedin, safe, blast)
by (metis Int_lower1 Un_subset_iff closedin_closed subset_antisym)
text \<open>These "transitivity" results are handy too\<close>
lemma openin_trans[trans]:
"openin (top_of_set T) S \<Longrightarrow> openin (top_of_set U) T \<Longrightarrow>
openin (top_of_set U) S"
by (metis openin_Int_open openin_open)
lemma openin_open_trans: "openin (top_of_set T) S \<Longrightarrow> open T \<Longrightarrow> open S"
by (auto simp: openin_open intro: openin_trans)
lemma closedin_trans[trans]:
"closedin (top_of_set T) S \<Longrightarrow> closedin (top_of_set U) T \<Longrightarrow>
closedin (top_of_set U) S"
by (auto simp: closedin_closed closed_Inter Int_assoc)
lemma closedin_closed_trans: "closedin (top_of_set T) S \<Longrightarrow> closed T \<Longrightarrow> closed S"
by (auto simp: closedin_closed intro: closedin_trans)
lemma openin_subtopology_Int_subset:
"\<lbrakk>openin (top_of_set u) (u \<inter> S); v \<subseteq> u\<rbrakk> \<Longrightarrow> openin (top_of_set v) (v \<inter> S)"
by (auto simp: openin_subtopology)
lemma openin_open_eq: "open s \<Longrightarrow> (openin (top_of_set s) t \<longleftrightarrow> open t \<and> t \<subseteq> s)"
using open_subset openin_open_trans openin_subset by fastforce
subsection\<open>Derived set (set of limit points)\<close>
definition derived_set_of :: "'a topology \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "derived'_set'_of" 80)
where "X derived_set_of S \<equiv>
{x \<in> topspace X.
(\<forall>T. x \<in> T \<and> openin X T \<longrightarrow> (\<exists>y\<noteq>x. y \<in> S \<and> y \<in> T))}"
lemma derived_set_of_restrict [simp]:
"X derived_set_of (topspace X \<inter> S) = X derived_set_of S"
by (simp add: derived_set_of_def) (metis openin_subset subset_iff)
lemma in_derived_set_of:
"x \<in> X derived_set_of S \<longleftrightarrow> x \<in> topspace X \<and> (\<forall>T. x \<in> T \<and> openin X T \<longrightarrow> (\<exists>y\<noteq>x. y \<in> S \<and> y \<in> T))"
by (simp add: derived_set_of_def)
lemma derived_set_of_subset_topspace:
"X derived_set_of S \<subseteq> topspace X"
by (auto simp add: derived_set_of_def)
lemma derived_set_of_subtopology:
"(subtopology X U) derived_set_of S = U \<inter> (X derived_set_of (U \<inter> S))"
by (simp add: derived_set_of_def openin_subtopology) blast
lemma derived_set_of_subset_subtopology:
"(subtopology X S) derived_set_of T \<subseteq> S"
by (simp add: derived_set_of_subtopology)
lemma derived_set_of_empty [simp]: "X derived_set_of {} = {}"
by (auto simp: derived_set_of_def)
lemma derived_set_of_mono:
"S \<subseteq> T \<Longrightarrow> X derived_set_of S \<subseteq> X derived_set_of T"
unfolding derived_set_of_def by blast
lemma derived_set_of_Un:
"X derived_set_of (S \<union> T) = X derived_set_of S \<union> X derived_set_of T" (is "?lhs = ?rhs")
proof
show "?lhs \<subseteq> ?rhs"
apply (clarsimp simp: in_derived_set_of)
by (metis IntE IntI openin_Int)
show "?rhs \<subseteq> ?lhs"
by (simp add: derived_set_of_mono)
qed
lemma derived_set_of_Union:
"finite \<F> \<Longrightarrow> X derived_set_of (\<Union>\<F>) = (\<Union>S \<in> \<F>. X derived_set_of S)"
proof (induction \<F> rule: finite_induct)
case (insert S \<F>)
then show ?case
by (simp add: derived_set_of_Un)
qed auto
lemma derived_set_of_topspace:
"X derived_set_of (topspace X) = {x \<in> topspace X. \<not> openin X {x}}"
apply (auto simp: in_derived_set_of)
by (metis Set.set_insert all_not_in_conv insertCI openin_subset subsetCE)
lemma discrete_topology_unique_derived_set:
"discrete_topology U = X \<longleftrightarrow> topspace X = U \<and> X derived_set_of U = {}"
by (auto simp: discrete_topology_unique derived_set_of_topspace)
lemma subtopology_eq_discrete_topology_eq:
"subtopology X U = discrete_topology U \<longleftrightarrow> U \<subseteq> topspace X \<and> U \<inter> X derived_set_of U = {}"
using discrete_topology_unique_derived_set [of U "subtopology X U"]
by (auto simp: eq_commute derived_set_of_subtopology)
lemma subtopology_eq_discrete_topology:
"S \<subseteq> topspace X \<and> S \<inter> X derived_set_of S = {}
\<Longrightarrow> subtopology X S = discrete_topology S"
by (simp add: subtopology_eq_discrete_topology_eq)
lemma subtopology_eq_discrete_topology_gen:
"S \<inter> X derived_set_of S = {} \<Longrightarrow> subtopology X S = discrete_topology(topspace X \<inter> S)"
by (metis Int_lower1 derived_set_of_restrict inf_assoc inf_bot_right subtopology_eq_discrete_topology_eq subtopology_subtopology subtopology_topspace)
lemma subtopology_discrete_topology [simp]: "subtopology (discrete_topology U) S = discrete_topology(U \<inter> S)"
proof -
have "(\<lambda>T. \<exists>Sa. T = Sa \<inter> S \<and> Sa \<subseteq> U) = (\<lambda>Sa. Sa \<subseteq> U \<and> Sa \<subseteq> S)"
by force
then show ?thesis
by (simp add: subtopology_def) (simp add: discrete_topology_def)
qed
lemma openin_Int_derived_set_of_subset:
"openin X S \<Longrightarrow> S \<inter> X derived_set_of T \<subseteq> X derived_set_of (S \<inter> T)"
by (auto simp: derived_set_of_def)
lemma openin_Int_derived_set_of_eq:
"openin X S \<Longrightarrow> S \<inter> X derived_set_of T = S \<inter> X derived_set_of (S \<inter> T)"
apply auto
apply (meson IntI openin_Int_derived_set_of_subset subsetCE)
by (meson derived_set_of_mono inf_sup_ord(2) subset_eq)
subsection\<open> Closure with respect to a topological space\<close>
definition closure_of :: "'a topology \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixr "closure'_of" 80)
where "X closure_of S \<equiv> {x \<in> topspace X. \<forall>T. x \<in> T \<and> openin X T \<longrightarrow> (\<exists>y \<in> S. y \<in> T)}"
lemma closure_of_restrict: "X closure_of S = X closure_of (topspace X \<inter> S)"
unfolding closure_of_def
apply safe
apply (meson IntI openin_subset subset_iff)
by auto
lemma in_closure_of:
"x \<in> X closure_of S \<longleftrightarrow>
x \<in> topspace X \<and> (\<forall>T. x \<in> T \<and> openin X T \<longrightarrow> (\<exists>y. y \<in> S \<and> y \<in> T))"
by (auto simp: closure_of_def)
lemma closure_of: "X closure_of S = topspace X \<inter> (S \<union> X derived_set_of S)"
by (fastforce simp: in_closure_of in_derived_set_of)
lemma closure_of_alt: "X closure_of S = topspace X \<inter> S \<union> X derived_set_of S"
using derived_set_of_subset_topspace [of X S]
unfolding closure_of_def in_derived_set_of
by safe (auto simp: in_derived_set_of)
lemma derived_set_of_subset_closure_of:
"X derived_set_of S \<subseteq> X closure_of S"
by (fastforce simp: closure_of_def in_derived_set_of)
lemma closure_of_subtopology:
"(subtopology X U) closure_of S = U \<inter> (X closure_of (U \<inter> S))"
unfolding closure_of_def topspace_subtopology openin_subtopology
by safe (metis (full_types) IntI Int_iff inf.commute)+
lemma closure_of_empty [simp]: "X closure_of {} = {}"
by (simp add: closure_of_alt)
lemma closure_of_topspace [simp]: "X closure_of topspace X = topspace X"
by (simp add: closure_of)
lemma closure_of_UNIV [simp]: "X closure_of UNIV = topspace X"
by (simp add: closure_of)
lemma closure_of_subset_topspace: "X closure_of S \<subseteq> topspace X"
by (simp add: closure_of)
lemma closure_of_subset_subtopology: "(subtopology X S) closure_of T \<subseteq> S"
by (simp add: closure_of_subtopology)
lemma closure_of_mono: "S \<subseteq> T \<Longrightarrow> X closure_of S \<subseteq> X closure_of T"
by (fastforce simp add: closure_of_def)
lemma closure_of_subtopology_subset:
"(subtopology X U) closure_of S \<subseteq> (X closure_of S)"
unfolding closure_of_subtopology
by clarsimp (meson closure_of_mono contra_subsetD inf.cobounded2)
lemma closure_of_subtopology_mono:
"T \<subseteq> U \<Longrightarrow> (subtopology X T) closure_of S \<subseteq> (subtopology X U) closure_of S"
unfolding closure_of_subtopology
by auto (meson closure_of_mono inf_mono subset_iff)
lemma closure_of_Un [simp]: "X closure_of (S \<union> T) = X closure_of S \<union> X closure_of T"
by (simp add: Un_assoc Un_left_commute closure_of_alt derived_set_of_Un inf_sup_distrib1)
lemma closure_of_Union:
"finite \<F> \<Longrightarrow> X closure_of (\<Union>\<F>) = (\<Union>S \<in> \<F>. X closure_of S)"
by (induction \<F> rule: finite_induct) auto
lemma closure_of_subset: "S \<subseteq> topspace X \<Longrightarrow> S \<subseteq> X closure_of S"
by (auto simp: closure_of_def)
lemma closure_of_subset_Int: "topspace X \<inter> S \<subseteq> X closure_of S"
by (auto simp: closure_of_def)
lemma closure_of_subset_eq: "S \<subseteq> topspace X \<and> X closure_of S \<subseteq> S \<longleftrightarrow> closedin X S"
proof (cases "S \<subseteq> topspace X")
case True
then have "\<forall>x. x \<in> topspace X \<and> (\<forall>T. x \<in> T \<and> openin X T \<longrightarrow> (\<exists>y\<in>S. y \<in> T)) \<longrightarrow> x \<in> S
\<Longrightarrow> openin X (topspace X - S)"
apply (subst openin_subopen, safe)
by (metis DiffI subset_eq openin_subset [of X])
then show ?thesis
by (auto simp: closedin_def closure_of_def)
next
case False
then show ?thesis
by (simp add: closedin_def)
qed
lemma closure_of_eq: "X closure_of S = S \<longleftrightarrow> closedin X S"
proof (cases "S \<subseteq> topspace X")
case True
then show ?thesis
by (metis closure_of_subset closure_of_subset_eq set_eq_subset)
next
case False
then show ?thesis
using closure_of closure_of_subset_eq by fastforce
qed
lemma closedin_contains_derived_set:
"closedin X S \<longleftrightarrow> X derived_set_of S \<subseteq> S \<and> S \<subseteq> topspace X"
proof (intro iffI conjI)
show "closedin X S \<Longrightarrow> X derived_set_of S \<subseteq> S"
using closure_of_eq derived_set_of_subset_closure_of by fastforce
show "closedin X S \<Longrightarrow> S \<subseteq> topspace X"
using closedin_subset by blast
show "X derived_set_of S \<subseteq> S \<and> S \<subseteq> topspace X \<Longrightarrow> closedin X S"
by (metis closure_of closure_of_eq inf.absorb_iff2 sup.orderE)
qed
lemma derived_set_subset_gen:
"X derived_set_of S \<subseteq> S \<longleftrightarrow> closedin X (topspace X \<inter> S)"
by (simp add: closedin_contains_derived_set derived_set_of_restrict derived_set_of_subset_topspace)
lemma derived_set_subset: "S \<subseteq> topspace X \<Longrightarrow> (X derived_set_of S \<subseteq> S \<longleftrightarrow> closedin X S)"
by (simp add: closedin_contains_derived_set)
lemma closedin_derived_set:
"closedin (subtopology X T) S \<longleftrightarrow>
S \<subseteq> topspace X \<and> S \<subseteq> T \<and> (\<forall>x. x \<in> X derived_set_of S \<and> x \<in> T \<longrightarrow> x \<in> S)"
by (auto simp: closedin_contains_derived_set derived_set_of_subtopology Int_absorb1)
lemma closedin_Int_closure_of:
"closedin (subtopology X S) T \<longleftrightarrow> S \<inter> X closure_of T = T"
by (metis Int_left_absorb closure_of_eq closure_of_subtopology)
lemma closure_of_closedin: "closedin X S \<Longrightarrow> X closure_of S = S"
by (simp add: closure_of_eq)
lemma closure_of_eq_diff: "X closure_of S = topspace X - \<Union>{T. openin X T \<and> disjnt S T}"
by (auto simp: closure_of_def disjnt_iff)
lemma closedin_closure_of [simp]: "closedin X (X closure_of S)"
unfolding closure_of_eq_diff by blast
lemma closure_of_closure_of [simp]: "X closure_of (X closure_of S) = X closure_of S"
by (simp add: closure_of_eq)
lemma closure_of_hull:
assumes "S \<subseteq> topspace X" shows "X closure_of S = (closedin X) hull S"
proof (rule hull_unique [THEN sym])
show "S \<subseteq> X closure_of S"
by (simp add: closure_of_subset assms)
next
show "closedin X (X closure_of S)"
by simp
show "\<And>T. \<lbrakk>S \<subseteq> T; closedin X T\<rbrakk> \<Longrightarrow> X closure_of S \<subseteq> T"
by (metis closure_of_eq closure_of_mono)
qed
lemma closure_of_minimal:
"\<lbrakk>S \<subseteq> T; closedin X T\<rbrakk> \<Longrightarrow> (X closure_of S) \<subseteq> T"
by (metis closure_of_eq closure_of_mono)
lemma closure_of_minimal_eq:
"\<lbrakk>S \<subseteq> topspace X; closedin X T\<rbrakk> \<Longrightarrow> (X closure_of S) \<subseteq> T \<longleftrightarrow> S \<subseteq> T"
by (meson closure_of_minimal closure_of_subset subset_trans)
lemma closure_of_unique:
"\<lbrakk>S \<subseteq> T; closedin X T;
\<And>T'. \<lbrakk>S \<subseteq> T'; closedin X T'\<rbrakk> \<Longrightarrow> T \<subseteq> T'\<rbrakk>
\<Longrightarrow> X closure_of S = T"
by (meson closedin_closure_of closedin_subset closure_of_minimal closure_of_subset eq_iff order.trans)
lemma closure_of_eq_empty_gen: "X closure_of S = {} \<longleftrightarrow> disjnt (topspace X) S"
unfolding disjnt_def closure_of_restrict [where S=S]
using closure_of by fastforce
lemma closure_of_eq_empty: "S \<subseteq> topspace X \<Longrightarrow> X closure_of S = {} \<longleftrightarrow> S = {}"
using closure_of_subset by fastforce
lemma openin_Int_closure_of_subset:
assumes "openin X S"
shows "S \<inter> X closure_of T \<subseteq> X closure_of (S \<inter> T)"
proof -
have "S \<inter> X derived_set_of T = S \<inter> X derived_set_of (S \<inter> T)"
by (meson assms openin_Int_derived_set_of_eq)
moreover have "S \<inter> (S \<inter> T) = S \<inter> T"
by fastforce
ultimately show ?thesis
by (metis closure_of_alt inf.cobounded2 inf_left_commute inf_sup_distrib1)
qed
lemma closure_of_openin_Int_closure_of:
assumes "openin X S"
shows "X closure_of (S \<inter> X closure_of T) = X closure_of (S \<inter> T)"
proof
show "X closure_of (S \<inter> X closure_of T) \<subseteq> X closure_of (S \<inter> T)"
by (simp add: assms closure_of_minimal openin_Int_closure_of_subset)
next
show "X closure_of (S \<inter> T) \<subseteq> X closure_of (S \<inter> X closure_of T)"
by (metis Int_lower1 Int_subset_iff assms closedin_closure_of closure_of_minimal_eq closure_of_mono inf_le2 le_infI1 openin_subset)
qed
lemma openin_Int_closure_of_eq:
"openin X S \<Longrightarrow> S \<inter> X closure_of T = S \<inter> X closure_of (S \<inter> T)"
apply (rule equalityI)
apply (simp add: openin_Int_closure_of_subset)
by (meson closure_of_mono inf.cobounded2 inf_mono subset_refl)
lemma openin_Int_closure_of_eq_empty:
"openin X S \<Longrightarrow> S \<inter> X closure_of T = {} \<longleftrightarrow> S \<inter> T = {}"
apply (subst openin_Int_closure_of_eq, auto)
by (meson IntI closure_of_subset_Int disjoint_iff_not_equal openin_subset subset_eq)
lemma closure_of_openin_Int_superset:
"openin X S \<and> S \<subseteq> X closure_of T
\<Longrightarrow> X closure_of (S \<inter> T) = X closure_of S"
by (metis closure_of_openin_Int_closure_of inf.orderE)
lemma closure_of_openin_subtopology_Int_closure_of:
assumes S: "openin (subtopology X U) S" and "T \<subseteq> U"
shows "X closure_of (S \<inter> X closure_of T) = X closure_of (S \<inter> T)" (is "?lhs = ?rhs")
proof
obtain S0 where S0: "openin X S0" "S = S0 \<inter> U"
using assms by (auto simp: openin_subtopology)
show "?lhs \<subseteq> ?rhs"
proof -
have "S0 \<inter> X closure_of T = S0 \<inter> X closure_of (S0 \<inter> T)"
by (meson S0(1) openin_Int_closure_of_eq)
moreover have "S0 \<inter> T = S0 \<inter> U \<inter> T"
using \<open>T \<subseteq> U\<close> by fastforce
ultimately have "S \<inter> X closure_of T \<subseteq> X closure_of (S \<inter> T)"
using S0(2) by auto
then show ?thesis
by (meson closedin_closure_of closure_of_minimal)
qed
next
show "?rhs \<subseteq> ?lhs"
proof -
have "T \<inter> S \<subseteq> T \<union> X derived_set_of T"
by force
then show ?thesis
by (metis Int_subset_iff S closure_of closure_of_mono inf.cobounded2 inf.coboundedI2 inf_commute openin_closedin_eq topspace_subtopology)
qed
qed
lemma closure_of_subtopology_open:
"openin X U \<or> S \<subseteq> U \<Longrightarrow> (subtopology X U) closure_of S = U \<inter> X closure_of S"
by (metis closure_of_subtopology inf_absorb2 openin_Int_closure_of_eq)
lemma discrete_topology_closure_of:
"(discrete_topology U) closure_of S = U \<inter> S"
by (metis closedin_discrete_topology closure_of_restrict closure_of_unique discrete_topology_unique inf_sup_ord(1) order_refl)
text\<open> Interior with respect to a topological space. \<close>
definition interior_of :: "'a topology \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixr "interior'_of" 80)
where "X interior_of S \<equiv> {x. \<exists>T. openin X T \<and> x \<in> T \<and> T \<subseteq> S}"
lemma interior_of_restrict:
"X interior_of S = X interior_of (topspace X \<inter> S)"
using openin_subset by (auto simp: interior_of_def)
lemma interior_of_eq: "(X interior_of S = S) \<longleftrightarrow> openin X S"
unfolding interior_of_def using openin_subopen by blast
lemma interior_of_openin: "openin X S \<Longrightarrow> X interior_of S = S"
by (simp add: interior_of_eq)
lemma interior_of_empty [simp]: "X interior_of {} = {}"
by (simp add: interior_of_eq)
lemma interior_of_topspace [simp]: "X interior_of (topspace X) = topspace X"
by (simp add: interior_of_eq)
lemma openin_interior_of [simp]: "openin X (X interior_of S)"
unfolding interior_of_def
using openin_subopen by fastforce
lemma interior_of_interior_of [simp]:
"X interior_of X interior_of S = X interior_of S"
by (simp add: interior_of_eq)
lemma interior_of_subset: "X interior_of S \<subseteq> S"
by (auto simp: interior_of_def)
lemma interior_of_subset_closure_of: "X interior_of S \<subseteq> X closure_of S"
by (metis closure_of_subset_Int dual_order.trans interior_of_restrict interior_of_subset)
lemma subset_interior_of_eq: "S \<subseteq> X interior_of S \<longleftrightarrow> openin X S"
by (metis interior_of_eq interior_of_subset subset_antisym)
lemma interior_of_mono: "S \<subseteq> T \<Longrightarrow> X interior_of S \<subseteq> X interior_of T"
by (auto simp: interior_of_def)
lemma interior_of_maximal: "\<lbrakk>T \<subseteq> S; openin X T\<rbrakk> \<Longrightarrow> T \<subseteq> X interior_of S"
by (auto simp: interior_of_def)
lemma interior_of_maximal_eq: "openin X T \<Longrightarrow> T \<subseteq> X interior_of S \<longleftrightarrow> T \<subseteq> S"
by (meson interior_of_maximal interior_of_subset order_trans)
lemma interior_of_unique:
"\<lbrakk>T \<subseteq> S; openin X T; \<And>T'. \<lbrakk>T' \<subseteq> S; openin X T'\<rbrakk> \<Longrightarrow> T' \<subseteq> T\<rbrakk> \<Longrightarrow> X interior_of S = T"
by (simp add: interior_of_maximal_eq interior_of_subset subset_antisym)
lemma interior_of_subset_topspace: "X interior_of S \<subseteq> topspace X"
by (simp add: openin_subset)
lemma interior_of_subset_subtopology: "(subtopology X S) interior_of T \<subseteq> S"
by (meson openin_imp_subset openin_interior_of)
lemma interior_of_Int: "X interior_of (S \<inter> T) = X interior_of S \<inter> X interior_of T"
apply (rule equalityI)
apply (simp add: interior_of_mono)
apply (auto simp: interior_of_maximal_eq openin_Int interior_of_subset le_infI1 le_infI2)
done
lemma interior_of_Inter_subset: "X interior_of (\<Inter>\<F>) \<subseteq> (\<Inter>S \<in> \<F>. X interior_of S)"
by (simp add: INT_greatest Inf_lower interior_of_mono)
lemma union_interior_of_subset:
"X interior_of S \<union> X interior_of T \<subseteq> X interior_of (S \<union> T)"
by (simp add: interior_of_mono)
lemma interior_of_eq_empty:
"X interior_of S = {} \<longleftrightarrow> (\<forall>T. openin X T \<and> T \<subseteq> S \<longrightarrow> T = {})"
by (metis bot.extremum_uniqueI interior_of_maximal interior_of_subset openin_interior_of)
lemma interior_of_eq_empty_alt:
"X interior_of S = {} \<longleftrightarrow> (\<forall>T. openin X T \<and> T \<noteq> {} \<longrightarrow> T - S \<noteq> {})"
by (auto simp: interior_of_eq_empty)
lemma interior_of_Union_openin_subsets:
"\<Union>{T. openin X T \<and> T \<subseteq> S} = X interior_of S"
by (rule interior_of_unique [symmetric]) auto
lemma interior_of_complement:
"X interior_of (topspace X - S) = topspace X - X closure_of S"
by (auto simp: interior_of_def closure_of_def)
lemma interior_of_closure_of:
"X interior_of S = topspace X - X closure_of (topspace X - S)"
unfolding interior_of_complement [symmetric]
by (metis Diff_Diff_Int interior_of_restrict)
lemma closure_of_interior_of:
"X closure_of S = topspace X - X interior_of (topspace X - S)"
by (simp add: interior_of_complement Diff_Diff_Int closure_of)
lemma closure_of_complement: "X closure_of (topspace X - S) = topspace X - X interior_of S"
unfolding interior_of_def closure_of_def
by (blast dest: openin_subset)
lemma interior_of_eq_empty_complement:
"X interior_of S = {} \<longleftrightarrow> X closure_of (topspace X - S) = topspace X"
using interior_of_subset_topspace [of X S] closure_of_complement by fastforce
lemma closure_of_eq_topspace:
"X closure_of S = topspace X \<longleftrightarrow> X interior_of (topspace X - S) = {}"
using closure_of_subset_topspace [of X S] interior_of_complement by fastforce
lemma interior_of_subtopology_subset:
"U \<inter> X interior_of S \<subseteq> (subtopology X U) interior_of S"
by (auto simp: interior_of_def openin_subtopology)
lemma interior_of_subtopology_subsets:
"T \<subseteq> U \<Longrightarrow> T \<inter> (subtopology X U) interior_of S \<subseteq> (subtopology X T) interior_of S"
by (metis inf.absorb_iff2 interior_of_subtopology_subset subtopology_subtopology)
lemma interior_of_subtopology_mono:
"\<lbrakk>S \<subseteq> T; T \<subseteq> U\<rbrakk> \<Longrightarrow> (subtopology X U) interior_of S \<subseteq> (subtopology X T) interior_of S"
by (metis dual_order.trans inf.orderE inf_commute interior_of_subset interior_of_subtopology_subsets)
lemma interior_of_subtopology_open:
assumes "openin X U"
shows "(subtopology X U) interior_of S = U \<inter> X interior_of S"
proof -
have "\<forall>A. U \<inter> X closure_of (U \<inter> A) = U \<inter> X closure_of A"
using assms openin_Int_closure_of_eq by blast
then have "topspace X \<inter> U - U \<inter> X closure_of (topspace X \<inter> U - S) = U \<inter> (topspace X - X closure_of (topspace X - S))"
by (metis (no_types) Diff_Int_distrib Int_Diff inf_commute)
then show ?thesis
unfolding interior_of_closure_of closure_of_subtopology_open topspace_subtopology
using openin_Int_closure_of_eq [OF assms]
by (metis assms closure_of_subtopology_open)
qed
lemma dense_intersects_open:
"X closure_of S = topspace X \<longleftrightarrow> (\<forall>T. openin X T \<and> T \<noteq> {} \<longrightarrow> S \<inter> T \<noteq> {})"
proof -
have "X closure_of S = topspace X \<longleftrightarrow> (topspace X - X interior_of (topspace X - S) = topspace X)"
by (simp add: closure_of_interior_of)
also have "\<dots> \<longleftrightarrow> X interior_of (topspace X - S) = {}"
by (simp add: closure_of_complement interior_of_eq_empty_complement)
also have "\<dots> \<longleftrightarrow> (\<forall>T. openin X T \<and> T \<noteq> {} \<longrightarrow> S \<inter> T \<noteq> {})"
unfolding interior_of_eq_empty_alt
using openin_subset by fastforce
finally show ?thesis .
qed
lemma interior_of_closedin_union_empty_interior_of:
assumes "closedin X S" and disj: "X interior_of T = {}"
shows "X interior_of (S \<union> T) = X interior_of S"
proof -
have "X closure_of (topspace X - T) = topspace X"
by (metis Diff_Diff_Int disj closure_of_eq_topspace closure_of_restrict interior_of_closure_of)
then show ?thesis
unfolding interior_of_closure_of
by (metis Diff_Un Diff_subset assms(1) closedin_def closure_of_openin_Int_superset)
qed
lemma interior_of_union_eq_empty:
"closedin X S
\<Longrightarrow> (X interior_of (S \<union> T) = {} \<longleftrightarrow>
X interior_of S = {} \<and> X interior_of T = {})"
by (metis interior_of_closedin_union_empty_interior_of le_sup_iff subset_empty union_interior_of_subset)
lemma discrete_topology_interior_of [simp]:
"(discrete_topology U) interior_of S = U \<inter> S"
by (simp add: interior_of_restrict [of _ S] interior_of_eq)
subsection \<open>Frontier with respect to topological space \<close>
definition frontier_of :: "'a topology \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixr "frontier'_of" 80)
where "X frontier_of S \<equiv> X closure_of S - X interior_of S"
lemma frontier_of_closures:
"X frontier_of S = X closure_of S \<inter> X closure_of (topspace X - S)"
by (metis Diff_Diff_Int closure_of_complement closure_of_subset_topspace double_diff frontier_of_def interior_of_subset_closure_of)
lemma interior_of_union_frontier_of [simp]:
"X interior_of S \<union> X frontier_of S = X closure_of S"
by (simp add: frontier_of_def interior_of_subset_closure_of subset_antisym)
lemma frontier_of_restrict: "X frontier_of S = X frontier_of (topspace X \<inter> S)"
by (metis closure_of_restrict frontier_of_def interior_of_restrict)
lemma closedin_frontier_of: "closedin X (X frontier_of S)"
by (simp add: closedin_Int frontier_of_closures)
lemma frontier_of_subset_topspace: "X frontier_of S \<subseteq> topspace X"
by (simp add: closedin_frontier_of closedin_subset)
lemma frontier_of_subset_subtopology: "(subtopology X S) frontier_of T \<subseteq> S"
by (metis (no_types) closedin_derived_set closedin_frontier_of)
lemma frontier_of_subtopology_subset:
"U \<inter> (subtopology X U) frontier_of S \<subseteq> (X frontier_of S)"
proof -
have "U \<inter> X interior_of S - subtopology X U interior_of S = {}"
by (simp add: interior_of_subtopology_subset)
moreover have "X closure_of S \<inter> subtopology X U closure_of S = subtopology X U closure_of S"
by (meson closure_of_subtopology_subset inf.absorb_iff2)
ultimately show ?thesis
unfolding frontier_of_def
by blast
qed
lemma frontier_of_subtopology_mono:
"\<lbrakk>S \<subseteq> T; T \<subseteq> U\<rbrakk> \<Longrightarrow> (subtopology X T) frontier_of S \<subseteq> (subtopology X U) frontier_of S"
by (simp add: frontier_of_def Diff_mono closure_of_subtopology_mono interior_of_subtopology_mono)
lemma clopenin_eq_frontier_of:
"closedin X S \<and> openin X S \<longleftrightarrow> S \<subseteq> topspace X \<and> X frontier_of S = {}"
proof (cases "S \<subseteq> topspace X")
case True
then show ?thesis
by (metis Diff_eq_empty_iff closure_of_eq closure_of_subset_eq frontier_of_def interior_of_eq interior_of_subset interior_of_union_frontier_of sup_bot_right)
next
case False
then show ?thesis
by (simp add: frontier_of_closures openin_closedin_eq)
qed
lemma frontier_of_eq_empty:
"S \<subseteq> topspace X \<Longrightarrow> (X frontier_of S = {} \<longleftrightarrow> closedin X S \<and> openin X S)"
by (simp add: clopenin_eq_frontier_of)
lemma frontier_of_openin:
"openin X S \<Longrightarrow> X frontier_of S = X closure_of S - S"
by (metis (no_types) frontier_of_def interior_of_eq)
lemma frontier_of_openin_straddle_Int:
assumes "openin X U" "U \<inter> X frontier_of S \<noteq> {}"
shows "U \<inter> S \<noteq> {}" "U - S \<noteq> {}"
proof -
have "U \<inter> (X closure_of S \<inter> X closure_of (topspace X - S)) \<noteq> {}"
using assms by (simp add: frontier_of_closures)
then show "U \<inter> S \<noteq> {}"
using assms openin_Int_closure_of_eq_empty by fastforce
show "U - S \<noteq> {}"
proof -
have "\<exists>A. X closure_of (A - S) \<inter> U \<noteq> {}"
using \<open>U \<inter> (X closure_of S \<inter> X closure_of (topspace X - S)) \<noteq> {}\<close> by blast
then have "\<not> U \<subseteq> S"
by (metis Diff_disjoint Diff_eq_empty_iff Int_Diff assms(1) inf_commute openin_Int_closure_of_eq_empty)
then show ?thesis
by blast
qed
qed
lemma frontier_of_subset_closedin: "closedin X S \<Longrightarrow> (X frontier_of S) \<subseteq> S"
using closure_of_eq frontier_of_def by fastforce
lemma frontier_of_empty [simp]: "X frontier_of {} = {}"
by (simp add: frontier_of_def)
lemma frontier_of_topspace [simp]: "X frontier_of topspace X = {}"
by (simp add: frontier_of_def)
lemma frontier_of_subset_eq:
assumes "S \<subseteq> topspace X"
shows "(X frontier_of S) \<subseteq> S \<longleftrightarrow> closedin X S"
proof
show "X frontier_of S \<subseteq> S \<Longrightarrow> closedin X S"
by (metis assms closure_of_subset_eq interior_of_subset interior_of_union_frontier_of le_sup_iff)
show "closedin X S \<Longrightarrow> X frontier_of S \<subseteq> S"
by (simp add: frontier_of_subset_closedin)
qed
lemma frontier_of_complement: "X frontier_of (topspace X - S) = X frontier_of S"
by (metis Diff_Diff_Int closure_of_restrict frontier_of_closures inf_commute)
lemma frontier_of_disjoint_eq:
assumes "S \<subseteq> topspace X"
shows "((X frontier_of S) \<inter> S = {} \<longleftrightarrow> openin X S)"
proof
assume "X frontier_of S \<inter> S = {}"
then have "closedin X (topspace X - S)"
using assms closure_of_subset frontier_of_def interior_of_eq interior_of_subset by fastforce
then show "openin X S"
using assms by (simp add: openin_closedin)
next
show "openin X S \<Longrightarrow> X frontier_of S \<inter> S = {}"
by (simp add: Diff_Diff_Int closedin_def frontier_of_openin inf.absorb_iff2 inf_commute)
qed
lemma frontier_of_disjoint_eq_alt:
"S \<subseteq> (topspace X - X frontier_of S) \<longleftrightarrow> openin X S"
proof (cases "S \<subseteq> topspace X")
case True
show ?thesis
using True frontier_of_disjoint_eq by auto
next
case False
then show ?thesis
by (meson Diff_subset openin_subset subset_trans)
qed
lemma frontier_of_Int:
"X frontier_of (S \<inter> T) =
X closure_of (S \<inter> T) \<inter> (X frontier_of S \<union> X frontier_of T)"
proof -
have *: "U \<subseteq> S \<and> U \<subseteq> T \<Longrightarrow> U \<inter> (S \<inter> A \<union> T \<inter> B) = U \<inter> (A \<union> B)" for U S T A B :: "'a set"
by blast
show ?thesis
by (simp add: frontier_of_closures closure_of_mono Diff_Int * flip: closure_of_Un)
qed
lemma frontier_of_Int_subset: "X frontier_of (S \<inter> T) \<subseteq> X frontier_of S \<union> X frontier_of T"
by (simp add: frontier_of_Int)
lemma frontier_of_Int_closedin:
"\<lbrakk>closedin X S; closedin X T\<rbrakk> \<Longrightarrow> X frontier_of(S \<inter> T) = X frontier_of S \<inter> T \<union> S \<inter> X frontier_of T"
apply (simp add: frontier_of_Int closedin_Int closure_of_closedin)
using frontier_of_subset_closedin by blast
lemma frontier_of_Un_subset: "X frontier_of(S \<union> T) \<subseteq> X frontier_of S \<union> X frontier_of T"
by (metis Diff_Un frontier_of_Int_subset frontier_of_complement)
lemma frontier_of_Union_subset:
"finite \<F> \<Longrightarrow> X frontier_of (\<Union>\<F>) \<subseteq> (\<Union>T \<in> \<F>. X frontier_of T)"
proof (induction \<F> rule: finite_induct)
case (insert A \<F>)
then show ?case
using frontier_of_Un_subset by fastforce
qed simp
lemma frontier_of_frontier_of_subset:
"X frontier_of (X frontier_of S) \<subseteq> X frontier_of S"
by (simp add: closedin_frontier_of frontier_of_subset_closedin)
lemma frontier_of_subtopology_open:
"openin X U \<Longrightarrow> (subtopology X U) frontier_of S = U \<inter> X frontier_of S"
by (simp add: Diff_Int_distrib closure_of_subtopology_open frontier_of_def interior_of_subtopology_open)
lemma discrete_topology_frontier_of [simp]:
"(discrete_topology U) frontier_of S = {}"
by (simp add: Diff_eq discrete_topology_closure_of frontier_of_closures)
subsection\<open>Locally finite collections\<close>
definition locally_finite_in
where
"locally_finite_in X \<A> \<longleftrightarrow>
(\<Union>\<A> \<subseteq> topspace X) \<and>
(\<forall>x \<in> topspace X. \<exists>V. openin X V \<and> x \<in> V \<and> finite {U \<in> \<A>. U \<inter> V \<noteq> {}})"
lemma finite_imp_locally_finite_in:
"\<lbrakk>finite \<A>; \<Union>\<A> \<subseteq> topspace X\<rbrakk> \<Longrightarrow> locally_finite_in X \<A>"
by (auto simp: locally_finite_in_def)
lemma locally_finite_in_subset:
assumes "locally_finite_in X \<A>" "\<B> \<subseteq> \<A>"
shows "locally_finite_in X \<B>"
proof -
have "finite {U \<in> \<A>. U \<inter> V \<noteq> {}} \<Longrightarrow> finite {U \<in> \<B>. U \<inter> V \<noteq> {}}" for V
apply (erule rev_finite_subset) using \<open>\<B> \<subseteq> \<A>\<close> by blast
then show ?thesis
using assms unfolding locally_finite_in_def by (fastforce simp add:)
qed
lemma locally_finite_in_refinement:
assumes \<A>: "locally_finite_in X \<A>" and f: "\<And>S. S \<in> \<A> \<Longrightarrow> f S \<subseteq> S"
shows "locally_finite_in X (f ` \<A>)"
proof -
show ?thesis
unfolding locally_finite_in_def
proof safe
fix x
assume "x \<in> topspace X"
then obtain V where "openin X V" "x \<in> V" "finite {U \<in> \<A>. U \<inter> V \<noteq> {}}"
using \<A> unfolding locally_finite_in_def by blast
moreover have "{U \<in> \<A>. f U \<inter> V \<noteq> {}} \<subseteq> {U \<in> \<A>. U \<inter> V \<noteq> {}}" for V
using f by blast
ultimately have "finite {U \<in> \<A>. f U \<inter> V \<noteq> {}}"
using finite_subset by blast
moreover have "f ` {U \<in> \<A>. f U \<inter> V \<noteq> {}} = {U \<in> f ` \<A>. U \<inter> V \<noteq> {}}"
by blast
ultimately have "finite {U \<in> f ` \<A>. U \<inter> V \<noteq> {}}"
by (metis (no_types, lifting) finite_imageI)
then show "\<exists>V. openin X V \<and> x \<in> V \<and> finite {U \<in> f ` \<A>. U \<inter> V \<noteq> {}}"
using \<open>openin X V\<close> \<open>x \<in> V\<close> by blast
next
show "\<And>x xa. \<lbrakk>xa \<in> \<A>; x \<in> f xa\<rbrakk> \<Longrightarrow> x \<in> topspace X"
by (meson Sup_upper \<A> f locally_finite_in_def subset_iff)
qed
qed
lemma locally_finite_in_subtopology:
assumes \<A>: "locally_finite_in X \<A>" "\<Union>\<A> \<subseteq> S"
shows "locally_finite_in (subtopology X S) \<A>"
unfolding locally_finite_in_def
proof safe
fix x
assume x: "x \<in> topspace (subtopology X S)"
then obtain V where "openin X V" "x \<in> V" and fin: "finite {U \<in> \<A>. U \<inter> V \<noteq> {}}"
using \<A> unfolding locally_finite_in_def topspace_subtopology by blast
show "\<exists>V. openin (subtopology X S) V \<and> x \<in> V \<and> finite {U \<in> \<A>. U \<inter> V \<noteq> {}}"
proof (intro exI conjI)
show "openin (subtopology X S) (S \<inter> V)"
by (simp add: \<open>openin X V\<close> openin_subtopology_Int2)
have "{U \<in> \<A>. U \<inter> (S \<inter> V) \<noteq> {}} \<subseteq> {U \<in> \<A>. U \<inter> V \<noteq> {}}"
by auto
with fin show "finite {U \<in> \<A>. U \<inter> (S \<inter> V) \<noteq> {}}"
using finite_subset by auto
show "x \<in> S \<inter> V"
using x \<open>x \<in> V\<close> by (simp)
qed
next
show "\<And>x A. \<lbrakk>x \<in> A; A \<in> \<A>\<rbrakk> \<Longrightarrow> x \<in> topspace (subtopology X S)"
using assms unfolding locally_finite_in_def topspace_subtopology by blast
qed
lemma closedin_locally_finite_Union:
assumes clo: "\<And>S. S \<in> \<A> \<Longrightarrow> closedin X S" and \<A>: "locally_finite_in X \<A>"
shows "closedin X (\<Union>\<A>)"
using \<A> unfolding locally_finite_in_def closedin_def
proof clarify
show "openin X (topspace X - \<Union>\<A>)"
proof (subst openin_subopen, clarify)
fix x
assume "x \<in> topspace X" and "x \<notin> \<Union>\<A>"
then obtain V where "openin X V" "x \<in> V" and fin: "finite {U \<in> \<A>. U \<inter> V \<noteq> {}}"
using \<A> unfolding locally_finite_in_def by blast
let ?T = "V - \<Union>{S \<in> \<A>. S \<inter> V \<noteq> {}}"
show "\<exists>T. openin X T \<and> x \<in> T \<and> T \<subseteq> topspace X - \<Union>\<A>"
proof (intro exI conjI)
show "openin X ?T"
by (metis (no_types, lifting) fin \<open>openin X V\<close> clo closedin_Union mem_Collect_eq openin_diff)
show "x \<in> ?T"
using \<open>x \<notin> \<Union>\<A>\<close> \<open>x \<in> V\<close> by auto
show "?T \<subseteq> topspace X - \<Union>\<A>"
using \<open>openin X V\<close> openin_subset by auto
qed
qed
qed
lemma locally_finite_in_closure:
assumes \<A>: "locally_finite_in X \<A>"
shows "locally_finite_in X ((\<lambda>S. X closure_of S) ` \<A>)"
using \<A> unfolding locally_finite_in_def
proof (intro conjI; clarsimp)
fix x A
assume "x \<in> X closure_of A"
then show "x \<in> topspace X"
by (meson in_closure_of)
next
fix x
assume "x \<in> topspace X" and "\<Union>\<A> \<subseteq> topspace X"
then obtain V where V: "openin X V" "x \<in> V" and fin: "finite {U \<in> \<A>. U \<inter> V \<noteq> {}}"
using \<A> unfolding locally_finite_in_def by blast
have eq: "{y \<in> f ` \<A>. Q y} = f ` {x. x \<in> \<A> \<and> Q(f x)}" for f Q
by blast
have eq2: "{A \<in> \<A>. X closure_of A \<inter> V \<noteq> {}} = {A \<in> \<A>. A \<inter> V \<noteq> {}}"
using openin_Int_closure_of_eq_empty V by blast
have "finite {U \<in> (closure_of) X ` \<A>. U \<inter> V \<noteq> {}}"
by (simp add: eq eq2 fin)
with V show "\<exists>V. openin X V \<and> x \<in> V \<and> finite {U \<in> (closure_of) X ` \<A>. U \<inter> V \<noteq> {}}"
by blast
qed
lemma closedin_Union_locally_finite_closure:
"locally_finite_in X \<A> \<Longrightarrow> closedin X (\<Union>((\<lambda>S. X closure_of S) ` \<A>))"
by (metis (mono_tags) closedin_closure_of closedin_locally_finite_Union imageE locally_finite_in_closure)
lemma closure_of_Union_subset: "\<Union>((\<lambda>S. X closure_of S) ` \<A>) \<subseteq> X closure_of (\<Union>\<A>)"
by clarify (meson Union_upper closure_of_mono subsetD)
lemma closure_of_locally_finite_Union:
"locally_finite_in X \<A> \<Longrightarrow> X closure_of (\<Union>\<A>) = \<Union>((\<lambda>S. X closure_of S) ` \<A>)"
apply (rule closure_of_unique)
apply (simp add: SUP_upper2 Sup_le_iff closure_of_subset locally_finite_in_def)
apply (simp add: closedin_Union_locally_finite_closure)
by (simp add: Sup_le_iff closure_of_minimal)
subsection\<^marker>\<open>tag important\<close> \<open>Continuous maps\<close>
text \<open>We will need to deal with continuous maps in terms of topologies and not in terms
of type classes, as defined below.\<close>
definition continuous_map where
"continuous_map X Y f \<equiv>
(\<forall>x \<in> topspace X. f x \<in> topspace Y) \<and>
(\<forall>U. openin Y U \<longrightarrow> openin X {x \<in> topspace X. f x \<in> U})"
lemma continuous_map:
"continuous_map X Y f \<longleftrightarrow>
f ` (topspace X) \<subseteq> topspace Y \<and> (\<forall>U. openin Y U \<longrightarrow> openin X {x \<in> topspace X. f x \<in> U})"
by (auto simp: continuous_map_def)
lemma continuous_map_image_subset_topspace:
"continuous_map X Y f \<Longrightarrow> f ` (topspace X) \<subseteq> topspace Y"
by (auto simp: continuous_map_def)
lemma continuous_map_on_empty: "topspace X = {} \<Longrightarrow> continuous_map X Y f"
by (auto simp: continuous_map_def)
lemma continuous_map_closedin:
"continuous_map X Y f \<longleftrightarrow>
(\<forall>x \<in> topspace X. f x \<in> topspace Y) \<and>
(\<forall>C. closedin Y C \<longrightarrow> closedin X {x \<in> topspace X. f x \<in> C})"
proof -
have "(\<forall>U. openin Y U \<longrightarrow> openin X {x \<in> topspace X. f x \<in> U}) =
(\<forall>C. closedin Y C \<longrightarrow> closedin X {x \<in> topspace X. f x \<in> C})"
if "\<And>x. x \<in> topspace X \<Longrightarrow> f x \<in> topspace Y"
proof -
have eq: "{x \<in> topspace X. f x \<in> topspace Y \<and> f x \<notin> C} = (topspace X - {x \<in> topspace X. f x \<in> C})" for C
using that by blast
show ?thesis
proof (intro iffI allI impI)
fix C
assume "\<forall>U. openin Y U \<longrightarrow> openin X {x \<in> topspace X. f x \<in> U}" and "closedin Y C"
then have "openin X {x \<in> topspace X. f x \<in> topspace Y - C}" by blast
then show "closedin X {x \<in> topspace X. f x \<in> C}"
by (auto simp add: closedin_def eq)
next
fix U
assume "\<forall>C. closedin Y C \<longrightarrow> closedin X {x \<in> topspace X. f x \<in> C}" and "openin Y U"
then have "closedin X {x \<in> topspace X. f x \<in> topspace Y - U}" by blast
then show "openin X {x \<in> topspace X. f x \<in> U}"
by (auto simp add: openin_closedin_eq eq)
qed
qed
then show ?thesis
by (auto simp: continuous_map_def)
qed
lemma openin_continuous_map_preimage:
"\<lbrakk>continuous_map X Y f; openin Y U\<rbrakk> \<Longrightarrow> openin X {x \<in> topspace X. f x \<in> U}"
by (simp add: continuous_map_def)
lemma closedin_continuous_map_preimage:
"\<lbrakk>continuous_map X Y f; closedin Y C\<rbrakk> \<Longrightarrow> closedin X {x \<in> topspace X. f x \<in> C}"
by (simp add: continuous_map_closedin)
lemma openin_continuous_map_preimage_gen:
assumes "continuous_map X Y f" "openin X U" "openin Y V"
shows "openin X {x \<in> U. f x \<in> V}"
proof -
have eq: "{x \<in> U. f x \<in> V} = U \<inter> {x \<in> topspace X. f x \<in> V}"
using assms(2) openin_closedin_eq by fastforce
show ?thesis
unfolding eq
using assms openin_continuous_map_preimage by fastforce
qed
lemma closedin_continuous_map_preimage_gen:
assumes "continuous_map X Y f" "closedin X U" "closedin Y V"
shows "closedin X {x \<in> U. f x \<in> V}"
proof -
have eq: "{x \<in> U. f x \<in> V} = U \<inter> {x \<in> topspace X. f x \<in> V}"
using assms(2) closedin_def by fastforce
show ?thesis
unfolding eq
using assms closedin_continuous_map_preimage by fastforce
qed
lemma continuous_map_image_closure_subset:
assumes "continuous_map X Y f"
shows "f ` (X closure_of S) \<subseteq> Y closure_of f ` S"
proof -
have *: "f ` (topspace X) \<subseteq> topspace Y"
by (meson assms continuous_map)
have "X closure_of T \<subseteq> {x \<in> X closure_of T. f x \<in> Y closure_of (f ` T)}" if "T \<subseteq> topspace X" for T
proof (rule closure_of_minimal)
show "T \<subseteq> {x \<in> X closure_of T. f x \<in> Y closure_of f ` T}"
using closure_of_subset * that by (fastforce simp: in_closure_of)
next
show "closedin X {x \<in> X closure_of T. f x \<in> Y closure_of f ` T}"
using assms closedin_continuous_map_preimage_gen by fastforce
qed
then have "f ` (X closure_of (topspace X \<inter> S)) \<subseteq> Y closure_of (f ` (topspace X \<inter> S))"
by blast
also have "\<dots> \<subseteq> Y closure_of (topspace Y \<inter> f ` S)"
using * by (blast intro!: closure_of_mono)
finally have "f ` (X closure_of (topspace X \<inter> S)) \<subseteq> Y closure_of (topspace Y \<inter> f ` S)" .
then show ?thesis
by (metis closure_of_restrict)
qed
lemma continuous_map_subset_aux1: "continuous_map X Y f \<Longrightarrow>
(\<forall>S. f ` (X closure_of S) \<subseteq> Y closure_of f ` S)"
using continuous_map_image_closure_subset by blast
lemma continuous_map_subset_aux2:
assumes "\<forall>S. S \<subseteq> topspace X \<longrightarrow> f ` (X closure_of S) \<subseteq> Y closure_of f ` S"
shows "continuous_map X Y f"
unfolding continuous_map_closedin
proof (intro conjI ballI allI impI)
fix x
assume "x \<in> topspace X"
then show "f x \<in> topspace Y"
using assms closure_of_subset_topspace by fastforce
next
fix C
assume "closedin Y C"
then show "closedin X {x \<in> topspace X. f x \<in> C}"
proof (clarsimp simp flip: closure_of_subset_eq, intro conjI)
fix x
assume x: "x \<in> X closure_of {x \<in> topspace X. f x \<in> C}"
and "C \<subseteq> topspace Y" and "Y closure_of C \<subseteq> C"
show "x \<in> topspace X"
by (meson x in_closure_of)
have "{a \<in> topspace X. f a \<in> C} \<subseteq> topspace X"
by simp
moreover have "Y closure_of f ` {a \<in> topspace X. f a \<in> C} \<subseteq> C"
by (simp add: \<open>closedin Y C\<close> closure_of_minimal image_subset_iff)
ultimately have "f ` (X closure_of {a \<in> topspace X. f a \<in> C}) \<subseteq> C"
using assms by blast
then show "f x \<in> C"
using x by auto
qed
qed
lemma continuous_map_eq_image_closure_subset:
"continuous_map X Y f \<longleftrightarrow> (\<forall>S. f ` (X closure_of S) \<subseteq> Y closure_of f ` S)"
using continuous_map_subset_aux1 continuous_map_subset_aux2 by metis
lemma continuous_map_eq_image_closure_subset_alt:
"continuous_map X Y f \<longleftrightarrow> (\<forall>S. S \<subseteq> topspace X \<longrightarrow> f ` (X closure_of S) \<subseteq> Y closure_of f ` S)"
using continuous_map_subset_aux1 continuous_map_subset_aux2 by metis
lemma continuous_map_eq_image_closure_subset_gen:
"continuous_map X Y f \<longleftrightarrow>
f ` (topspace X) \<subseteq> topspace Y \<and>
(\<forall>S. f ` (X closure_of S) \<subseteq> Y closure_of f ` S)"
using continuous_map_subset_aux1 continuous_map_subset_aux2 continuous_map_image_subset_topspace by metis
lemma continuous_map_closure_preimage_subset:
"continuous_map X Y f
\<Longrightarrow> X closure_of {x \<in> topspace X. f x \<in> T}
\<subseteq> {x \<in> topspace X. f x \<in> Y closure_of T}"
unfolding continuous_map_closedin
by (rule closure_of_minimal) (use in_closure_of in \<open>fastforce+\<close>)
lemma continuous_map_frontier_frontier_preimage_subset:
assumes "continuous_map X Y f"
shows "X frontier_of {x \<in> topspace X. f x \<in> T} \<subseteq> {x \<in> topspace X. f x \<in> Y frontier_of T}"
proof -
have eq: "topspace X - {x \<in> topspace X. f x \<in> T} = {x \<in> topspace X. f x \<in> topspace Y - T}"
using assms unfolding continuous_map_def by blast
have "X closure_of {x \<in> topspace X. f x \<in> T} \<subseteq> {x \<in> topspace X. f x \<in> Y closure_of T}"
by (simp add: assms continuous_map_closure_preimage_subset)
moreover
have "X closure_of (topspace X - {x \<in> topspace X. f x \<in> T}) \<subseteq> {x \<in> topspace X. f x \<in> Y closure_of (topspace Y - T)}"
using continuous_map_closure_preimage_subset [OF assms] eq by presburger
ultimately show ?thesis
by (auto simp: frontier_of_closures)
qed
lemma topology_finer_continuous_id:
"topspace X = topspace Y \<Longrightarrow> ((\<forall>S. openin X S \<longrightarrow> openin Y S) \<longleftrightarrow> continuous_map Y X id)"
unfolding continuous_map_def
apply auto
using openin_subopen openin_subset apply fastforce
using openin_subopen topspace_def by fastforce
lemma continuous_map_const [simp]:
"continuous_map X Y (\<lambda>x. C) \<longleftrightarrow> topspace X = {} \<or> C \<in> topspace Y"
proof (cases "topspace X = {}")
case False
show ?thesis
proof (cases "C \<in> topspace Y")
case True
with openin_subopen show ?thesis
by (auto simp: continuous_map_def)
next
case False
then show ?thesis
unfolding continuous_map_def by fastforce
qed
qed (auto simp: continuous_map_on_empty)
declare continuous_map_const [THEN iffD2, continuous_intros]
lemma continuous_map_compose [continuous_intros]:
assumes f: "continuous_map X X' f" and g: "continuous_map X' X'' g"
shows "continuous_map X X'' (g \<circ> f)"
unfolding continuous_map_def
proof (intro conjI ballI allI impI)
fix x
assume "x \<in> topspace X"
then show "(g \<circ> f) x \<in> topspace X''"
using assms unfolding continuous_map_def by force
next
fix U
assume "openin X'' U"
have eq: "{x \<in> topspace X. (g \<circ> f) x \<in> U} = {x \<in> topspace X. f x \<in> {y. y \<in> topspace X' \<and> g y \<in> U}}"
by auto (meson f continuous_map_def)
show "openin X {x \<in> topspace X. (g \<circ> f) x \<in> U}"
unfolding eq
using assms unfolding continuous_map_def
using \<open>openin X'' U\<close> by blast
qed
lemma continuous_map_eq:
assumes "continuous_map X X' f" and "\<And>x. x \<in> topspace X \<Longrightarrow> f x = g x" shows "continuous_map X X' g"
proof -
have eq: "{x \<in> topspace X. f x \<in> U} = {x \<in> topspace X. g x \<in> U}" for U
using assms by auto
show ?thesis
using assms by (simp add: continuous_map_def eq)
qed
lemma restrict_continuous_map [simp]:
"topspace X \<subseteq> S \<Longrightarrow> continuous_map X X' (restrict f S) \<longleftrightarrow> continuous_map X X' f"
by (auto simp: elim!: continuous_map_eq)
lemma continuous_map_in_subtopology:
"continuous_map X (subtopology X' S) f \<longleftrightarrow> continuous_map X X' f \<and> f ` (topspace X) \<subseteq> S"
(is "?lhs = ?rhs")
proof
assume L: ?lhs
show ?rhs
proof -
have "\<And>A. f ` (X closure_of A) \<subseteq> subtopology X' S closure_of f ` A"
by (meson L continuous_map_image_closure_subset)
then show ?thesis
by (metis (no_types) closure_of_subset_subtopology closure_of_subtopology_subset closure_of_topspace continuous_map_eq_image_closure_subset dual_order.trans)
qed
next
assume R: ?rhs
then have eq: "{x \<in> topspace X. f x \<in> U} = {x \<in> topspace X. f x \<in> U \<and> f x \<in> S}" for U
by auto
show ?lhs
using R
unfolding continuous_map
by (auto simp: openin_subtopology eq)
qed
lemma continuous_map_from_subtopology:
"continuous_map X X' f \<Longrightarrow> continuous_map (subtopology X S) X' f"
by (auto simp: continuous_map openin_subtopology)
lemma continuous_map_into_fulltopology:
"continuous_map X (subtopology X' T) f \<Longrightarrow> continuous_map X X' f"
by (auto simp: continuous_map_in_subtopology)
lemma continuous_map_into_subtopology:
"\<lbrakk>continuous_map X X' f; f ` topspace X \<subseteq> T\<rbrakk> \<Longrightarrow> continuous_map X (subtopology X' T) f"
by (auto simp: continuous_map_in_subtopology)
lemma continuous_map_from_subtopology_mono:
"\<lbrakk>continuous_map (subtopology X T) X' f; S \<subseteq> T\<rbrakk>
\<Longrightarrow> continuous_map (subtopology X S) X' f"
by (metis inf.absorb_iff2 continuous_map_from_subtopology subtopology_subtopology)
lemma continuous_map_from_discrete_topology [simp]:
"continuous_map (discrete_topology U) X f \<longleftrightarrow> f ` U \<subseteq> topspace X"
by (auto simp: continuous_map_def)
lemma continuous_map_iff_continuous [simp]: "continuous_map (top_of_set S) euclidean g = continuous_on S g"
by (fastforce simp add: continuous_map openin_subtopology continuous_on_open_invariant)
lemma continuous_map_iff_continuous2 [simp]: "continuous_map euclidean euclidean g = continuous_on UNIV g"
by (metis continuous_map_iff_continuous subtopology_UNIV)
lemma continuous_map_openin_preimage_eq:
"continuous_map X Y f \<longleftrightarrow>
f ` (topspace X) \<subseteq> topspace Y \<and> (\<forall>U. openin Y U \<longrightarrow> openin X (topspace X \<inter> f -` U))"
by (auto simp: continuous_map_def vimage_def Int_def)
lemma continuous_map_closedin_preimage_eq:
"continuous_map X Y f \<longleftrightarrow>
f ` (topspace X) \<subseteq> topspace Y \<and> (\<forall>U. closedin Y U \<longrightarrow> closedin X (topspace X \<inter> f -` U))"
by (auto simp: continuous_map_closedin vimage_def Int_def)
lemma continuous_map_square_root: "continuous_map euclideanreal euclideanreal sqrt"
by (simp add: continuous_at_imp_continuous_on isCont_real_sqrt)
lemma continuous_map_sqrt [continuous_intros]:
"continuous_map X euclideanreal f \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. sqrt(f x))"
by (meson continuous_map_compose continuous_map_eq continuous_map_square_root o_apply)
lemma continuous_map_id [simp, continuous_intros]: "continuous_map X X id"
unfolding continuous_map_def using openin_subopen topspace_def by fastforce
declare continuous_map_id [unfolded id_def, simp, continuous_intros]
lemma continuous_map_id_subt [simp]: "continuous_map (subtopology X S) X id"
by (simp add: continuous_map_from_subtopology)
declare continuous_map_id_subt [unfolded id_def, simp]
lemma\<^marker>\<open>tag important\<close> continuous_map_alt:
"continuous_map T1 T2 f
= ((\<forall>U. openin T2 U \<longrightarrow> openin T1 (f -` U \<inter> topspace T1)) \<and> f ` topspace T1 \<subseteq> topspace T2)"
by (auto simp: continuous_map_def vimage_def image_def Collect_conj_eq inf_commute)
lemma continuous_map_open [intro]:
"continuous_map T1 T2 f \<Longrightarrow> openin T2 U \<Longrightarrow> openin T1 (f-`U \<inter> topspace(T1))"
unfolding continuous_map_alt by auto
lemma continuous_map_preimage_topspace [intro]:
assumes "continuous_map T1 T2 f"
shows "f-`(topspace T2) \<inter> topspace T1 = topspace T1"
using assms unfolding continuous_map_def by auto
subsection\<open>Open and closed maps (not a priori assumed continuous)\<close>
definition open_map :: "'a topology \<Rightarrow> 'b topology \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
where "open_map X1 X2 f \<equiv> \<forall>U. openin X1 U \<longrightarrow> openin X2 (f ` U)"
definition closed_map :: "'a topology \<Rightarrow> 'b topology \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
where "closed_map X1 X2 f \<equiv> \<forall>U. closedin X1 U \<longrightarrow> closedin X2 (f ` U)"
lemma open_map_imp_subset_topspace:
"open_map X1 X2 f \<Longrightarrow> f ` (topspace X1) \<subseteq> topspace X2"
unfolding open_map_def by (simp add: openin_subset)
lemma open_map_on_empty:
"topspace X = {} \<Longrightarrow> open_map X Y f"
by (metis empty_iff imageE in_mono open_map_def openin_subopen openin_subset)
lemma closed_map_on_empty:
"topspace X = {} \<Longrightarrow> closed_map X Y f"
by (simp add: closed_map_def closedin_topspace_empty)
lemma closed_map_const:
"closed_map X Y (\<lambda>x. c) \<longleftrightarrow> topspace X = {} \<or> closedin Y {c}"
proof (cases "topspace X = {}")
case True
then show ?thesis
by (simp add: closed_map_on_empty)
next
case False
then show ?thesis
by (auto simp: closed_map_def image_constant_conv)
qed
lemma open_map_imp_subset:
"\<lbrakk>open_map X1 X2 f; S \<subseteq> topspace X1\<rbrakk> \<Longrightarrow> f ` S \<subseteq> topspace X2"
by (meson order_trans open_map_imp_subset_topspace subset_image_iff)
lemma topology_finer_open_id:
"(\<forall>S. openin X S \<longrightarrow> openin X' S) \<longleftrightarrow> open_map X X' id"
unfolding open_map_def by auto
lemma open_map_id: "open_map X X id"
unfolding open_map_def by auto
lemma open_map_eq:
"\<lbrakk>open_map X X' f; \<And>x. x \<in> topspace X \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> open_map X X' g"
unfolding open_map_def
by (metis image_cong openin_subset subset_iff)
lemma open_map_inclusion_eq:
"open_map (subtopology X S) X id \<longleftrightarrow> openin X (topspace X \<inter> S)"
proof -
have *: "openin X (T \<inter> S)" if "openin X (S \<inter> topspace X)" "openin X T" for T
proof -
have "T \<subseteq> topspace X"
using that by (simp add: openin_subset)
with that show "openin X (T \<inter> S)"
by (metis inf.absorb1 inf.left_commute inf_commute openin_Int)
qed
show ?thesis
by (fastforce simp add: open_map_def Int_commute openin_subtopology_alt intro: *)
qed
lemma open_map_inclusion:
"openin X S \<Longrightarrow> open_map (subtopology X S) X id"
by (simp add: open_map_inclusion_eq openin_Int)
lemma open_map_compose:
"\<lbrakk>open_map X X' f; open_map X' X'' g\<rbrakk> \<Longrightarrow> open_map X X'' (g \<circ> f)"
by (metis (no_types, lifting) image_comp open_map_def)
lemma closed_map_imp_subset_topspace:
"closed_map X1 X2 f \<Longrightarrow> f ` (topspace X1) \<subseteq> topspace X2"
by (simp add: closed_map_def closedin_subset)
lemma closed_map_imp_subset:
"\<lbrakk>closed_map X1 X2 f; S \<subseteq> topspace X1\<rbrakk> \<Longrightarrow> f ` S \<subseteq> topspace X2"
using closed_map_imp_subset_topspace by blast
lemma topology_finer_closed_id:
"(\<forall>S. closedin X S \<longrightarrow> closedin X' S) \<longleftrightarrow> closed_map X X' id"
by (simp add: closed_map_def)
lemma closed_map_id: "closed_map X X id"
by (simp add: closed_map_def)
lemma closed_map_eq:
"\<lbrakk>closed_map X X' f; \<And>x. x \<in> topspace X \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> closed_map X X' g"
unfolding closed_map_def
by (metis image_cong closedin_subset subset_iff)
lemma closed_map_compose:
"\<lbrakk>closed_map X X' f; closed_map X' X'' g\<rbrakk> \<Longrightarrow> closed_map X X'' (g \<circ> f)"
by (metis (no_types, lifting) closed_map_def image_comp)
lemma closed_map_inclusion_eq:
"closed_map (subtopology X S) X id \<longleftrightarrow>
closedin X (topspace X \<inter> S)"
proof -
have *: "closedin X (T \<inter> S)" if "closedin X (S \<inter> topspace X)" "closedin X T" for T
proof -
have "T \<subseteq> topspace X"
using that by (simp add: closedin_subset)
with that show "closedin X (T \<inter> S)"
by (metis inf.absorb1 inf.left_commute inf_commute closedin_Int)
qed
show ?thesis
by (fastforce simp add: closed_map_def Int_commute closedin_subtopology_alt intro: *)
qed
lemma closed_map_inclusion: "closedin X S \<Longrightarrow> closed_map (subtopology X S) X id"
by (simp add: closed_map_inclusion_eq closedin_Int)
lemma open_map_into_subtopology:
"\<lbrakk>open_map X X' f; f ` topspace X \<subseteq> S\<rbrakk> \<Longrightarrow> open_map X (subtopology X' S) f"
unfolding open_map_def openin_subtopology
using openin_subset by fastforce
lemma closed_map_into_subtopology:
"\<lbrakk>closed_map X X' f; f ` topspace X \<subseteq> S\<rbrakk> \<Longrightarrow> closed_map X (subtopology X' S) f"
unfolding closed_map_def closedin_subtopology
using closedin_subset by fastforce
lemma open_map_into_discrete_topology:
"open_map X (discrete_topology U) f \<longleftrightarrow> f ` (topspace X) \<subseteq> U"
unfolding open_map_def openin_discrete_topology using openin_subset by blast
lemma closed_map_into_discrete_topology:
"closed_map X (discrete_topology U) f \<longleftrightarrow> f ` (topspace X) \<subseteq> U"
unfolding closed_map_def closedin_discrete_topology using closedin_subset by blast
lemma bijective_open_imp_closed_map:
"\<lbrakk>open_map X X' f; f ` (topspace X) = topspace X'; inj_on f (topspace X)\<rbrakk> \<Longrightarrow> closed_map X X' f"
unfolding open_map_def closed_map_def closedin_def
by auto (metis Diff_subset inj_on_image_set_diff)
lemma bijective_closed_imp_open_map:
"\<lbrakk>closed_map X X' f; f ` (topspace X) = topspace X'; inj_on f (topspace X)\<rbrakk> \<Longrightarrow> open_map X X' f"
unfolding closed_map_def open_map_def openin_closedin_eq
by auto (metis Diff_subset inj_on_image_set_diff)
lemma open_map_from_subtopology:
"\<lbrakk>open_map X X' f; openin X U\<rbrakk> \<Longrightarrow> open_map (subtopology X U) X' f"
unfolding open_map_def openin_subtopology_alt by blast
lemma closed_map_from_subtopology:
"\<lbrakk>closed_map X X' f; closedin X U\<rbrakk> \<Longrightarrow> closed_map (subtopology X U) X' f"
unfolding closed_map_def closedin_subtopology_alt by blast
lemma open_map_restriction:
"\<lbrakk>open_map X X' f; {x. x \<in> topspace X \<and> f x \<in> V} = U\<rbrakk>
\<Longrightarrow> open_map (subtopology X U) (subtopology X' V) f"
unfolding open_map_def openin_subtopology_alt
apply clarify
apply (rename_tac T)
apply (rule_tac x="f ` T" in image_eqI)
using openin_closedin_eq by fastforce+
lemma closed_map_restriction:
"\<lbrakk>closed_map X X' f; {x. x \<in> topspace X \<and> f x \<in> V} = U\<rbrakk>
\<Longrightarrow> closed_map (subtopology X U) (subtopology X' V) f"
unfolding closed_map_def closedin_subtopology_alt
apply clarify
apply (rename_tac T)
apply (rule_tac x="f ` T" in image_eqI)
using closedin_def by fastforce+
subsection\<open>Quotient maps\<close>
definition quotient_map where
"quotient_map X X' f \<longleftrightarrow>
f ` (topspace X) = topspace X' \<and>
(\<forall>U. U \<subseteq> topspace X' \<longrightarrow> (openin X {x. x \<in> topspace X \<and> f x \<in> U} \<longleftrightarrow> openin X' U))"
lemma quotient_map_eq:
assumes "quotient_map X X' f" "\<And>x. x \<in> topspace X \<Longrightarrow> f x = g x"
shows "quotient_map X X' g"
proof -
have eq: "{x \<in> topspace X. f x \<in> U} = {x \<in> topspace X. g x \<in> U}" for U
using assms by auto
show ?thesis
using assms
unfolding quotient_map_def
by (metis (mono_tags, lifting) eq image_cong)
qed
lemma quotient_map_compose:
assumes f: "quotient_map X X' f" and g: "quotient_map X' X'' g"
shows "quotient_map X X'' (g \<circ> f)"
unfolding quotient_map_def
proof (intro conjI allI impI)
show "(g \<circ> f) ` topspace X = topspace X''"
using assms by (simp only: image_comp [symmetric]) (simp add: quotient_map_def)
next
fix U''
assume "U'' \<subseteq> topspace X''"
define U' where "U' \<equiv> {y \<in> topspace X'. g y \<in> U''}"
have "U' \<subseteq> topspace X'"
by (auto simp add: U'_def)
then have U': "openin X {x \<in> topspace X. f x \<in> U'} = openin X' U'"
using assms unfolding quotient_map_def by simp
have eq: "{x \<in> topspace X. f x \<in> topspace X' \<and> g (f x) \<in> U''} = {x \<in> topspace X. (g \<circ> f) x \<in> U''}"
using f quotient_map_def by fastforce
have "openin X {x \<in> topspace X. (g \<circ> f) x \<in> U''} = openin X {x \<in> topspace X. f x \<in> U'}"
using assms by (simp add: quotient_map_def U'_def eq)
also have "\<dots> = openin X'' U''"
using U'_def \<open>U'' \<subseteq> topspace X''\<close> U' g quotient_map_def by fastforce
finally show "openin X {x \<in> topspace X. (g \<circ> f) x \<in> U''} = openin X'' U''" .
qed
lemma quotient_map_from_composition:
assumes f: "continuous_map X X' f" and g: "continuous_map X' X'' g" and gf: "quotient_map X X'' (g \<circ> f)"
shows "quotient_map X' X'' g"
unfolding quotient_map_def
proof (intro conjI allI impI)
show "g ` topspace X' = topspace X''"
using assms unfolding continuous_map_def quotient_map_def by fastforce
next
fix U'' :: "'c set"
assume U'': "U'' \<subseteq> topspace X''"
have eq: "{x \<in> topspace X. g (f x) \<in> U''} = {x \<in> topspace X. f x \<in> {y. y \<in> topspace X' \<and> g y \<in> U''}}"
using continuous_map_def f by fastforce
show "openin X' {x \<in> topspace X'. g x \<in> U''} = openin X'' U''"
using assms unfolding continuous_map_def quotient_map_def
by (metis (mono_tags, lifting) Collect_cong U'' comp_apply eq)
qed
lemma quotient_imp_continuous_map:
"quotient_map X X' f \<Longrightarrow> continuous_map X X' f"
by (simp add: continuous_map openin_subset quotient_map_def)
lemma quotient_imp_surjective_map:
"quotient_map X X' f \<Longrightarrow> f ` (topspace X) = topspace X'"
by (simp add: quotient_map_def)
lemma quotient_map_closedin:
"quotient_map X X' f \<longleftrightarrow>
f ` (topspace X) = topspace X' \<and>
(\<forall>U. U \<subseteq> topspace X' \<longrightarrow> (closedin X {x. x \<in> topspace X \<and> f x \<in> U} \<longleftrightarrow> closedin X' U))"
proof -
have eq: "(topspace X - {x \<in> topspace X. f x \<in> U'}) = {x \<in> topspace X. f x \<in> topspace X' \<and> f x \<notin> U'}"
if "f ` topspace X = topspace X'" "U' \<subseteq> topspace X'" for U'
using that by auto
have "(\<forall>U\<subseteq>topspace X'. openin X {x \<in> topspace X. f x \<in> U} = openin X' U) =
(\<forall>U\<subseteq>topspace X'. closedin X {x \<in> topspace X. f x \<in> U} = closedin X' U)"
if "f ` topspace X = topspace X'"
proof (rule iffI; intro allI impI subsetI)
fix U'
assume *[rule_format]: "\<forall>U\<subseteq>topspace X'. openin X {x \<in> topspace X. f x \<in> U} = openin X' U"
and U': "U' \<subseteq> topspace X'"
show "closedin X {x \<in> topspace X. f x \<in> U'} = closedin X' U'"
using U' by (auto simp add: closedin_def simp flip: * [of "topspace X' - U'"] eq [OF that])
next
fix U' :: "'b set"
assume *[rule_format]: "\<forall>U\<subseteq>topspace X'. closedin X {x \<in> topspace X. f x \<in> U} = closedin X' U"
and U': "U' \<subseteq> topspace X'"
show "openin X {x \<in> topspace X. f x \<in> U'} = openin X' U'"
using U' by (auto simp add: openin_closedin_eq simp flip: * [of "topspace X' - U'"] eq [OF that])
qed
then show ?thesis
unfolding quotient_map_def by force
qed
lemma continuous_open_imp_quotient_map:
assumes "continuous_map X X' f" and om: "open_map X X' f" and feq: "f ` (topspace X) = topspace X'"
shows "quotient_map X X' f"
proof -
{ fix U
assume U: "U \<subseteq> topspace X'" and "openin X {x \<in> topspace X. f x \<in> U}"
then have ope: "openin X' (f ` {x \<in> topspace X. f x \<in> U})"
using om unfolding open_map_def by blast
then have "openin X' U"
using U feq by (subst openin_subopen) force
}
moreover have "openin X {x \<in> topspace X. f x \<in> U}" if "U \<subseteq> topspace X'" and "openin X' U" for U
using that assms unfolding continuous_map_def by blast
ultimately show ?thesis
unfolding quotient_map_def using assms by blast
qed
lemma continuous_closed_imp_quotient_map:
assumes "continuous_map X X' f" and om: "closed_map X X' f" and feq: "f ` (topspace X) = topspace X'"
shows "quotient_map X X' f"
proof -
have "f ` {x \<in> topspace X. f x \<in> U} = U" if "U \<subseteq> topspace X'" for U
using that feq by auto
with assms show ?thesis
unfolding quotient_map_closedin closed_map_def continuous_map_closedin by auto
qed
lemma continuous_open_quotient_map:
"\<lbrakk>continuous_map X X' f; open_map X X' f\<rbrakk> \<Longrightarrow> quotient_map X X' f \<longleftrightarrow> f ` (topspace X) = topspace X'"
by (meson continuous_open_imp_quotient_map quotient_map_def)
lemma continuous_closed_quotient_map:
"\<lbrakk>continuous_map X X' f; closed_map X X' f\<rbrakk> \<Longrightarrow> quotient_map X X' f \<longleftrightarrow> f ` (topspace X) = topspace X'"
by (meson continuous_closed_imp_quotient_map quotient_map_def)
lemma injective_quotient_map:
assumes "inj_on f (topspace X)"
shows "quotient_map X X' f \<longleftrightarrow>
continuous_map X X' f \<and> open_map X X' f \<and> closed_map X X' f \<and> f ` (topspace X) = topspace X'"
(is "?lhs = ?rhs")
proof
assume L: ?lhs
have "open_map X X' f"
proof (clarsimp simp add: open_map_def)
fix U
assume "openin X U"
then have "U \<subseteq> topspace X"
by (simp add: openin_subset)
moreover have "{x \<in> topspace X. f x \<in> f ` U} = U"
using \<open>U \<subseteq> topspace X\<close> assms inj_onD by fastforce
ultimately show "openin X' (f ` U)"
using L unfolding quotient_map_def
by (metis (no_types, lifting) Collect_cong \<open>openin X U\<close> image_mono)
qed
moreover have "closed_map X X' f"
proof (clarsimp simp add: closed_map_def)
fix U
assume "closedin X U"
then have "U \<subseteq> topspace X"
by (simp add: closedin_subset)
moreover have "{x \<in> topspace X. f x \<in> f ` U} = U"
using \<open>U \<subseteq> topspace X\<close> assms inj_onD by fastforce
ultimately show "closedin X' (f ` U)"
using L unfolding quotient_map_closedin
by (metis (no_types, lifting) Collect_cong \<open>closedin X U\<close> image_mono)
qed
ultimately show ?rhs
using L by (simp add: quotient_imp_continuous_map quotient_imp_surjective_map)
next
assume ?rhs
then show ?lhs
by (simp add: continuous_closed_imp_quotient_map)
qed
lemma continuous_compose_quotient_map:
assumes f: "quotient_map X X' f" and g: "continuous_map X X'' (g \<circ> f)"
shows "continuous_map X' X'' g"
unfolding quotient_map_def continuous_map_def
proof (intro conjI ballI allI impI)
show "\<And>x'. x' \<in> topspace X' \<Longrightarrow> g x' \<in> topspace X''"
using assms unfolding quotient_map_def
by (metis (no_types, hide_lams) continuous_map_image_subset_topspace image_comp image_subset_iff)
next
fix U'' :: "'c set"
assume U'': "openin X'' U''"
have "f ` topspace X = topspace X'"
by (simp add: f quotient_imp_surjective_map)
then have eq: "{x \<in> topspace X. f x \<in> topspace X' \<and> g (f x) \<in> U} = {x \<in> topspace X. g (f x) \<in> U}" for U
by auto
have "openin X {x \<in> topspace X. f x \<in> topspace X' \<and> g (f x) \<in> U''}"
unfolding eq using U'' g openin_continuous_map_preimage by fastforce
then have *: "openin X {x \<in> topspace X. f x \<in> {x \<in> topspace X'. g x \<in> U''}}"
by auto
show "openin X' {x \<in> topspace X'. g x \<in> U''}"
using f unfolding quotient_map_def
by (metis (no_types) Collect_subset *)
qed
lemma continuous_compose_quotient_map_eq:
"quotient_map X X' f \<Longrightarrow> continuous_map X X'' (g \<circ> f) \<longleftrightarrow> continuous_map X' X'' g"
using continuous_compose_quotient_map continuous_map_compose quotient_imp_continuous_map by blast
lemma quotient_map_compose_eq:
"quotient_map X X' f \<Longrightarrow> quotient_map X X'' (g \<circ> f) \<longleftrightarrow> quotient_map X' X'' g"
apply safe
apply (meson continuous_compose_quotient_map_eq quotient_imp_continuous_map quotient_map_from_composition)
by (simp add: quotient_map_compose)
lemma quotient_map_restriction:
assumes quo: "quotient_map X Y f" and U: "{x \<in> topspace X. f x \<in> V} = U" and disj: "openin Y V \<or> closedin Y V"
shows "quotient_map (subtopology X U) (subtopology Y V) f"
using disj
proof
assume V: "openin Y V"
with U have sub: "U \<subseteq> topspace X" "V \<subseteq> topspace Y"
by (auto simp: openin_subset)
have fim: "f ` topspace X = topspace Y"
and Y: "\<And>U. U \<subseteq> topspace Y \<Longrightarrow> openin X {x \<in> topspace X. f x \<in> U} = openin Y U"
using quo unfolding quotient_map_def by auto
have "openin X U"
using U V Y sub(2) by blast
show ?thesis
unfolding quotient_map_def
proof (intro conjI allI impI)
show "f ` topspace (subtopology X U) = topspace (subtopology Y V)"
using sub U fim by (auto)
next
fix Y' :: "'b set"
assume "Y' \<subseteq> topspace (subtopology Y V)"
then have "Y' \<subseteq> topspace Y" "Y' \<subseteq> V"
by (simp_all)
then have eq: "{x \<in> topspace X. x \<in> U \<and> f x \<in> Y'} = {x \<in> topspace X. f x \<in> Y'}"
using U by blast
then show "openin (subtopology X U) {x \<in> topspace (subtopology X U). f x \<in> Y'} = openin (subtopology Y V) Y'"
using U V Y \<open>openin X U\<close> \<open>Y' \<subseteq> topspace Y\<close> \<open>Y' \<subseteq> V\<close>
by (simp add: openin_open_subtopology eq) (auto simp: openin_closedin_eq)
qed
next
assume V: "closedin Y V"
with U have sub: "U \<subseteq> topspace X" "V \<subseteq> topspace Y"
by (auto simp: closedin_subset)
have fim: "f ` topspace X = topspace Y"
and Y: "\<And>U. U \<subseteq> topspace Y \<Longrightarrow> closedin X {x \<in> topspace X. f x \<in> U} = closedin Y U"
using quo unfolding quotient_map_closedin by auto
have "closedin X U"
using U V Y sub(2) by blast
show ?thesis
unfolding quotient_map_closedin
proof (intro conjI allI impI)
show "f ` topspace (subtopology X U) = topspace (subtopology Y V)"
using sub U fim by (auto)
next
fix Y' :: "'b set"
assume "Y' \<subseteq> topspace (subtopology Y V)"
then have "Y' \<subseteq> topspace Y" "Y' \<subseteq> V"
by (simp_all)
then have eq: "{x \<in> topspace X. x \<in> U \<and> f x \<in> Y'} = {x \<in> topspace X. f x \<in> Y'}"
using U by blast
then show "closedin (subtopology X U) {x \<in> topspace (subtopology X U). f x \<in> Y'} = closedin (subtopology Y V) Y'"
using U V Y \<open>closedin X U\<close> \<open>Y' \<subseteq> topspace Y\<close> \<open>Y' \<subseteq> V\<close>
by (simp add: closedin_closed_subtopology eq) (auto simp: closedin_def)
qed
qed
lemma quotient_map_saturated_open:
"quotient_map X Y f \<longleftrightarrow>
continuous_map X Y f \<and> f ` (topspace X) = topspace Y \<and>
(\<forall>U. openin X U \<and> {x \<in> topspace X. f x \<in> f ` U} \<subseteq> U \<longrightarrow> openin Y (f ` U))"
(is "?lhs = ?rhs")
proof
assume L: ?lhs
then have fim: "f ` topspace X = topspace Y"
and Y: "\<And>U. U \<subseteq> topspace Y \<Longrightarrow> openin Y U = openin X {x \<in> topspace X. f x \<in> U}"
unfolding quotient_map_def by auto
show ?rhs
proof (intro conjI allI impI)
show "continuous_map X Y f"
by (simp add: L quotient_imp_continuous_map)
show "f ` topspace X = topspace Y"
by (simp add: fim)
next
fix U :: "'a set"
assume U: "openin X U \<and> {x \<in> topspace X. f x \<in> f ` U} \<subseteq> U"
then have sub: "f ` U \<subseteq> topspace Y" and eq: "{x \<in> topspace X. f x \<in> f ` U} = U"
using fim openin_subset by fastforce+
show "openin Y (f ` U)"
by (simp add: sub Y eq U)
qed
next
assume ?rhs
then have YX: "\<And>U. openin Y U \<Longrightarrow> openin X {x \<in> topspace X. f x \<in> U}"
and fim: "f ` topspace X = topspace Y"
and XY: "\<And>U. \<lbrakk>openin X U; {x \<in> topspace X. f x \<in> f ` U} \<subseteq> U\<rbrakk> \<Longrightarrow> openin Y (f ` U)"
by (auto simp: quotient_map_def continuous_map_def)
show ?lhs
proof (simp add: quotient_map_def fim, intro allI impI iffI)
fix U :: "'b set"
assume "U \<subseteq> topspace Y" and X: "openin X {x \<in> topspace X. f x \<in> U}"
have feq: "f ` {x \<in> topspace X. f x \<in> U} = U"
using \<open>U \<subseteq> topspace Y\<close> fim by auto
show "openin Y U"
using XY [OF X] by (simp add: feq)
next
fix U :: "'b set"
assume "U \<subseteq> topspace Y" and Y: "openin Y U"
show "openin X {x \<in> topspace X. f x \<in> U}"
by (metis YX [OF Y])
qed
qed
subsection\<open> Separated Sets\<close>
definition separatedin :: "'a topology \<Rightarrow> 'a set \<Rightarrow> 'a set \<Rightarrow> bool"
where "separatedin X S T \<equiv>
S \<subseteq> topspace X \<and> T \<subseteq> topspace X \<and>
S \<inter> X closure_of T = {} \<and> T \<inter> X closure_of S = {}"
lemma separatedin_empty [simp]:
"separatedin X S {} \<longleftrightarrow> S \<subseteq> topspace X"
"separatedin X {} S \<longleftrightarrow> S \<subseteq> topspace X"
by (simp_all add: separatedin_def)
lemma separatedin_refl [simp]:
"separatedin X S S \<longleftrightarrow> S = {}"
proof -
have "\<And>x. \<lbrakk>separatedin X S S; x \<in> S\<rbrakk> \<Longrightarrow> False"
by (metis all_not_in_conv closure_of_subset inf.orderE separatedin_def)
then show ?thesis
by auto
qed
lemma separatedin_sym:
"separatedin X S T \<longleftrightarrow> separatedin X T S"
by (auto simp: separatedin_def)
lemma separatedin_imp_disjoint:
"separatedin X S T \<Longrightarrow> disjnt S T"
by (meson closure_of_subset disjnt_def disjnt_subset2 separatedin_def)
lemma separatedin_mono:
"\<lbrakk>separatedin X S T; S' \<subseteq> S; T' \<subseteq> T\<rbrakk> \<Longrightarrow> separatedin X S' T'"
unfolding separatedin_def
using closure_of_mono by blast
lemma separatedin_open_sets:
"\<lbrakk>openin X S; openin X T\<rbrakk> \<Longrightarrow> separatedin X S T \<longleftrightarrow> disjnt S T"
unfolding disjnt_def separatedin_def
by (auto simp: openin_Int_closure_of_eq_empty openin_subset)
lemma separatedin_closed_sets:
"\<lbrakk>closedin X S; closedin X T\<rbrakk> \<Longrightarrow> separatedin X S T \<longleftrightarrow> disjnt S T"
by (metis closedin_def closure_of_eq disjnt_def inf_commute separatedin_def)
lemma separatedin_subtopology:
"separatedin (subtopology X U) S T \<longleftrightarrow> S \<subseteq> U \<and> T \<subseteq> U \<and> separatedin X S T"
apply (simp add: separatedin_def closure_of_subtopology)
apply (safe; metis Int_absorb1 inf.assoc inf.orderE insert_disjoint(2) mk_disjoint_insert)
done
lemma separatedin_discrete_topology:
"separatedin (discrete_topology U) S T \<longleftrightarrow> S \<subseteq> U \<and> T \<subseteq> U \<and> disjnt S T"
by (metis openin_discrete_topology separatedin_def separatedin_open_sets topspace_discrete_topology)
lemma separated_eq_distinguishable:
"separatedin X {x} {y} \<longleftrightarrow>
x \<in> topspace X \<and> y \<in> topspace X \<and>
(\<exists>U. openin X U \<and> x \<in> U \<and> (y \<notin> U)) \<and>
(\<exists>v. openin X v \<and> y \<in> v \<and> (x \<notin> v))"
by (force simp: separatedin_def closure_of_def)
lemma separatedin_Un [simp]:
"separatedin X S (T \<union> U) \<longleftrightarrow> separatedin X S T \<and> separatedin X S U"
"separatedin X (S \<union> T) U \<longleftrightarrow> separatedin X S U \<and> separatedin X T U"
by (auto simp: separatedin_def)
lemma separatedin_Union:
"finite \<F> \<Longrightarrow> separatedin X S (\<Union>\<F>) \<longleftrightarrow> S \<subseteq> topspace X \<and> (\<forall>T \<in> \<F>. separatedin X S T)"
"finite \<F> \<Longrightarrow> separatedin X (\<Union>\<F>) S \<longleftrightarrow> (\<forall>T \<in> \<F>. separatedin X S T) \<and> S \<subseteq> topspace X"
by (auto simp: separatedin_def closure_of_Union)
lemma separatedin_openin_diff:
"\<lbrakk>openin X S; openin X T\<rbrakk> \<Longrightarrow> separatedin X (S - T) (T - S)"
unfolding separatedin_def
apply (intro conjI)
apply (meson Diff_subset openin_subset subset_trans)+
using openin_Int_closure_of_eq_empty by fastforce+
lemma separatedin_closedin_diff:
"\<lbrakk>closedin X S; closedin X T\<rbrakk> \<Longrightarrow> separatedin X (S - T) (T - S)"
apply (simp add: separatedin_def Diff_Int_distrib2 closure_of_minimal inf_absorb2)
apply (meson Diff_subset closedin_subset subset_trans)
done
lemma separation_closedin_Un_gen:
"separatedin X S T \<longleftrightarrow>
S \<subseteq> topspace X \<and> T \<subseteq> topspace X \<and> disjnt S T \<and>
closedin (subtopology X (S \<union> T)) S \<and>
closedin (subtopology X (S \<union> T)) T"
apply (simp add: separatedin_def closedin_Int_closure_of disjnt_iff)
using closure_of_subset apply blast
done
lemma separation_openin_Un_gen:
"separatedin X S T \<longleftrightarrow>
S \<subseteq> topspace X \<and> T \<subseteq> topspace X \<and> disjnt S T \<and>
openin (subtopology X (S \<union> T)) S \<and>
openin (subtopology X (S \<union> T)) T"
unfolding openin_closedin_eq topspace_subtopology separation_closedin_Un_gen disjnt_def
by (auto simp: Diff_triv Int_commute Un_Diff inf_absorb1 topspace_def)
subsection\<open>Homeomorphisms\<close>
text\<open>(1-way and 2-way versions may be useful in places)\<close>
definition homeomorphic_map :: "'a topology \<Rightarrow> 'b topology \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
where
"homeomorphic_map X Y f \<equiv> quotient_map X Y f \<and> inj_on f (topspace X)"
definition homeomorphic_maps :: "'a topology \<Rightarrow> 'b topology \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> bool"
where
"homeomorphic_maps X Y f g \<equiv>
continuous_map X Y f \<and> continuous_map Y X g \<and>
(\<forall>x \<in> topspace X. g(f x) = x) \<and> (\<forall>y \<in> topspace Y. f(g y) = y)"
lemma homeomorphic_map_eq:
"\<lbrakk>homeomorphic_map X Y f; \<And>x. x \<in> topspace X \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> homeomorphic_map X Y g"
by (meson homeomorphic_map_def inj_on_cong quotient_map_eq)
lemma homeomorphic_maps_eq:
"\<lbrakk>homeomorphic_maps X Y f g;
\<And>x. x \<in> topspace X \<Longrightarrow> f x = f' x; \<And>y. y \<in> topspace Y \<Longrightarrow> g y = g' y\<rbrakk>
\<Longrightarrow> homeomorphic_maps X Y f' g'"
apply (simp add: homeomorphic_maps_def)
by (metis continuous_map_eq continuous_map_eq_image_closure_subset_gen image_subset_iff)
lemma homeomorphic_maps_sym:
"homeomorphic_maps X Y f g \<longleftrightarrow> homeomorphic_maps Y X g f"
by (auto simp: homeomorphic_maps_def)
lemma homeomorphic_maps_id:
"homeomorphic_maps X Y id id \<longleftrightarrow> Y = X"
(is "?lhs = ?rhs")
proof
assume L: ?lhs
then have "topspace X = topspace Y"
by (auto simp: homeomorphic_maps_def continuous_map_def)
with L show ?rhs
unfolding homeomorphic_maps_def
by (metis topology_finer_continuous_id topology_eq)
next
assume ?rhs
then show ?lhs
unfolding homeomorphic_maps_def by auto
qed
lemma homeomorphic_map_id [simp]: "homeomorphic_map X Y id \<longleftrightarrow> Y = X"
(is "?lhs = ?rhs")
proof
assume L: ?lhs
then have eq: "topspace X = topspace Y"
by (auto simp: homeomorphic_map_def continuous_map_def quotient_map_def)
then have "\<And>S. openin X S \<longrightarrow> openin Y S"
by (meson L homeomorphic_map_def injective_quotient_map topology_finer_open_id)
then show ?rhs
using L unfolding homeomorphic_map_def
by (metis eq quotient_imp_continuous_map topology_eq topology_finer_continuous_id)
next
assume ?rhs
then show ?lhs
unfolding homeomorphic_map_def
by (simp add: closed_map_id continuous_closed_imp_quotient_map)
qed
-lemma homeomorphic_maps_i [simp]:"homeomorphic_maps X Y id id \<longleftrightarrow> Y = X"
- by (metis (full_types) eq_id_iff homeomorphic_maps_id)
-
-lemma homeomorphic_map_i [simp]: "homeomorphic_map X Y id \<longleftrightarrow> Y = X"
- by (metis (no_types) eq_id_iff homeomorphic_map_id)
-
lemma homeomorphic_map_compose:
assumes "homeomorphic_map X Y f" "homeomorphic_map Y X'' g"
shows "homeomorphic_map X X'' (g \<circ> f)"
proof -
have "inj_on g (f ` topspace X)"
by (metis (no_types) assms homeomorphic_map_def quotient_imp_surjective_map)
then show ?thesis
using assms by (meson comp_inj_on homeomorphic_map_def quotient_map_compose_eq)
qed
lemma homeomorphic_maps_compose:
"homeomorphic_maps X Y f h \<and>
homeomorphic_maps Y X'' g k
\<Longrightarrow> homeomorphic_maps X X'' (g \<circ> f) (h \<circ> k)"
unfolding homeomorphic_maps_def
by (auto simp: continuous_map_compose; simp add: continuous_map_def)
lemma homeomorphic_eq_everything_map:
"homeomorphic_map X Y f \<longleftrightarrow>
continuous_map X Y f \<and> open_map X Y f \<and> closed_map X Y f \<and>
f ` (topspace X) = topspace Y \<and> inj_on f (topspace X)"
unfolding homeomorphic_map_def
by (force simp: injective_quotient_map intro: injective_quotient_map)
lemma homeomorphic_imp_continuous_map:
"homeomorphic_map X Y f \<Longrightarrow> continuous_map X Y f"
by (simp add: homeomorphic_eq_everything_map)
lemma homeomorphic_imp_open_map:
"homeomorphic_map X Y f \<Longrightarrow> open_map X Y f"
by (simp add: homeomorphic_eq_everything_map)
lemma homeomorphic_imp_closed_map:
"homeomorphic_map X Y f \<Longrightarrow> closed_map X Y f"
by (simp add: homeomorphic_eq_everything_map)
lemma homeomorphic_imp_surjective_map:
"homeomorphic_map X Y f \<Longrightarrow> f ` (topspace X) = topspace Y"
by (simp add: homeomorphic_eq_everything_map)
lemma homeomorphic_imp_injective_map:
"homeomorphic_map X Y f \<Longrightarrow> inj_on f (topspace X)"
by (simp add: homeomorphic_eq_everything_map)
lemma bijective_open_imp_homeomorphic_map:
"\<lbrakk>continuous_map X Y f; open_map X Y f; f ` (topspace X) = topspace Y; inj_on f (topspace X)\<rbrakk>
\<Longrightarrow> homeomorphic_map X Y f"
by (simp add: homeomorphic_map_def continuous_open_imp_quotient_map)
lemma bijective_closed_imp_homeomorphic_map:
"\<lbrakk>continuous_map X Y f; closed_map X Y f; f ` (topspace X) = topspace Y; inj_on f (topspace X)\<rbrakk>
\<Longrightarrow> homeomorphic_map X Y f"
by (simp add: continuous_closed_quotient_map homeomorphic_map_def)
lemma open_eq_continuous_inverse_map:
assumes X: "\<And>x. x \<in> topspace X \<Longrightarrow> f x \<in> topspace Y \<and> g(f x) = x"
and Y: "\<And>y. y \<in> topspace Y \<Longrightarrow> g y \<in> topspace X \<and> f(g y) = y"
shows "open_map X Y f \<longleftrightarrow> continuous_map Y X g"
proof -
have eq: "{x \<in> topspace Y. g x \<in> U} = f ` U" if "openin X U" for U
using openin_subset [OF that] by (force simp: X Y image_iff)
show ?thesis
by (auto simp: Y open_map_def continuous_map_def eq)
qed
lemma closed_eq_continuous_inverse_map:
assumes X: "\<And>x. x \<in> topspace X \<Longrightarrow> f x \<in> topspace Y \<and> g(f x) = x"
and Y: "\<And>y. y \<in> topspace Y \<Longrightarrow> g y \<in> topspace X \<and> f(g y) = y"
shows "closed_map X Y f \<longleftrightarrow> continuous_map Y X g"
proof -
have eq: "{x \<in> topspace Y. g x \<in> U} = f ` U" if "closedin X U" for U
using closedin_subset [OF that] by (force simp: X Y image_iff)
show ?thesis
by (auto simp: Y closed_map_def continuous_map_closedin eq)
qed
lemma homeomorphic_maps_map:
"homeomorphic_maps X Y f g \<longleftrightarrow>
homeomorphic_map X Y f \<and> homeomorphic_map Y X g \<and>
(\<forall>x \<in> topspace X. g(f x) = x) \<and> (\<forall>y \<in> topspace Y. f(g y) = y)"
(is "?lhs = ?rhs")
proof
assume ?lhs
then have L: "continuous_map X Y f" "continuous_map Y X g" "\<forall>x\<in>topspace X. g (f x) = x" "\<forall>x'\<in>topspace Y. f (g x') = x'"
by (auto simp: homeomorphic_maps_def)
show ?rhs
proof (intro conjI bijective_open_imp_homeomorphic_map L)
show "open_map X Y f"
using L using open_eq_continuous_inverse_map [of concl: X Y f g] by (simp add: continuous_map_def)
show "open_map Y X g"
using L using open_eq_continuous_inverse_map [of concl: Y X g f] by (simp add: continuous_map_def)
show "f ` topspace X = topspace Y" "g ` topspace Y = topspace X"
using L by (force simp: continuous_map_closedin)+
show "inj_on f (topspace X)" "inj_on g (topspace Y)"
using L unfolding inj_on_def by metis+
qed
next
assume ?rhs
then show ?lhs
by (auto simp: homeomorphic_maps_def homeomorphic_imp_continuous_map)
qed
lemma homeomorphic_maps_imp_map:
"homeomorphic_maps X Y f g \<Longrightarrow> homeomorphic_map X Y f"
using homeomorphic_maps_map by blast
lemma homeomorphic_map_maps:
"homeomorphic_map X Y f \<longleftrightarrow> (\<exists>g. homeomorphic_maps X Y f g)"
(is "?lhs = ?rhs")
proof
assume ?lhs
then have L: "continuous_map X Y f" "open_map X Y f" "closed_map X Y f"
"f ` (topspace X) = topspace Y" "inj_on f (topspace X)"
by (auto simp: homeomorphic_eq_everything_map)
have X: "\<And>x. x \<in> topspace X \<Longrightarrow> f x \<in> topspace Y \<and> inv_into (topspace X) f (f x) = x"
using L by auto
have Y: "\<And>y. y \<in> topspace Y \<Longrightarrow> inv_into (topspace X) f y \<in> topspace X \<and> f (inv_into (topspace X) f y) = y"
by (simp add: L f_inv_into_f inv_into_into)
have "homeomorphic_maps X Y f (inv_into (topspace X) f)"
unfolding homeomorphic_maps_def
proof (intro conjI L)
show "continuous_map Y X (inv_into (topspace X) f)"
by (simp add: L X Y flip: open_eq_continuous_inverse_map [where f=f])
next
show "\<forall>x\<in>topspace X. inv_into (topspace X) f (f x) = x"
"\<forall>y\<in>topspace Y. f (inv_into (topspace X) f y) = y"
using X Y by auto
qed
then show ?rhs
by metis
next
assume ?rhs
then show ?lhs
using homeomorphic_maps_map by blast
qed
lemma homeomorphic_maps_involution:
"\<lbrakk>continuous_map X X f; \<And>x. x \<in> topspace X \<Longrightarrow> f(f x) = x\<rbrakk> \<Longrightarrow> homeomorphic_maps X X f f"
by (auto simp: homeomorphic_maps_def)
lemma homeomorphic_map_involution:
"\<lbrakk>continuous_map X X f; \<And>x. x \<in> topspace X \<Longrightarrow> f(f x) = x\<rbrakk> \<Longrightarrow> homeomorphic_map X X f"
using homeomorphic_maps_involution homeomorphic_maps_map by blast
lemma homeomorphic_map_openness:
assumes hom: "homeomorphic_map X Y f" and U: "U \<subseteq> topspace X"
shows "openin Y (f ` U) \<longleftrightarrow> openin X U"
proof -
obtain g where "homeomorphic_maps X Y f g"
using assms by (auto simp: homeomorphic_map_maps)
then have g: "homeomorphic_map Y X g" and gf: "\<And>x. x \<in> topspace X \<Longrightarrow> g(f x) = x"
by (auto simp: homeomorphic_maps_map)
then have "openin X U \<Longrightarrow> openin Y (f ` U)"
using hom homeomorphic_imp_open_map open_map_def by blast
show "openin Y (f ` U) = openin X U"
proof
assume L: "openin Y (f ` U)"
have "U = g ` (f ` U)"
using U gf by force
then show "openin X U"
by (metis L homeomorphic_imp_open_map open_map_def g)
next
assume "openin X U"
then show "openin Y (f ` U)"
using hom homeomorphic_imp_open_map open_map_def by blast
qed
qed
lemma homeomorphic_map_closedness:
assumes hom: "homeomorphic_map X Y f" and U: "U \<subseteq> topspace X"
shows "closedin Y (f ` U) \<longleftrightarrow> closedin X U"
proof -
obtain g where "homeomorphic_maps X Y f g"
using assms by (auto simp: homeomorphic_map_maps)
then have g: "homeomorphic_map Y X g" and gf: "\<And>x. x \<in> topspace X \<Longrightarrow> g(f x) = x"
by (auto simp: homeomorphic_maps_map)
then have "closedin X U \<Longrightarrow> closedin Y (f ` U)"
using hom homeomorphic_imp_closed_map closed_map_def by blast
show "closedin Y (f ` U) = closedin X U"
proof
assume L: "closedin Y (f ` U)"
have "U = g ` (f ` U)"
using U gf by force
then show "closedin X U"
by (metis L homeomorphic_imp_closed_map closed_map_def g)
next
assume "closedin X U"
then show "closedin Y (f ` U)"
using hom homeomorphic_imp_closed_map closed_map_def by blast
qed
qed
lemma homeomorphic_map_openness_eq:
"homeomorphic_map X Y f \<Longrightarrow> openin X U \<longleftrightarrow> U \<subseteq> topspace X \<and> openin Y (f ` U)"
by (meson homeomorphic_map_openness openin_closedin_eq)
lemma homeomorphic_map_closedness_eq:
"homeomorphic_map X Y f \<Longrightarrow> closedin X U \<longleftrightarrow> U \<subseteq> topspace X \<and> closedin Y (f ` U)"
by (meson closedin_subset homeomorphic_map_closedness)
lemma all_openin_homeomorphic_image:
assumes "homeomorphic_map X Y f"
shows "(\<forall>V. openin Y V \<longrightarrow> P V) \<longleftrightarrow> (\<forall>U. openin X U \<longrightarrow> P(f ` U))" (is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
by (meson assms homeomorphic_map_openness_eq)
next
assume ?rhs
then show ?lhs
by (metis (no_types, lifting) assms homeomorphic_imp_surjective_map homeomorphic_map_openness openin_subset subset_image_iff)
qed
lemma all_closedin_homeomorphic_image:
assumes "homeomorphic_map X Y f"
shows "(\<forall>V. closedin Y V \<longrightarrow> P V) \<longleftrightarrow> (\<forall>U. closedin X U \<longrightarrow> P(f ` U))" (is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
by (meson assms homeomorphic_map_closedness_eq)
next
assume ?rhs
then show ?lhs
by (metis (no_types, lifting) assms homeomorphic_imp_surjective_map homeomorphic_map_closedness closedin_subset subset_image_iff)
qed
lemma homeomorphic_map_derived_set_of:
assumes hom: "homeomorphic_map X Y f" and S: "S \<subseteq> topspace X"
shows "Y derived_set_of (f ` S) = f ` (X derived_set_of S)"
proof -
have fim: "f ` (topspace X) = topspace Y" and inj: "inj_on f (topspace X)"
using hom by (auto simp: homeomorphic_eq_everything_map)
have iff: "(\<forall>T. x \<in> T \<and> openin X T \<longrightarrow> (\<exists>y. y \<noteq> x \<and> y \<in> S \<and> y \<in> T)) =
(\<forall>T. T \<subseteq> topspace Y \<longrightarrow> f x \<in> T \<longrightarrow> openin Y T \<longrightarrow> (\<exists>y. y \<noteq> f x \<and> y \<in> f ` S \<and> y \<in> T))"
if "x \<in> topspace X" for x
proof -
have 1: "(x \<in> T \<and> openin X T) = (T \<subseteq> topspace X \<and> f x \<in> f ` T \<and> openin Y (f ` T))" for T
by (meson hom homeomorphic_map_openness_eq inj inj_on_image_mem_iff that)
have 2: "(\<exists>y. y \<noteq> x \<and> y \<in> S \<and> y \<in> T) = (\<exists>y. y \<noteq> f x \<and> y \<in> f ` S \<and> y \<in> f ` T)" (is "?lhs = ?rhs")
if "T \<subseteq> topspace X \<and> f x \<in> f ` T \<and> openin Y (f ` T)" for T
proof
show "?lhs \<Longrightarrow> ?rhs"
by (meson "1" imageI inj inj_on_eq_iff inj_on_subset that)
show "?rhs \<Longrightarrow> ?lhs"
using S inj inj_onD that by fastforce
qed
show ?thesis
apply (simp flip: fim add: all_subset_image)
apply (simp flip: imp_conjL)
by (intro all_cong1 imp_cong 1 2)
qed
have *: "\<lbrakk>T = f ` S; \<And>x. x \<in> S \<Longrightarrow> P x \<longleftrightarrow> Q(f x)\<rbrakk> \<Longrightarrow> {y. y \<in> T \<and> Q y} = f ` {x \<in> S. P x}" for T S P Q
by auto
show ?thesis
unfolding derived_set_of_def
apply (rule *)
using fim apply blast
using iff openin_subset by force
qed
lemma homeomorphic_map_closure_of:
assumes hom: "homeomorphic_map X Y f" and S: "S \<subseteq> topspace X"
shows "Y closure_of (f ` S) = f ` (X closure_of S)"
unfolding closure_of
using homeomorphic_imp_surjective_map [OF hom] S
by (auto simp: in_derived_set_of homeomorphic_map_derived_set_of [OF assms])
lemma homeomorphic_map_interior_of:
assumes hom: "homeomorphic_map X Y f" and S: "S \<subseteq> topspace X"
shows "Y interior_of (f ` S) = f ` (X interior_of S)"
proof -
{ fix y
assume "y \<in> topspace Y" and "y \<notin> Y closure_of (topspace Y - f ` S)"
then have "y \<in> f ` (topspace X - X closure_of (topspace X - S))"
using homeomorphic_eq_everything_map [THEN iffD1, OF hom] homeomorphic_map_closure_of [OF hom]
by (metis DiffI Diff_subset S closure_of_subset_topspace inj_on_image_set_diff) }
moreover
{ fix x
assume "x \<in> topspace X"
then have "f x \<in> topspace Y"
using hom homeomorphic_imp_surjective_map by blast }
moreover
{ fix x
assume "x \<in> topspace X" and "x \<notin> X closure_of (topspace X - S)" and "f x \<in> Y closure_of (topspace Y - f ` S)"
then have "False"
using homeomorphic_map_closure_of [OF hom] hom
unfolding homeomorphic_eq_everything_map
by (metis (no_types, lifting) Diff_subset S closure_of_subset_topspace inj_on_image_mem_iff_alt inj_on_image_set_diff) }
ultimately show ?thesis
by (auto simp: interior_of_closure_of)
qed
lemma homeomorphic_map_frontier_of:
assumes hom: "homeomorphic_map X Y f" and S: "S \<subseteq> topspace X"
shows "Y frontier_of (f ` S) = f ` (X frontier_of S)"
unfolding frontier_of_def
proof (intro equalityI subsetI DiffI)
fix y
assume "y \<in> Y closure_of f ` S - Y interior_of f ` S"
then show "y \<in> f ` (X closure_of S - X interior_of S)"
using S hom homeomorphic_map_closure_of homeomorphic_map_interior_of by fastforce
next
fix y
assume "y \<in> f ` (X closure_of S - X interior_of S)"
then show "y \<in> Y closure_of f ` S"
using S hom homeomorphic_map_closure_of by fastforce
next
fix x
assume "x \<in> f ` (X closure_of S - X interior_of S)"
then obtain y where y: "x = f y" "y \<in> X closure_of S" "y \<notin> X interior_of S"
by blast
then have "y \<in> topspace X"
by (simp add: in_closure_of)
then have "f y \<notin> f ` (X interior_of S)"
by (meson hom homeomorphic_eq_everything_map inj_on_image_mem_iff_alt interior_of_subset_topspace y(3))
then show "x \<notin> Y interior_of f ` S"
using S hom homeomorphic_map_interior_of y(1) by blast
qed
lemma homeomorphic_maps_subtopologies:
"\<lbrakk>homeomorphic_maps X Y f g; f ` (topspace X \<inter> S) = topspace Y \<inter> T\<rbrakk>
\<Longrightarrow> homeomorphic_maps (subtopology X S) (subtopology Y T) f g"
unfolding homeomorphic_maps_def
by (force simp: continuous_map_from_subtopology continuous_map_in_subtopology)
lemma homeomorphic_maps_subtopologies_alt:
"\<lbrakk>homeomorphic_maps X Y f g; f ` (topspace X \<inter> S) \<subseteq> T; g ` (topspace Y \<inter> T) \<subseteq> S\<rbrakk>
\<Longrightarrow> homeomorphic_maps (subtopology X S) (subtopology Y T) f g"
unfolding homeomorphic_maps_def
by (force simp: continuous_map_from_subtopology continuous_map_in_subtopology)
lemma homeomorphic_map_subtopologies:
"\<lbrakk>homeomorphic_map X Y f; f ` (topspace X \<inter> S) = topspace Y \<inter> T\<rbrakk>
\<Longrightarrow> homeomorphic_map (subtopology X S) (subtopology Y T) f"
by (meson homeomorphic_map_maps homeomorphic_maps_subtopologies)
lemma homeomorphic_map_subtopologies_alt:
"\<lbrakk>homeomorphic_map X Y f;
\<And>x. \<lbrakk>x \<in> topspace X; f x \<in> topspace Y\<rbrakk> \<Longrightarrow> f x \<in> T \<longleftrightarrow> x \<in> S\<rbrakk>
\<Longrightarrow> homeomorphic_map (subtopology X S) (subtopology Y T) f"
unfolding homeomorphic_map_maps
apply (erule ex_forward)
apply (rule homeomorphic_maps_subtopologies)
apply (auto simp: homeomorphic_maps_def continuous_map_def)
by (metis IntI image_iff)
subsection\<open>Relation of homeomorphism between topological spaces\<close>
definition homeomorphic_space (infixr "homeomorphic'_space" 50)
where "X homeomorphic_space Y \<equiv> \<exists>f g. homeomorphic_maps X Y f g"
lemma homeomorphic_space_refl: "X homeomorphic_space X"
by (meson homeomorphic_maps_id homeomorphic_space_def)
lemma homeomorphic_space_sym:
"X homeomorphic_space Y \<longleftrightarrow> Y homeomorphic_space X"
unfolding homeomorphic_space_def by (metis homeomorphic_maps_sym)
lemma homeomorphic_space_trans [trans]:
"\<lbrakk>X1 homeomorphic_space X2; X2 homeomorphic_space X3\<rbrakk> \<Longrightarrow> X1 homeomorphic_space X3"
unfolding homeomorphic_space_def by (metis homeomorphic_maps_compose)
lemma homeomorphic_space:
"X homeomorphic_space Y \<longleftrightarrow> (\<exists>f. homeomorphic_map X Y f)"
by (simp add: homeomorphic_map_maps homeomorphic_space_def)
lemma homeomorphic_maps_imp_homeomorphic_space:
"homeomorphic_maps X Y f g \<Longrightarrow> X homeomorphic_space Y"
unfolding homeomorphic_space_def by metis
lemma homeomorphic_map_imp_homeomorphic_space:
"homeomorphic_map X Y f \<Longrightarrow> X homeomorphic_space Y"
unfolding homeomorphic_map_maps
using homeomorphic_space_def by blast
lemma homeomorphic_empty_space:
"X homeomorphic_space Y \<Longrightarrow> topspace X = {} \<longleftrightarrow> topspace Y = {}"
by (metis homeomorphic_imp_surjective_map homeomorphic_space image_is_empty)
lemma homeomorphic_empty_space_eq:
assumes "topspace X = {}"
shows "X homeomorphic_space Y \<longleftrightarrow> topspace Y = {}"
proof -
have "\<forall>f t. continuous_map X (t::'b topology) f"
using assms continuous_map_on_empty by blast
then show ?thesis
by (metis (no_types) assms continuous_map_on_empty empty_iff homeomorphic_empty_space homeomorphic_maps_def homeomorphic_space_def)
qed
subsection\<open>Connected topological spaces\<close>
definition connected_space :: "'a topology \<Rightarrow> bool" where
"connected_space X \<equiv>
\<not>(\<exists>E1 E2. openin X E1 \<and> openin X E2 \<and>
topspace X \<subseteq> E1 \<union> E2 \<and> E1 \<inter> E2 = {} \<and> E1 \<noteq> {} \<and> E2 \<noteq> {})"
definition connectedin :: "'a topology \<Rightarrow> 'a set \<Rightarrow> bool" where
"connectedin X S \<equiv> S \<subseteq> topspace X \<and> connected_space (subtopology X S)"
lemma connected_spaceD:
"\<lbrakk>connected_space X;
openin X U; openin X V; topspace X \<subseteq> U \<union> V; U \<inter> V = {}; U \<noteq> {}; V \<noteq> {}\<rbrakk> \<Longrightarrow> False"
by (auto simp: connected_space_def)
lemma connectedin_subset_topspace: "connectedin X S \<Longrightarrow> S \<subseteq> topspace X"
by (simp add: connectedin_def)
lemma connectedin_topspace:
"connectedin X (topspace X) \<longleftrightarrow> connected_space X"
by (simp add: connectedin_def)
lemma connected_space_subtopology:
"connectedin X S \<Longrightarrow> connected_space (subtopology X S)"
by (simp add: connectedin_def)
lemma connectedin_subtopology:
"connectedin (subtopology X S) T \<longleftrightarrow> connectedin X T \<and> T \<subseteq> S"
by (force simp: connectedin_def subtopology_subtopology inf_absorb2)
lemma connected_space_eq:
"connected_space X \<longleftrightarrow>
(\<nexists>E1 E2. openin X E1 \<and> openin X E2 \<and> E1 \<union> E2 = topspace X \<and> E1 \<inter> E2 = {} \<and> E1 \<noteq> {} \<and> E2 \<noteq> {})"
unfolding connected_space_def
by (metis openin_Un openin_subset subset_antisym)
lemma connected_space_closedin:
"connected_space X \<longleftrightarrow>
(\<nexists>E1 E2. closedin X E1 \<and> closedin X E2 \<and> topspace X \<subseteq> E1 \<union> E2 \<and>
E1 \<inter> E2 = {} \<and> E1 \<noteq> {} \<and> E2 \<noteq> {})" (is "?lhs = ?rhs")
proof
assume ?lhs
then have L: "\<And>E1 E2. \<lbrakk>openin X E1; E1 \<inter> E2 = {}; topspace X \<subseteq> E1 \<union> E2; openin X E2\<rbrakk> \<Longrightarrow> E1 = {} \<or> E2 = {}"
by (simp add: connected_space_def)
show ?rhs
unfolding connected_space_def
proof clarify
fix E1 E2
assume "closedin X E1" and "closedin X E2" and "topspace X \<subseteq> E1 \<union> E2" and "E1 \<inter> E2 = {}"
and "E1 \<noteq> {}" and "E2 \<noteq> {}"
have "E1 \<union> E2 = topspace X"
by (meson Un_subset_iff \<open>closedin X E1\<close> \<open>closedin X E2\<close> \<open>topspace X \<subseteq> E1 \<union> E2\<close> closedin_def subset_antisym)
then have "topspace X - E2 = E1"
using \<open>E1 \<inter> E2 = {}\<close> by fastforce
then have "topspace X = E1"
using \<open>E1 \<noteq> {}\<close> L \<open>closedin X E1\<close> \<open>closedin X E2\<close> by blast
then show "False"
using \<open>E1 \<inter> E2 = {}\<close> \<open>E1 \<union> E2 = topspace X\<close> \<open>E2 \<noteq> {}\<close> by blast
qed
next
assume R: ?rhs
show ?lhs
unfolding connected_space_def
proof clarify
fix E1 E2
assume "openin X E1" and "openin X E2" and "topspace X \<subseteq> E1 \<union> E2" and "E1 \<inter> E2 = {}"
and "E1 \<noteq> {}" and "E2 \<noteq> {}"
have "E1 \<union> E2 = topspace X"
by (meson Un_subset_iff \<open>openin X E1\<close> \<open>openin X E2\<close> \<open>topspace X \<subseteq> E1 \<union> E2\<close> openin_closedin_eq subset_antisym)
then have "topspace X - E2 = E1"
using \<open>E1 \<inter> E2 = {}\<close> by fastforce
then have "topspace X = E1"
using \<open>E1 \<noteq> {}\<close> R \<open>openin X E1\<close> \<open>openin X E2\<close> by blast
then show "False"
using \<open>E1 \<inter> E2 = {}\<close> \<open>E1 \<union> E2 = topspace X\<close> \<open>E2 \<noteq> {}\<close> by blast
qed
qed
lemma connected_space_closedin_eq:
"connected_space X \<longleftrightarrow>
(\<nexists>E1 E2. closedin X E1 \<and> closedin X E2 \<and>
E1 \<union> E2 = topspace X \<and> E1 \<inter> E2 = {} \<and> E1 \<noteq> {} \<and> E2 \<noteq> {})"
apply (simp add: connected_space_closedin)
apply (intro all_cong)
using closedin_subset apply blast
done
lemma connected_space_clopen_in:
"connected_space X \<longleftrightarrow>
(\<forall>T. openin X T \<and> closedin X T \<longrightarrow> T = {} \<or> T = topspace X)"
proof -
have eq: "openin X E1 \<and> openin X E2 \<and> E1 \<union> E2 = topspace X \<and> E1 \<inter> E2 = {} \<and> P
\<longleftrightarrow> E2 = topspace X - E1 \<and> openin X E1 \<and> openin X E2 \<and> P" for E1 E2 P
using openin_subset by blast
show ?thesis
unfolding connected_space_eq eq closedin_def
by (auto simp: openin_closedin_eq)
qed
lemma connectedin:
"connectedin X S \<longleftrightarrow>
S \<subseteq> topspace X \<and>
(\<nexists>E1 E2.
openin X E1 \<and> openin X E2 \<and>
S \<subseteq> E1 \<union> E2 \<and> E1 \<inter> E2 \<inter> S = {} \<and> E1 \<inter> S \<noteq> {} \<and> E2 \<inter> S \<noteq> {})"
proof -
have *: "(\<exists>E1:: 'a set. \<exists>E2:: 'a set. (\<exists>T1:: 'a set. P1 T1 \<and> E1 = f1 T1) \<and> (\<exists>T2:: 'a set. P2 T2 \<and> E2 = f2 T2) \<and>
R E1 E2) \<longleftrightarrow> (\<exists>T1 T2. P1 T1 \<and> P2 T2 \<and> R(f1 T1) (f2 T2))" for P1 f1 P2 f2 R
by auto
show ?thesis
unfolding connectedin_def connected_space_def openin_subtopology topspace_subtopology Not_eq_iff *
apply (intro conj_cong arg_cong [where f=Not] ex_cong1 refl)
apply (blast elim: dest!: openin_subset)+
done
qed
lemma connectedin_iff_connected [simp]: "connectedin euclidean S \<longleftrightarrow> connected S"
by (simp add: connected_def connectedin)
lemma connectedin_closedin:
"connectedin X S \<longleftrightarrow>
S \<subseteq> topspace X \<and>
\<not>(\<exists>E1 E2. closedin X E1 \<and> closedin X E2 \<and>
S \<subseteq> (E1 \<union> E2) \<and>
(E1 \<inter> E2 \<inter> S = {}) \<and>
\<not>(E1 \<inter> S = {}) \<and> \<not>(E2 \<inter> S = {}))"
proof -
have *: "(\<exists>E1:: 'a set. \<exists>E2:: 'a set. (\<exists>T1:: 'a set. P1 T1 \<and> E1 = f1 T1) \<and> (\<exists>T2:: 'a set. P2 T2 \<and> E2 = f2 T2) \<and>
R E1 E2) \<longleftrightarrow> (\<exists>T1 T2. P1 T1 \<and> P2 T2 \<and> R(f1 T1) (f2 T2))" for P1 f1 P2 f2 R
by auto
show ?thesis
unfolding connectedin_def connected_space_closedin closedin_subtopology topspace_subtopology Not_eq_iff *
apply (intro conj_cong arg_cong [where f=Not] ex_cong1 refl)
apply (blast elim: dest!: openin_subset)+
done
qed
lemma connectedin_empty [simp]: "connectedin X {}"
by (simp add: connectedin)
lemma connected_space_topspace_empty:
"topspace X = {} \<Longrightarrow> connected_space X"
using connectedin_topspace by fastforce
lemma connectedin_sing [simp]: "connectedin X {a} \<longleftrightarrow> a \<in> topspace X"
by (simp add: connectedin)
lemma connectedin_absolute [simp]:
"connectedin (subtopology X S) S \<longleftrightarrow> connectedin X S"
apply (simp only: connectedin_def topspace_subtopology subtopology_subtopology)
apply (intro conj_cong imp_cong arg_cong [where f=Not] all_cong1 ex_cong1 refl)
by auto
lemma connectedin_Union:
assumes \<U>: "\<And>S. S \<in> \<U> \<Longrightarrow> connectedin X S" and ne: "\<Inter>\<U> \<noteq> {}"
shows "connectedin X (\<Union>\<U>)"
proof -
have "\<Union>\<U> \<subseteq> topspace X"
using \<U> by (simp add: Union_least connectedin_def)
moreover have False
if "openin X E1" "openin X E2" and cover: "\<Union>\<U> \<subseteq> E1 \<union> E2" and disj: "E1 \<inter> E2 \<inter> \<Union>\<U> = {}"
and overlap1: "E1 \<inter> \<Union>\<U> \<noteq> {}" and overlap2: "E2 \<inter> \<Union>\<U> \<noteq> {}"
for E1 E2
proof -
have disjS: "E1 \<inter> E2 \<inter> S = {}" if "S \<in> \<U>" for S
using Diff_triv that disj by auto
have coverS: "S \<subseteq> E1 \<union> E2" if "S \<in> \<U>" for S
using that cover by blast
have "\<U> \<noteq> {}"
using overlap1 by blast
obtain a where a: "\<And>U. U \<in> \<U> \<Longrightarrow> a \<in> U"
using ne by force
with \<open>\<U> \<noteq> {}\<close> have "a \<in> \<Union>\<U>"
by blast
then consider "a \<in> E1" | "a \<in> E2"
using \<open>\<Union>\<U> \<subseteq> E1 \<union> E2\<close> by auto
then show False
proof cases
case 1
then obtain b S where "b \<in> E2" "b \<in> S" "S \<in> \<U>"
using overlap2 by blast
then show ?thesis
using "1" \<open>openin X E1\<close> \<open>openin X E2\<close> disjS coverS a [OF \<open>S \<in> \<U>\<close>] \<U>[OF \<open>S \<in> \<U>\<close>]
unfolding connectedin
by (meson disjoint_iff_not_equal)
next
case 2
then obtain b S where "b \<in> E1" "b \<in> S" "S \<in> \<U>"
using overlap1 by blast
then show ?thesis
using "2" \<open>openin X E1\<close> \<open>openin X E2\<close> disjS coverS a [OF \<open>S \<in> \<U>\<close>] \<U>[OF \<open>S \<in> \<U>\<close>]
unfolding connectedin
by (meson disjoint_iff_not_equal)
qed
qed
ultimately show ?thesis
unfolding connectedin by blast
qed
lemma connectedin_Un:
"\<lbrakk>connectedin X S; connectedin X T; S \<inter> T \<noteq> {}\<rbrakk> \<Longrightarrow> connectedin X (S \<union> T)"
using connectedin_Union [of "{S,T}"] by auto
lemma connected_space_subconnected:
"connected_space X \<longleftrightarrow> (\<forall>x \<in> topspace X. \<forall>y \<in> topspace X. \<exists>S. connectedin X S \<and> x \<in> S \<and> y \<in> S)" (is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
using connectedin_topspace by blast
next
assume R [rule_format]: ?rhs
have False if "openin X U" "openin X V" and disj: "U \<inter> V = {}" and cover: "topspace X \<subseteq> U \<union> V"
and "U \<noteq> {}" "V \<noteq> {}" for U V
proof -
obtain u v where "u \<in> U" "v \<in> V"
using \<open>U \<noteq> {}\<close> \<open>V \<noteq> {}\<close> by auto
then obtain T where "u \<in> T" "v \<in> T" and T: "connectedin X T"
using R [of u v] that
by (meson \<open>openin X U\<close> \<open>openin X V\<close> subsetD openin_subset)
then show False
using that unfolding connectedin
by (metis IntI \<open>u \<in> U\<close> \<open>v \<in> V\<close> empty_iff inf_bot_left subset_trans)
qed
then show ?lhs
by (auto simp: connected_space_def)
qed
lemma connectedin_intermediate_closure_of:
assumes "connectedin X S" "S \<subseteq> T" "T \<subseteq> X closure_of S"
shows "connectedin X T"
proof -
have S: "S \<subseteq> topspace X"and T: "T \<subseteq> topspace X"
using assms by (meson closure_of_subset_topspace dual_order.trans)+
show ?thesis
using assms
apply (simp add: connectedin closure_of_subset_topspace S T)
apply (elim all_forward imp_forward2 asm_rl)
apply (blast dest: openin_Int_closure_of_eq_empty [of X _ S])+
done
qed
lemma connectedin_closure_of:
"connectedin X S \<Longrightarrow> connectedin X (X closure_of S)"
by (meson closure_of_subset connectedin_def connectedin_intermediate_closure_of subset_refl)
lemma connectedin_separation:
"connectedin X S \<longleftrightarrow>
S \<subseteq> topspace X \<and>
(\<nexists>C1 C2. C1 \<union> C2 = S \<and> C1 \<noteq> {} \<and> C2 \<noteq> {} \<and> C1 \<inter> X closure_of C2 = {} \<and> C2 \<inter> X closure_of C1 = {})" (is "?lhs = ?rhs")
unfolding connectedin_def connected_space_closedin_eq closedin_Int_closure_of topspace_subtopology
apply (intro conj_cong refl arg_cong [where f=Not])
apply (intro ex_cong1 iffI, blast)
using closure_of_subset_Int by force
lemma connectedin_eq_not_separated:
"connectedin X S \<longleftrightarrow>
S \<subseteq> topspace X \<and>
(\<nexists>C1 C2. C1 \<union> C2 = S \<and> C1 \<noteq> {} \<and> C2 \<noteq> {} \<and> separatedin X C1 C2)"
apply (simp add: separatedin_def connectedin_separation)
apply (intro conj_cong all_cong1 refl, blast)
done
lemma connectedin_eq_not_separated_subset:
"connectedin X S \<longleftrightarrow>
S \<subseteq> topspace X \<and> (\<nexists>C1 C2. S \<subseteq> C1 \<union> C2 \<and> S \<inter> C1 \<noteq> {} \<and> S \<inter> C2 \<noteq> {} \<and> separatedin X C1 C2)"
proof -
have *: "\<forall>C1 C2. S \<subseteq> C1 \<union> C2 \<longrightarrow> S \<inter> C1 = {} \<or> S \<inter> C2 = {} \<or> \<not> separatedin X C1 C2"
if "\<And>C1 C2. C1 \<union> C2 = S \<longrightarrow> C1 = {} \<or> C2 = {} \<or> \<not> separatedin X C1 C2"
proof (intro allI)
fix C1 C2
show "S \<subseteq> C1 \<union> C2 \<longrightarrow> S \<inter> C1 = {} \<or> S \<inter> C2 = {} \<or> \<not> separatedin X C1 C2"
using that [of "S \<inter> C1" "S \<inter> C2"]
by (auto simp: separatedin_mono)
qed
show ?thesis
apply (simp add: connectedin_eq_not_separated)
apply (intro conj_cong refl iffI *)
apply (blast elim!: all_forward)+
done
qed
lemma connected_space_eq_not_separated:
"connected_space X \<longleftrightarrow>
(\<nexists>C1 C2. C1 \<union> C2 = topspace X \<and> C1 \<noteq> {} \<and> C2 \<noteq> {} \<and> separatedin X C1 C2)"
by (simp add: connectedin_eq_not_separated flip: connectedin_topspace)
lemma connected_space_eq_not_separated_subset:
"connected_space X \<longleftrightarrow>
(\<nexists>C1 C2. topspace X \<subseteq> C1 \<union> C2 \<and> C1 \<noteq> {} \<and> C2 \<noteq> {} \<and> separatedin X C1 C2)"
apply (simp add: connected_space_eq_not_separated)
apply (intro all_cong1)
by (metis Un_absorb dual_order.antisym separatedin_def subset_refl sup_mono)
lemma connectedin_subset_separated_union:
"\<lbrakk>connectedin X C; separatedin X S T; C \<subseteq> S \<union> T\<rbrakk> \<Longrightarrow> C \<subseteq> S \<or> C \<subseteq> T"
unfolding connectedin_eq_not_separated_subset by blast
lemma connectedin_nonseparated_union:
"\<lbrakk>connectedin X S; connectedin X T; \<not>separatedin X S T\<rbrakk> \<Longrightarrow> connectedin X (S \<union> T)"
apply (simp add: connectedin_eq_not_separated_subset, auto)
apply (metis (no_types, hide_lams) Diff_subset_conv Diff_triv disjoint_iff_not_equal separatedin_mono sup_commute)
apply (metis (no_types, hide_lams) Diff_subset_conv Diff_triv disjoint_iff_not_equal separatedin_mono separatedin_sym sup_commute)
by (meson disjoint_iff_not_equal)
lemma connected_space_closures:
"connected_space X \<longleftrightarrow>
(\<nexists>e1 e2. e1 \<union> e2 = topspace X \<and> X closure_of e1 \<inter> X closure_of e2 = {} \<and> e1 \<noteq> {} \<and> e2 \<noteq> {})"
(is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
unfolding connected_space_closedin_eq
by (metis Un_upper1 Un_upper2 closedin_closure_of closure_of_Un closure_of_eq_empty closure_of_topspace)
next
assume ?rhs
then show ?lhs
unfolding connected_space_closedin_eq
by (metis closure_of_eq)
qed
lemma connectedin_inter_frontier_of:
assumes "connectedin X S" "S \<inter> T \<noteq> {}" "S - T \<noteq> {}"
shows "S \<inter> X frontier_of T \<noteq> {}"
proof -
have "S \<subseteq> topspace X" and *:
"\<And>E1 E2. openin X E1 \<longrightarrow> openin X E2 \<longrightarrow> E1 \<inter> E2 \<inter> S = {} \<longrightarrow> S \<subseteq> E1 \<union> E2 \<longrightarrow> E1 \<inter> S = {} \<or> E2 \<inter> S = {}"
using \<open>connectedin X S\<close> by (auto simp: connectedin)
have "S - (topspace X \<inter> T) \<noteq> {}"
using assms(3) by blast
moreover
have "S \<inter> topspace X \<inter> T \<noteq> {}"
using assms(1) assms(2) connectedin by fastforce
moreover
have False if "S \<inter> T \<noteq> {}" "S - T \<noteq> {}" "T \<subseteq> topspace X" "S \<inter> X frontier_of T = {}" for T
proof -
have null: "S \<inter> (X closure_of T - X interior_of T) = {}"
using that unfolding frontier_of_def by blast
have 1: "X interior_of T \<inter> (topspace X - X closure_of T) \<inter> S = {}"
by (metis Diff_disjoint inf_bot_left interior_of_Int interior_of_complement interior_of_empty)
have 2: "S \<subseteq> X interior_of T \<union> (topspace X - X closure_of T)"
using that \<open>S \<subseteq> topspace X\<close> null by auto
have 3: "S \<inter> X interior_of T \<noteq> {}"
using closure_of_subset that(1) that(3) null by fastforce
show ?thesis
using null \<open>S \<subseteq> topspace X\<close> that * [of "X interior_of T" "topspace X - X closure_of T"]
apply (clarsimp simp add: openin_diff 1 2)
apply (simp add: Int_commute Diff_Int_distrib 3)
by (metis Int_absorb2 contra_subsetD interior_of_subset)
qed
ultimately show ?thesis
by (metis Int_lower1 frontier_of_restrict inf_assoc)
qed
lemma connectedin_continuous_map_image:
assumes f: "continuous_map X Y f" and "connectedin X S"
shows "connectedin Y (f ` S)"
proof -
have "S \<subseteq> topspace X" and *:
"\<And>E1 E2. openin X E1 \<longrightarrow> openin X E2 \<longrightarrow> E1 \<inter> E2 \<inter> S = {} \<longrightarrow> S \<subseteq> E1 \<union> E2 \<longrightarrow> E1 \<inter> S = {} \<or> E2 \<inter> S = {}"
using \<open>connectedin X S\<close> by (auto simp: connectedin)
show ?thesis
unfolding connectedin connected_space_def
proof (intro conjI notI; clarify)
show "f x \<in> topspace Y" if "x \<in> S" for x
using \<open>S \<subseteq> topspace X\<close> continuous_map_image_subset_topspace f that by blast
next
fix U V
let ?U = "{x \<in> topspace X. f x \<in> U}"
let ?V = "{x \<in> topspace X. f x \<in> V}"
assume UV: "openin Y U" "openin Y V" "f ` S \<subseteq> U \<union> V" "U \<inter> V \<inter> f ` S = {}" "U \<inter> f ` S \<noteq> {}" "V \<inter> f ` S \<noteq> {}"
then have 1: "?U \<inter> ?V \<inter> S = {}"
by auto
have 2: "openin X ?U" "openin X ?V"
using \<open>openin Y U\<close> \<open>openin Y V\<close> continuous_map f by fastforce+
show "False"
using * [of ?U ?V] UV \<open>S \<subseteq> topspace X\<close>
by (auto simp: 1 2)
qed
qed
lemma homeomorphic_connected_space:
"X homeomorphic_space Y \<Longrightarrow> connected_space X \<longleftrightarrow> connected_space Y"
unfolding homeomorphic_space_def homeomorphic_maps_def
apply safe
apply (metis connectedin_continuous_map_image connected_space_subconnected continuous_map_image_subset_topspace image_eqI image_subset_iff)
by (metis (no_types, hide_lams) connectedin_continuous_map_image connectedin_topspace continuous_map_def continuous_map_image_subset_topspace imageI set_eq_subset subsetI)
lemma homeomorphic_map_connectedness:
assumes f: "homeomorphic_map X Y f" and U: "U \<subseteq> topspace X"
shows "connectedin Y (f ` U) \<longleftrightarrow> connectedin X U"
proof -
have 1: "f ` U \<subseteq> topspace Y \<longleftrightarrow> U \<subseteq> topspace X"
using U f homeomorphic_imp_surjective_map by blast
moreover have "connected_space (subtopology Y (f ` U)) \<longleftrightarrow> connected_space (subtopology X U)"
proof (rule homeomorphic_connected_space)
have "f ` U \<subseteq> topspace Y"
by (simp add: U 1)
then have "topspace Y \<inter> f ` U = f ` U"
by (simp add: subset_antisym)
then show "subtopology Y (f ` U) homeomorphic_space subtopology X U"
by (metis (no_types) Int_subset_iff U f homeomorphic_map_imp_homeomorphic_space homeomorphic_map_subtopologies homeomorphic_space_sym subset_antisym subset_refl)
qed
ultimately show ?thesis
by (auto simp: connectedin_def)
qed
lemma homeomorphic_map_connectedness_eq:
"homeomorphic_map X Y f
\<Longrightarrow> connectedin X U \<longleftrightarrow>
U \<subseteq> topspace X \<and> connectedin Y (f ` U)"
using homeomorphic_map_connectedness connectedin_subset_topspace by metis
lemma connectedin_discrete_topology:
"connectedin (discrete_topology U) S \<longleftrightarrow> S \<subseteq> U \<and> (\<exists>a. S \<subseteq> {a})"
proof (cases "S \<subseteq> U")
case True
show ?thesis
proof (cases "S = {}")
case False
moreover have "connectedin (discrete_topology U) S \<longleftrightarrow> (\<exists>a. S = {a})"
apply safe
using False connectedin_inter_frontier_of insert_Diff apply fastforce
using True by auto
ultimately show ?thesis
by auto
qed simp
next
case False
then show ?thesis
by (simp add: connectedin_def)
qed
lemma connected_space_discrete_topology:
"connected_space (discrete_topology U) \<longleftrightarrow> (\<exists>a. U \<subseteq> {a})"
by (metis connectedin_discrete_topology connectedin_topspace order_refl topspace_discrete_topology)
subsection\<open>Compact sets\<close>
definition compactin where
"compactin X S \<longleftrightarrow>
S \<subseteq> topspace X \<and>
(\<forall>\<U>. (\<forall>U \<in> \<U>. openin X U) \<and> S \<subseteq> \<Union>\<U>
\<longrightarrow> (\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> S \<subseteq> \<Union>\<F>))"
definition compact_space where
"compact_space X \<equiv> compactin X (topspace X)"
lemma compact_space_alt:
"compact_space X \<longleftrightarrow>
(\<forall>\<U>. (\<forall>U \<in> \<U>. openin X U) \<and> topspace X \<subseteq> \<Union>\<U>
\<longrightarrow> (\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> topspace X \<subseteq> \<Union>\<F>))"
by (simp add: compact_space_def compactin_def)
lemma compact_space:
"compact_space X \<longleftrightarrow>
(\<forall>\<U>. (\<forall>U \<in> \<U>. openin X U) \<and> \<Union>\<U> = topspace X
\<longrightarrow> (\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> \<Union>\<F> = topspace X))"
unfolding compact_space_alt
using openin_subset by fastforce
lemma compactinD:
"\<lbrakk>compactin X S; \<And>U. U \<in> \<U> \<Longrightarrow> openin X U; S \<subseteq> \<Union>\<U>\<rbrakk> \<Longrightarrow> \<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> S \<subseteq> \<Union>\<F>"
by (auto simp: compactin_def)
lemma compactin_euclidean_iff [simp]: "compactin euclidean S \<longleftrightarrow> compact S"
by (simp add: compact_eq_Heine_Borel compactin_def) meson
lemma compactin_absolute [simp]:
"compactin (subtopology X S) S \<longleftrightarrow> compactin X S"
proof -
have eq: "(\<forall>U \<in> \<U>. \<exists>Y. openin X Y \<and> U = Y \<inter> S) \<longleftrightarrow> \<U> \<subseteq> (\<lambda>Y. Y \<inter> S) ` {y. openin X y}" for \<U>
by auto
show ?thesis
by (auto simp: compactin_def openin_subtopology eq imp_conjL all_subset_image ex_finite_subset_image)
qed
lemma compactin_subspace: "compactin X S \<longleftrightarrow> S \<subseteq> topspace X \<and> compact_space (subtopology X S)"
unfolding compact_space_def topspace_subtopology
by (metis compactin_absolute compactin_def inf.absorb2)
lemma compact_space_subtopology: "compactin X S \<Longrightarrow> compact_space (subtopology X S)"
by (simp add: compactin_subspace)
lemma compactin_subtopology: "compactin (subtopology X S) T \<longleftrightarrow> compactin X T \<and> T \<subseteq> S"
apply (simp add: compactin_subspace)
by (metis inf.orderE inf_commute subtopology_subtopology)
lemma compactin_subset_topspace: "compactin X S \<Longrightarrow> S \<subseteq> topspace X"
by (simp add: compactin_subspace)
lemma compactin_contractive:
"\<lbrakk>compactin X' S; topspace X' = topspace X;
\<And>U. openin X U \<Longrightarrow> openin X' U\<rbrakk> \<Longrightarrow> compactin X S"
by (simp add: compactin_def)
lemma finite_imp_compactin:
"\<lbrakk>S \<subseteq> topspace X; finite S\<rbrakk> \<Longrightarrow> compactin X S"
by (metis compactin_subspace compact_space finite_UnionD inf.absorb_iff2 order_refl topspace_subtopology)
lemma compactin_empty [iff]: "compactin X {}"
by (simp add: finite_imp_compactin)
lemma compact_space_topspace_empty:
"topspace X = {} \<Longrightarrow> compact_space X"
by (simp add: compact_space_def)
lemma finite_imp_compactin_eq:
"finite S \<Longrightarrow> (compactin X S \<longleftrightarrow> S \<subseteq> topspace X)"
using compactin_subset_topspace finite_imp_compactin by blast
lemma compactin_sing [simp]: "compactin X {a} \<longleftrightarrow> a \<in> topspace X"
by (simp add: finite_imp_compactin_eq)
lemma closed_compactin:
assumes XK: "compactin X K" and "C \<subseteq> K" and XC: "closedin X C"
shows "compactin X C"
unfolding compactin_def
proof (intro conjI allI impI)
show "C \<subseteq> topspace X"
by (simp add: XC closedin_subset)
next
fix \<U> :: "'a set set"
assume \<U>: "Ball \<U> (openin X) \<and> C \<subseteq> \<Union>\<U>"
have "(\<forall>U\<in>insert (topspace X - C) \<U>. openin X U)"
using XC \<U> by blast
moreover have "K \<subseteq> \<Union>(insert (topspace X - C) \<U>)"
using \<U> XK compactin_subset_topspace by fastforce
ultimately obtain \<F> where "finite \<F>" "\<F> \<subseteq> insert (topspace X - C) \<U>" "K \<subseteq> \<Union>\<F>"
using assms unfolding compactin_def by metis
moreover have "openin X (topspace X - C)"
using XC by auto
ultimately show "\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> C \<subseteq> \<Union>\<F>"
using \<open>C \<subseteq> K\<close>
by (rule_tac x="\<F> - {topspace X - C}" in exI) auto
qed
lemma closedin_compact_space:
"\<lbrakk>compact_space X; closedin X S\<rbrakk> \<Longrightarrow> compactin X S"
by (simp add: closed_compactin closedin_subset compact_space_def)
lemma compact_Int_closedin:
assumes "compactin X S" "closedin X T" shows "compactin X (S \<inter> T)"
proof -
have "compactin (subtopology X S) (S \<inter> T)"
by (metis assms closedin_compact_space closedin_subtopology compactin_subspace inf_commute)
then show ?thesis
by (simp add: compactin_subtopology)
qed
lemma closed_Int_compactin: "\<lbrakk>closedin X S; compactin X T\<rbrakk> \<Longrightarrow> compactin X (S \<inter> T)"
by (metis compact_Int_closedin inf_commute)
lemma compactin_Un:
assumes S: "compactin X S" and T: "compactin X T" shows "compactin X (S \<union> T)"
unfolding compactin_def
proof (intro conjI allI impI)
show "S \<union> T \<subseteq> topspace X"
using assms by (auto simp: compactin_def)
next
fix \<U> :: "'a set set"
assume \<U>: "Ball \<U> (openin X) \<and> S \<union> T \<subseteq> \<Union>\<U>"
with S obtain \<F> where \<V>: "finite \<F>" "\<F> \<subseteq> \<U>" "S \<subseteq> \<Union>\<F>"
unfolding compactin_def by (meson sup.bounded_iff)
obtain \<W> where "finite \<W>" "\<W> \<subseteq> \<U>" "T \<subseteq> \<Union>\<W>"
using \<U> T
unfolding compactin_def by (meson sup.bounded_iff)
with \<V> show "\<exists>\<V>. finite \<V> \<and> \<V> \<subseteq> \<U> \<and> S \<union> T \<subseteq> \<Union>\<V>"
by (rule_tac x="\<F> \<union> \<W>" in exI) auto
qed
lemma compactin_Union:
"\<lbrakk>finite \<F>; \<And>S. S \<in> \<F> \<Longrightarrow> compactin X S\<rbrakk> \<Longrightarrow> compactin X (\<Union>\<F>)"
by (induction rule: finite_induct) (simp_all add: compactin_Un)
lemma compactin_subtopology_imp_compact:
assumes "compactin (subtopology X S) K" shows "compactin X K"
using assms
proof (clarsimp simp add: compactin_def)
fix \<U>
define \<V> where "\<V> \<equiv> (\<lambda>U. U \<inter> S) ` \<U>"
assume "K \<subseteq> topspace X" and "K \<subseteq> S" and "\<forall>x\<in>\<U>. openin X x" and "K \<subseteq> \<Union>\<U>"
then have "\<forall>V \<in> \<V>. openin (subtopology X S) V" "K \<subseteq> \<Union>\<V>"
unfolding \<V>_def by (auto simp: openin_subtopology)
moreover
assume "\<forall>\<U>. (\<forall>x\<in>\<U>. openin (subtopology X S) x) \<and> K \<subseteq> \<Union>\<U> \<longrightarrow> (\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> K \<subseteq> \<Union>\<F>)"
ultimately obtain \<F> where "finite \<F>" "\<F> \<subseteq> \<V>" "K \<subseteq> \<Union>\<F>"
by meson
then have \<F>: "\<exists>U. U \<in> \<U> \<and> V = U \<inter> S" if "V \<in> \<F>" for V
unfolding \<V>_def using that by blast
let ?\<F> = "(\<lambda>F. @U. U \<in> \<U> \<and> F = U \<inter> S) ` \<F>"
show "\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> K \<subseteq> \<Union>\<F>"
proof (intro exI conjI)
show "finite ?\<F>"
using \<open>finite \<F>\<close> by blast
show "?\<F> \<subseteq> \<U>"
using someI_ex [OF \<F>] by blast
show "K \<subseteq> \<Union>?\<F>"
proof clarsimp
fix x
assume "x \<in> K"
then show "\<exists>V \<in> \<F>. x \<in> (SOME U. U \<in> \<U> \<and> V = U \<inter> S)"
using \<open>K \<subseteq> \<Union>\<F>\<close> someI_ex [OF \<F>]
by (metis (no_types, lifting) IntD1 Union_iff subsetCE)
qed
qed
qed
lemma compact_imp_compactin_subtopology:
assumes "compactin X K" "K \<subseteq> S" shows "compactin (subtopology X S) K"
using assms
proof (clarsimp simp add: compactin_def)
fix \<U> :: "'a set set"
define \<V> where "\<V> \<equiv> {V. openin X V \<and> (\<exists>U \<in> \<U>. U = V \<inter> S)}"
assume "K \<subseteq> S" and "K \<subseteq> topspace X" and "\<forall>U\<in>\<U>. openin (subtopology X S) U" and "K \<subseteq> \<Union>\<U>"
then have "\<forall>V \<in> \<V>. openin X V" "K \<subseteq> \<Union>\<V>"
unfolding \<V>_def by (fastforce simp: subset_eq openin_subtopology)+
moreover
assume "\<forall>\<U>. (\<forall>U\<in>\<U>. openin X U) \<and> K \<subseteq> \<Union>\<U> \<longrightarrow> (\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> K \<subseteq> \<Union>\<F>)"
ultimately obtain \<F> where "finite \<F>" "\<F> \<subseteq> \<V>" "K \<subseteq> \<Union>\<F>"
by meson
let ?\<F> = "(\<lambda>F. F \<inter> S) ` \<F>"
show "\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> K \<subseteq> \<Union>\<F>"
proof (intro exI conjI)
show "finite ?\<F>"
using \<open>finite \<F>\<close> by blast
show "?\<F> \<subseteq> \<U>"
using \<V>_def \<open>\<F> \<subseteq> \<V>\<close> by blast
show "K \<subseteq> \<Union>?\<F>"
using \<open>K \<subseteq> \<Union>\<F>\<close> assms(2) by auto
qed
qed
proposition compact_space_fip:
"compact_space X \<longleftrightarrow>
(\<forall>\<U>. (\<forall>C\<in>\<U>. closedin X C) \<and> (\<forall>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<longrightarrow> \<Inter>\<F> \<noteq> {}) \<longrightarrow> \<Inter>\<U> \<noteq> {})"
(is "_ = ?rhs")
proof (cases "topspace X = {}")
case True
then show ?thesis
apply (clarsimp simp add: compact_space_def closedin_topspace_empty)
by (metis finite.emptyI finite_insert infinite_super insertI1 subsetI)
next
case False
show ?thesis
proof safe
fix \<U> :: "'a set set"
assume * [rule_format]: "\<forall>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<longrightarrow> \<Inter>\<F> \<noteq> {}"
define \<V> where "\<V> \<equiv> (\<lambda>S. topspace X - S) ` \<U>"
assume clo: "\<forall>C\<in>\<U>. closedin X C" and [simp]: "\<Inter>\<U> = {}"
then have "\<forall>V \<in> \<V>. openin X V" "topspace X \<subseteq> \<Union>\<V>"
by (auto simp: \<V>_def)
moreover assume [unfolded compact_space_alt, rule_format, of \<V>]: "compact_space X"
ultimately obtain \<F> where \<F>: "finite \<F>" "\<F> \<subseteq> \<U>" "topspace X \<subseteq> topspace X - \<Inter>\<F>"
by (auto simp: ex_finite_subset_image \<V>_def)
moreover have "\<F> \<noteq> {}"
using \<F> \<open>topspace X \<noteq> {}\<close> by blast
ultimately show "False"
using * [of \<F>]
by auto (metis Diff_iff Inter_iff clo closedin_def subsetD)
next
assume R [rule_format]: ?rhs
show "compact_space X"
unfolding compact_space_alt
proof clarify
fix \<U> :: "'a set set"
define \<V> where "\<V> \<equiv> (\<lambda>S. topspace X - S) ` \<U>"
assume "\<forall>C\<in>\<U>. openin X C" and "topspace X \<subseteq> \<Union>\<U>"
with \<open>topspace X \<noteq> {}\<close> have *: "\<forall>V \<in> \<V>. closedin X V" "\<U> \<noteq> {}"
by (auto simp: \<V>_def)
show "\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> topspace X \<subseteq> \<Union>\<F>"
proof (rule ccontr; simp)
assume "\<forall>\<F>\<subseteq>\<U>. finite \<F> \<longrightarrow> \<not> topspace X \<subseteq> \<Union>\<F>"
then have "\<forall>\<F>. finite \<F> \<and> \<F> \<subseteq> \<V> \<longrightarrow> \<Inter>\<F> \<noteq> {}"
by (simp add: \<V>_def all_finite_subset_image)
with \<open>topspace X \<subseteq> \<Union>\<U>\<close> show False
using R [of \<V>] * by (simp add: \<V>_def)
qed
qed
qed
qed
corollary compactin_fip:
"compactin X S \<longleftrightarrow>
S \<subseteq> topspace X \<and>
(\<forall>\<U>. (\<forall>C\<in>\<U>. closedin X C) \<and> (\<forall>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<longrightarrow> S \<inter> \<Inter>\<F> \<noteq> {}) \<longrightarrow> S \<inter> \<Inter>\<U> \<noteq> {})"
proof (cases "S = {}")
case False
show ?thesis
proof (cases "S \<subseteq> topspace X")
case True
then have "compactin X S \<longleftrightarrow>
(\<forall>\<U>. \<U> \<subseteq> (\<lambda>T. S \<inter> T) ` {T. closedin X T} \<longrightarrow>
(\<forall>\<F>. finite \<F> \<longrightarrow> \<F> \<subseteq> \<U> \<longrightarrow> \<Inter>\<F> \<noteq> {}) \<longrightarrow> \<Inter>\<U> \<noteq> {})"
by (simp add: compact_space_fip compactin_subspace closedin_subtopology image_def subset_eq Int_commute imp_conjL)
also have "\<dots> = (\<forall>\<U>\<subseteq>Collect (closedin X). (\<forall>\<F>. finite \<F> \<longrightarrow> \<F> \<subseteq> (\<inter>) S ` \<U> \<longrightarrow> \<Inter>\<F> \<noteq> {}) \<longrightarrow> \<Inter> ((\<inter>) S ` \<U>) \<noteq> {})"
by (simp add: all_subset_image)
also have "\<dots> = (\<forall>\<U>. (\<forall>C\<in>\<U>. closedin X C) \<and> (\<forall>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<longrightarrow> S \<inter> \<Inter>\<F> \<noteq> {}) \<longrightarrow> S \<inter> \<Inter>\<U> \<noteq> {})"
proof -
have eq: "((\<forall>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<longrightarrow> \<Inter> ((\<inter>) S ` \<F>) \<noteq> {}) \<longrightarrow> \<Inter> ((\<inter>) S ` \<U>) \<noteq> {}) \<longleftrightarrow>
((\<forall>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<longrightarrow> S \<inter> \<Inter>\<F> \<noteq> {}) \<longrightarrow> S \<inter> \<Inter>\<U> \<noteq> {})" for \<U>
by simp (use \<open>S \<noteq> {}\<close> in blast)
show ?thesis
apply (simp only: imp_conjL [symmetric] all_finite_subset_image eq)
apply (simp add: subset_eq)
done
qed
finally show ?thesis
using True by simp
qed (simp add: compactin_subspace)
qed force
corollary compact_space_imp_nest:
fixes C :: "nat \<Rightarrow> 'a set"
assumes "compact_space X" and clo: "\<And>n. closedin X (C n)"
and ne: "\<And>n. C n \<noteq> {}" and inc: "\<And>m n. m \<le> n \<Longrightarrow> C n \<subseteq> C m"
shows "(\<Inter>n. C n) \<noteq> {}"
proof -
let ?\<U> = "range (\<lambda>n. \<Inter>m \<le> n. C m)"
have "closedin X A" if "A \<in> ?\<U>" for A
using that clo by auto
moreover have "(\<Inter>n\<in>K. \<Inter>m \<le> n. C m) \<noteq> {}" if "finite K" for K
proof -
obtain n where "\<And>k. k \<in> K \<Longrightarrow> k \<le> n"
using Max.coboundedI \<open>finite K\<close> by blast
with inc have "C n \<subseteq> (\<Inter>n\<in>K. \<Inter>m \<le> n. C m)"
by blast
with ne [of n] show ?thesis
by blast
qed
ultimately show ?thesis
using \<open>compact_space X\<close> [unfolded compact_space_fip, rule_format, of ?\<U>]
by (simp add: all_finite_subset_image INT_extend_simps UN_atMost_UNIV del: INT_simps)
qed
lemma compactin_discrete_topology:
"compactin (discrete_topology X) S \<longleftrightarrow> S \<subseteq> X \<and> finite S" (is "?lhs = ?rhs")
proof (intro iffI conjI)
assume L: ?lhs
then show "S \<subseteq> X"
by (auto simp: compactin_def)
have *: "\<And>\<U>. Ball \<U> (openin (discrete_topology X)) \<and> S \<subseteq> \<Union>\<U> \<Longrightarrow>
(\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> S \<subseteq> \<Union>\<F>)"
using L by (auto simp: compactin_def)
show "finite S"
using * [of "(\<lambda>x. {x}) ` X"] \<open>S \<subseteq> X\<close>
by clarsimp (metis UN_singleton finite_subset_image infinite_super)
next
assume ?rhs
then show ?lhs
by (simp add: finite_imp_compactin)
qed
lemma compact_space_discrete_topology: "compact_space(discrete_topology X) \<longleftrightarrow> finite X"
by (simp add: compactin_discrete_topology compact_space_def)
lemma compact_space_imp_Bolzano_Weierstrass:
assumes "compact_space X" "infinite S" "S \<subseteq> topspace X"
shows "X derived_set_of S \<noteq> {}"
proof
assume X: "X derived_set_of S = {}"
then have "closedin X S"
by (simp add: closedin_contains_derived_set assms)
then have "compactin X S"
by (rule closedin_compact_space [OF \<open>compact_space X\<close>])
with X show False
by (metis \<open>infinite S\<close> compactin_subspace compact_space_discrete_topology inf_bot_right subtopology_eq_discrete_topology_eq)
qed
lemma compactin_imp_Bolzano_Weierstrass:
"\<lbrakk>compactin X S; infinite T \<and> T \<subseteq> S\<rbrakk> \<Longrightarrow> S \<inter> X derived_set_of T \<noteq> {}"
using compact_space_imp_Bolzano_Weierstrass [of "subtopology X S"]
by (simp add: compactin_subspace derived_set_of_subtopology inf_absorb2)
lemma compact_closure_of_imp_Bolzano_Weierstrass:
"\<lbrakk>compactin X (X closure_of S); infinite T; T \<subseteq> S; T \<subseteq> topspace X\<rbrakk> \<Longrightarrow> X derived_set_of T \<noteq> {}"
using closure_of_mono closure_of_subset compactin_imp_Bolzano_Weierstrass by fastforce
lemma discrete_compactin_eq_finite:
"S \<inter> X derived_set_of S = {} \<Longrightarrow> compactin X S \<longleftrightarrow> S \<subseteq> topspace X \<and> finite S"
apply (rule iffI)
using compactin_imp_Bolzano_Weierstrass compactin_subset_topspace apply blast
by (simp add: finite_imp_compactin_eq)
lemma discrete_compact_space_eq_finite:
"X derived_set_of (topspace X) = {} \<Longrightarrow> (compact_space X \<longleftrightarrow> finite(topspace X))"
by (metis compact_space_discrete_topology discrete_topology_unique_derived_set)
lemma image_compactin:
assumes cpt: "compactin X S" and cont: "continuous_map X Y f"
shows "compactin Y (f ` S)"
unfolding compactin_def
proof (intro conjI allI impI)
show "f ` S \<subseteq> topspace Y"
using compactin_subset_topspace cont continuous_map_image_subset_topspace cpt by blast
next
fix \<U> :: "'b set set"
assume \<U>: "Ball \<U> (openin Y) \<and> f ` S \<subseteq> \<Union>\<U>"
define \<V> where "\<V> \<equiv> (\<lambda>U. {x \<in> topspace X. f x \<in> U}) ` \<U>"
have "S \<subseteq> topspace X"
and *: "\<And>\<U>. \<lbrakk>\<forall>U\<in>\<U>. openin X U; S \<subseteq> \<Union>\<U>\<rbrakk> \<Longrightarrow> \<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> S \<subseteq> \<Union>\<F>"
using cpt by (auto simp: compactin_def)
obtain \<F> where \<F>: "finite \<F>" "\<F> \<subseteq> \<V>" "S \<subseteq> \<Union>\<F>"
proof -
have 1: "\<forall>U\<in>\<V>. openin X U"
unfolding \<V>_def using \<U> cont[unfolded continuous_map] by blast
have 2: "S \<subseteq> \<Union>\<V>"
unfolding \<V>_def using compactin_subset_topspace cpt \<U> by fastforce
show thesis
using * [OF 1 2] that by metis
qed
have "\<forall>v \<in> \<V>. \<exists>U. U \<in> \<U> \<and> v = {x \<in> topspace X. f x \<in> U}"
using \<V>_def by blast
then obtain U where U: "\<forall>v \<in> \<V>. U v \<in> \<U> \<and> v = {x \<in> topspace X. f x \<in> U v}"
by metis
show "\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> f ` S \<subseteq> \<Union>\<F>"
proof (intro conjI exI)
show "finite (U ` \<F>)"
by (simp add: \<open>finite \<F>\<close>)
next
show "U ` \<F> \<subseteq> \<U>"
using \<open>\<F> \<subseteq> \<V>\<close> U by auto
next
show "f ` S \<subseteq> \<Union> (U ` \<F>)"
using \<F>(2-3) U UnionE subset_eq U by fastforce
qed
qed
lemma homeomorphic_compact_space:
assumes "X homeomorphic_space Y"
shows "compact_space X \<longleftrightarrow> compact_space Y"
using homeomorphic_space_sym
by (metis assms compact_space_def homeomorphic_eq_everything_map homeomorphic_space image_compactin)
lemma homeomorphic_map_compactness:
assumes hom: "homeomorphic_map X Y f" and U: "U \<subseteq> topspace X"
shows "compactin Y (f ` U) \<longleftrightarrow> compactin X U"
proof -
have "f ` U \<subseteq> topspace Y"
using hom U homeomorphic_imp_surjective_map by blast
moreover have "homeomorphic_map (subtopology X U) (subtopology Y (f ` U)) f"
using U hom homeomorphic_imp_surjective_map by (blast intro: homeomorphic_map_subtopologies)
then have "compact_space (subtopology Y (f ` U)) = compact_space (subtopology X U)"
using homeomorphic_compact_space homeomorphic_map_imp_homeomorphic_space by blast
ultimately show ?thesis
by (simp add: compactin_subspace U)
qed
lemma homeomorphic_map_compactness_eq:
"homeomorphic_map X Y f
\<Longrightarrow> compactin X U \<longleftrightarrow> U \<subseteq> topspace X \<and> compactin Y (f ` U)"
by (meson compactin_subset_topspace homeomorphic_map_compactness)
subsection\<open>Embedding maps\<close>
definition embedding_map
where "embedding_map X Y f \<equiv> homeomorphic_map X (subtopology Y (f ` (topspace X))) f"
lemma embedding_map_eq:
"\<lbrakk>embedding_map X Y f; \<And>x. x \<in> topspace X \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> embedding_map X Y g"
unfolding embedding_map_def
by (metis homeomorphic_map_eq image_cong)
lemma embedding_map_compose:
assumes "embedding_map X X' f" "embedding_map X' X'' g"
shows "embedding_map X X'' (g \<circ> f)"
proof -
have hm: "homeomorphic_map X (subtopology X' (f ` topspace X)) f" "homeomorphic_map X' (subtopology X'' (g ` topspace X')) g"
using assms by (auto simp: embedding_map_def)
then obtain C where "g ` topspace X' \<inter> C = (g \<circ> f) ` topspace X"
by (metis (no_types) Int_absorb1 continuous_map_image_subset_topspace continuous_map_in_subtopology homeomorphic_eq_everything_map image_comp image_mono)
then have "homeomorphic_map (subtopology X' (f ` topspace X)) (subtopology X'' ((g \<circ> f) ` topspace X)) g"
by (metis hm homeomorphic_imp_surjective_map homeomorphic_map_subtopologies image_comp subtopology_subtopology topspace_subtopology)
then show ?thesis
unfolding embedding_map_def
using hm(1) homeomorphic_map_compose by blast
qed
lemma surjective_embedding_map:
"embedding_map X Y f \<and> f ` (topspace X) = topspace Y \<longleftrightarrow> homeomorphic_map X Y f"
by (force simp: embedding_map_def homeomorphic_eq_everything_map)
lemma embedding_map_in_subtopology:
"embedding_map X (subtopology Y S) f \<longleftrightarrow> embedding_map X Y f \<and> f ` (topspace X) \<subseteq> S"
apply (auto simp: embedding_map_def subtopology_subtopology Int_absorb1)
apply (metis (no_types) homeomorphic_imp_surjective_map subtopology_subtopology subtopology_topspace topspace_subtopology)
apply (simp add: continuous_map_def homeomorphic_eq_everything_map)
done
lemma injective_open_imp_embedding_map:
"\<lbrakk>continuous_map X Y f; open_map X Y f; inj_on f (topspace X)\<rbrakk> \<Longrightarrow> embedding_map X Y f"
unfolding embedding_map_def
apply (rule bijective_open_imp_homeomorphic_map)
using continuous_map_in_subtopology apply blast
apply (auto simp: continuous_map_in_subtopology open_map_into_subtopology continuous_map)
done
lemma injective_closed_imp_embedding_map:
"\<lbrakk>continuous_map X Y f; closed_map X Y f; inj_on f (topspace X)\<rbrakk> \<Longrightarrow> embedding_map X Y f"
unfolding embedding_map_def
apply (rule bijective_closed_imp_homeomorphic_map)
apply (simp_all add: continuous_map_into_subtopology closed_map_into_subtopology)
apply (simp add: continuous_map inf.absorb_iff2)
done
lemma embedding_map_imp_homeomorphic_space:
"embedding_map X Y f \<Longrightarrow> X homeomorphic_space (subtopology Y (f ` (topspace X)))"
unfolding embedding_map_def
using homeomorphic_space by blast
lemma embedding_imp_closed_map:
"\<lbrakk>embedding_map X Y f; closedin Y (f ` topspace X)\<rbrakk> \<Longrightarrow> closed_map X Y f"
unfolding closed_map_def
by (auto simp: closedin_closed_subtopology embedding_map_def homeomorphic_map_closedness_eq)
subsection\<open>Retraction and section maps\<close>
definition retraction_maps :: "'a topology \<Rightarrow> 'b topology \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> bool"
where "retraction_maps X Y f g \<equiv>
continuous_map X Y f \<and> continuous_map Y X g \<and> (\<forall>x \<in> topspace Y. f(g x) = x)"
definition section_map :: "'a topology \<Rightarrow> 'b topology \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
where "section_map X Y f \<equiv> \<exists>g. retraction_maps Y X g f"
definition retraction_map :: "'a topology \<Rightarrow> 'b topology \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
where "retraction_map X Y f \<equiv> \<exists>g. retraction_maps X Y f g"
lemma retraction_maps_eq:
"\<lbrakk>retraction_maps X Y f g; \<And>x. x \<in> topspace X \<Longrightarrow> f x = f' x; \<And>x. x \<in> topspace Y \<Longrightarrow> g x = g' x\<rbrakk>
\<Longrightarrow> retraction_maps X Y f' g'"
unfolding retraction_maps_def by (metis (no_types, lifting) continuous_map_def continuous_map_eq)
lemma section_map_eq:
"\<lbrakk>section_map X Y f; \<And>x. x \<in> topspace X \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> section_map X Y g"
unfolding section_map_def using retraction_maps_eq by blast
lemma retraction_map_eq:
"\<lbrakk>retraction_map X Y f; \<And>x. x \<in> topspace X \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> retraction_map X Y g"
unfolding retraction_map_def using retraction_maps_eq by blast
lemma homeomorphic_imp_retraction_maps:
"homeomorphic_maps X Y f g \<Longrightarrow> retraction_maps X Y f g"
by (simp add: homeomorphic_maps_def retraction_maps_def)
lemma section_and_retraction_eq_homeomorphic_map:
"section_map X Y f \<and> retraction_map X Y f \<longleftrightarrow> homeomorphic_map X Y f"
apply (auto simp: section_map_def retraction_map_def homeomorphic_map_maps retraction_maps_def homeomorphic_maps_def)
by (metis (full_types) continuous_map_image_subset_topspace image_subset_iff)
lemma section_imp_embedding_map:
"section_map X Y f \<Longrightarrow> embedding_map X Y f"
unfolding section_map_def embedding_map_def homeomorphic_map_maps retraction_maps_def homeomorphic_maps_def
by (force simp: continuous_map_in_subtopology continuous_map_from_subtopology)
lemma retraction_imp_quotient_map:
assumes "retraction_map X Y f"
shows "quotient_map X Y f"
unfolding quotient_map_def
proof (intro conjI subsetI allI impI)
show "f ` topspace X = topspace Y"
using assms by (force simp: retraction_map_def retraction_maps_def continuous_map_def)
next
fix U
assume U: "U \<subseteq> topspace Y"
have "openin Y U"
if "\<forall>x\<in>topspace Y. g x \<in> topspace X" "\<forall>x\<in>topspace Y. f (g x) = x"
"openin Y {x \<in> topspace Y. g x \<in> {x \<in> topspace X. f x \<in> U}}" for g
using openin_subopen U that by fastforce
then show "openin X {x \<in> topspace X. f x \<in> U} = openin Y U"
using assms by (auto simp: retraction_map_def retraction_maps_def continuous_map_def)
qed
lemma retraction_maps_compose:
"\<lbrakk>retraction_maps X Y f f'; retraction_maps Y Z g g'\<rbrakk> \<Longrightarrow> retraction_maps X Z (g \<circ> f) (f' \<circ> g')"
by (clarsimp simp: retraction_maps_def continuous_map_compose) (simp add: continuous_map_def)
lemma retraction_map_compose:
"\<lbrakk>retraction_map X Y f; retraction_map Y Z g\<rbrakk> \<Longrightarrow> retraction_map X Z (g \<circ> f)"
by (meson retraction_map_def retraction_maps_compose)
lemma section_map_compose:
"\<lbrakk>section_map X Y f; section_map Y Z g\<rbrakk> \<Longrightarrow> section_map X Z (g \<circ> f)"
by (meson retraction_maps_compose section_map_def)
lemma surjective_section_eq_homeomorphic_map:
"section_map X Y f \<and> f ` (topspace X) = topspace Y \<longleftrightarrow> homeomorphic_map X Y f"
by (meson section_and_retraction_eq_homeomorphic_map section_imp_embedding_map surjective_embedding_map)
lemma surjective_retraction_or_section_map:
"f ` (topspace X) = topspace Y \<Longrightarrow> retraction_map X Y f \<or> section_map X Y f \<longleftrightarrow> retraction_map X Y f"
using section_and_retraction_eq_homeomorphic_map surjective_section_eq_homeomorphic_map by fastforce
lemma retraction_imp_surjective_map:
"retraction_map X Y f \<Longrightarrow> f ` (topspace X) = topspace Y"
by (simp add: retraction_imp_quotient_map quotient_imp_surjective_map)
lemma section_imp_injective_map:
"\<lbrakk>section_map X Y f; x \<in> topspace X; y \<in> topspace X\<rbrakk> \<Longrightarrow> f x = f y \<longleftrightarrow> x = y"
by (metis (mono_tags, hide_lams) retraction_maps_def section_map_def)
lemma retraction_maps_to_retract_maps:
"retraction_maps X Y r s
\<Longrightarrow> retraction_maps X (subtopology X (s ` (topspace Y))) (s \<circ> r) id"
unfolding retraction_maps_def
by (auto simp: continuous_map_compose continuous_map_into_subtopology continuous_map_from_subtopology)
subsection \<open>Continuity\<close>
lemma continuous_on_open:
"continuous_on S f \<longleftrightarrow>
(\<forall>T. openin (top_of_set (f ` S)) T \<longrightarrow>
openin (top_of_set S) (S \<inter> f -` T))"
unfolding continuous_on_open_invariant openin_open Int_def vimage_def Int_commute
by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong)
lemma continuous_on_closed:
"continuous_on S f \<longleftrightarrow>
(\<forall>T. closedin (top_of_set (f ` S)) T \<longrightarrow>
closedin (top_of_set S) (S \<inter> f -` T))"
unfolding continuous_on_closed_invariant closedin_closed Int_def vimage_def Int_commute
by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong)
lemma continuous_on_imp_closedin:
assumes "continuous_on S f" "closedin (top_of_set (f ` S)) T"
shows "closedin (top_of_set S) (S \<inter> f -` T)"
using assms continuous_on_closed by blast
lemma continuous_map_subtopology_eu [simp]:
"continuous_map (top_of_set S) (subtopology euclidean T) h \<longleftrightarrow> continuous_on S h \<and> h ` S \<subseteq> T"
apply safe
apply (metis continuous_map_closedin_preimage_eq continuous_map_in_subtopology continuous_on_closed order_refl topspace_euclidean_subtopology)
apply (simp add: continuous_map_closedin_preimage_eq image_subset_iff)
by (metis (no_types, hide_lams) continuous_map_closedin_preimage_eq continuous_map_in_subtopology continuous_on_closed order_refl topspace_euclidean_subtopology)
lemma continuous_map_euclidean_top_of_set:
assumes eq: "f -` S = UNIV" and cont: "continuous_on UNIV f"
shows "continuous_map euclidean (top_of_set S) f"
by (simp add: cont continuous_map_into_subtopology eq image_subset_iff_subset_vimage)
subsection\<^marker>\<open>tag unimportant\<close> \<open>Half-global and completely global cases\<close>
lemma continuous_openin_preimage_gen:
assumes "continuous_on S f" "open T"
shows "openin (top_of_set S) (S \<inter> f -` T)"
proof -
have *: "(S \<inter> f -` T) = (S \<inter> f -` (T \<inter> f ` S))"
by auto
have "openin (top_of_set (f ` S)) (T \<inter> f ` S)"
using openin_open_Int[of T "f ` S", OF assms(2)] unfolding openin_open by auto
then show ?thesis
using assms(1)[unfolded continuous_on_open, THEN spec[where x="T \<inter> f ` S"]]
using * by auto
qed
lemma continuous_closedin_preimage:
assumes "continuous_on S f" and "closed T"
shows "closedin (top_of_set S) (S \<inter> f -` T)"
proof -
have *: "(S \<inter> f -` T) = (S \<inter> f -` (T \<inter> f ` S))"
by auto
have "closedin (top_of_set (f ` S)) (T \<inter> f ` S)"
using closedin_closed_Int[of T "f ` S", OF assms(2)]
by (simp add: Int_commute)
then show ?thesis
using assms(1)[unfolded continuous_on_closed, THEN spec[where x="T \<inter> f ` S"]]
using * by auto
qed
lemma continuous_openin_preimage_eq:
"continuous_on S f \<longleftrightarrow>
(\<forall>T. open T \<longrightarrow> openin (top_of_set S) (S \<inter> f -` T))"
apply safe
apply (simp add: continuous_openin_preimage_gen)
apply (fastforce simp add: continuous_on_open openin_open)
done
lemma continuous_closedin_preimage_eq:
"continuous_on S f \<longleftrightarrow>
(\<forall>T. closed T \<longrightarrow> closedin (top_of_set S) (S \<inter> f -` T))"
apply safe
apply (simp add: continuous_closedin_preimage)
apply (fastforce simp add: continuous_on_closed closedin_closed)
done
lemma continuous_open_preimage:
assumes contf: "continuous_on S f" and "open S" "open T"
shows "open (S \<inter> f -` T)"
proof-
obtain U where "open U" "(S \<inter> f -` T) = S \<inter> U"
using continuous_openin_preimage_gen[OF contf \<open>open T\<close>]
unfolding openin_open by auto
then show ?thesis
using open_Int[of S U, OF \<open>open S\<close>] by auto
qed
lemma continuous_closed_preimage:
assumes contf: "continuous_on S f" and "closed S" "closed T"
shows "closed (S \<inter> f -` T)"
proof-
obtain U where "closed U" "(S \<inter> f -` T) = S \<inter> U"
using continuous_closedin_preimage[OF contf \<open>closed T\<close>]
unfolding closedin_closed by auto
then show ?thesis using closed_Int[of S U, OF \<open>closed S\<close>] by auto
qed
lemma continuous_open_vimage: "open S \<Longrightarrow> (\<And>x. continuous (at x) f) \<Longrightarrow> open (f -` S)"
by (metis continuous_on_eq_continuous_within open_vimage)
lemma continuous_closed_vimage: "closed S \<Longrightarrow> (\<And>x. continuous (at x) f) \<Longrightarrow> closed (f -` S)"
by (simp add: closed_vimage continuous_on_eq_continuous_within)
lemma Times_in_interior_subtopology:
assumes "(x, y) \<in> U" "openin (top_of_set (S \<times> T)) U"
obtains V W where "openin (top_of_set S) V" "x \<in> V"
"openin (top_of_set T) W" "y \<in> W" "(V \<times> W) \<subseteq> U"
proof -
from assms obtain E where "open E" "U = S \<times> T \<inter> E" "(x, y) \<in> E" "x \<in> S" "y \<in> T"
by (auto simp: openin_open)
from open_prod_elim[OF \<open>open E\<close> \<open>(x, y) \<in> E\<close>]
obtain E1 E2 where "open E1" "open E2" "(x, y) \<in> E1 \<times> E2" "E1 \<times> E2 \<subseteq> E"
by blast
show ?thesis
proof
show "openin (top_of_set S) (E1 \<inter> S)"
"openin (top_of_set T) (E2 \<inter> T)"
using \<open>open E1\<close> \<open>open E2\<close>
by (auto simp: openin_open)
show "x \<in> E1 \<inter> S" "y \<in> E2 \<inter> T"
using \<open>(x, y) \<in> E1 \<times> E2\<close> \<open>x \<in> S\<close> \<open>y \<in> T\<close> by auto
show "(E1 \<inter> S) \<times> (E2 \<inter> T) \<subseteq> U"
using \<open>E1 \<times> E2 \<subseteq> E\<close> \<open>U = _\<close>
by (auto simp: )
qed
qed
lemma closedin_Times:
"closedin (top_of_set S) S' \<Longrightarrow> closedin (top_of_set T) T' \<Longrightarrow>
closedin (top_of_set (S \<times> T)) (S' \<times> T')"
unfolding closedin_closed using closed_Times by blast
lemma openin_Times:
"openin (top_of_set S) S' \<Longrightarrow> openin (top_of_set T) T' \<Longrightarrow>
openin (top_of_set (S \<times> T)) (S' \<times> T')"
unfolding openin_open using open_Times by blast
lemma openin_Times_eq:
fixes S :: "'a::topological_space set" and T :: "'b::topological_space set"
shows
"openin (top_of_set (S \<times> T)) (S' \<times> T') \<longleftrightarrow>
S' = {} \<or> T' = {} \<or> openin (top_of_set S) S' \<and> openin (top_of_set T) T'"
(is "?lhs = ?rhs")
proof (cases "S' = {} \<or> T' = {}")
case True
then show ?thesis by auto
next
case False
then obtain x y where "x \<in> S'" "y \<in> T'"
by blast
show ?thesis
proof
assume ?lhs
have "openin (top_of_set S) S'"
apply (subst openin_subopen, clarify)
apply (rule Times_in_interior_subtopology [OF _ \<open>?lhs\<close>])
using \<open>y \<in> T'\<close>
apply auto
done
moreover have "openin (top_of_set T) T'"
apply (subst openin_subopen, clarify)
apply (rule Times_in_interior_subtopology [OF _ \<open>?lhs\<close>])
using \<open>x \<in> S'\<close>
apply auto
done
ultimately show ?rhs
by simp
next
assume ?rhs
with False show ?lhs
by (simp add: openin_Times)
qed
qed
lemma Lim_transform_within_openin:
assumes f: "(f \<longlongrightarrow> l) (at a within T)"
and "openin (top_of_set T) S" "a \<in> S"
and eq: "\<And>x. \<lbrakk>x \<in> S; x \<noteq> a\<rbrakk> \<Longrightarrow> f x = g x"
shows "(g \<longlongrightarrow> l) (at a within T)"
proof -
have "\<forall>\<^sub>F x in at a within T. x \<in> T \<and> x \<noteq> a"
by (simp add: eventually_at_filter)
moreover
from \<open>openin _ _\<close> obtain U where "open U" "S = T \<inter> U"
by (auto simp: openin_open)
then have "a \<in> U" using \<open>a \<in> S\<close> by auto
from topological_tendstoD[OF tendsto_ident_at \<open>open U\<close> \<open>a \<in> U\<close>]
have "\<forall>\<^sub>F x in at a within T. x \<in> U" by auto
ultimately
have "\<forall>\<^sub>F x in at a within T. f x = g x"
by eventually_elim (auto simp: \<open>S = _\<close> eq)
with f show ?thesis
by (rule Lim_transform_eventually)
qed
lemma continuous_on_open_gen:
assumes "f ` S \<subseteq> T"
shows "continuous_on S f \<longleftrightarrow>
(\<forall>U. openin (top_of_set T) U
\<longrightarrow> openin (top_of_set S) (S \<inter> f -` U))"
(is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
by (clarsimp simp add: continuous_openin_preimage_eq openin_open)
(metis Int_assoc assms image_subset_iff_subset_vimage inf.absorb_iff1)
next
assume R [rule_format]: ?rhs
show ?lhs
proof (clarsimp simp add: continuous_openin_preimage_eq)
fix U::"'a set"
assume "open U"
then have "openin (top_of_set S) (S \<inter> f -` (U \<inter> T))"
by (metis R inf_commute openin_open)
then show "openin (top_of_set S) (S \<inter> f -` U)"
by (metis Int_assoc Int_commute assms image_subset_iff_subset_vimage inf.absorb_iff2 vimage_Int)
qed
qed
lemma continuous_openin_preimage:
"\<lbrakk>continuous_on S f; f ` S \<subseteq> T; openin (top_of_set T) U\<rbrakk>
\<Longrightarrow> openin (top_of_set S) (S \<inter> f -` U)"
by (simp add: continuous_on_open_gen)
lemma continuous_on_closed_gen:
assumes "f ` S \<subseteq> T"
shows "continuous_on S f \<longleftrightarrow>
(\<forall>U. closedin (top_of_set T) U
\<longrightarrow> closedin (top_of_set S) (S \<inter> f -` U))"
(is "?lhs = ?rhs")
proof -
have *: "U \<subseteq> T \<Longrightarrow> S \<inter> f -` (T - U) = S - (S \<inter> f -` U)" for U
using assms by blast
show ?thesis
proof
assume L: ?lhs
show ?rhs
proof clarify
fix U
assume "closedin (top_of_set T) U"
then show "closedin (top_of_set S) (S \<inter> f -` U)"
using L unfolding continuous_on_open_gen [OF assms]
by (metis * closedin_def inf_le1 topspace_euclidean_subtopology)
qed
next
assume R [rule_format]: ?rhs
show ?lhs
unfolding continuous_on_open_gen [OF assms]
by (metis * R inf_le1 openin_closedin_eq topspace_euclidean_subtopology)
qed
qed
lemma continuous_closedin_preimage_gen:
assumes "continuous_on S f" "f ` S \<subseteq> T" "closedin (top_of_set T) U"
shows "closedin (top_of_set S) (S \<inter> f -` U)"
using assms continuous_on_closed_gen by blast
lemma continuous_transform_within_openin:
assumes "continuous (at a within T) f"
and "openin (top_of_set T) S" "a \<in> S"
and eq: "\<And>x. x \<in> S \<Longrightarrow> f x = g x"
shows "continuous (at a within T) g"
using assms by (simp add: Lim_transform_within_openin continuous_within)
subsection\<^marker>\<open>tag important\<close> \<open>The topology generated by some (open) subsets\<close>
text \<open>In the definition below of a generated topology, the \<open>Empty\<close> case is not necessary,
as it follows from \<open>UN\<close> taking for \<open>K\<close> the empty set. However, it is convenient to have,
and is never a problem in proofs, so I prefer to write it down explicitly.
We do not require \<open>UNIV\<close> to be an open set, as this will not be the case in applications. (We are
thinking of a topology on a subset of \<open>UNIV\<close>, the remaining part of \<open>UNIV\<close> being irrelevant.)\<close>
inductive generate_topology_on for S where
Empty: "generate_topology_on S {}"
| Int: "generate_topology_on S a \<Longrightarrow> generate_topology_on S b \<Longrightarrow> generate_topology_on S (a \<inter> b)"
| UN: "(\<And>k. k \<in> K \<Longrightarrow> generate_topology_on S k) \<Longrightarrow> generate_topology_on S (\<Union>K)"
| Basis: "s \<in> S \<Longrightarrow> generate_topology_on S s"
lemma istopology_generate_topology_on:
"istopology (generate_topology_on S)"
unfolding istopology_def by (auto intro: generate_topology_on.intros)
text \<open>The basic property of the topology generated by a set \<open>S\<close> is that it is the
smallest topology containing all the elements of \<open>S\<close>:\<close>
lemma generate_topology_on_coarsest:
assumes "istopology T"
"\<And>s. s \<in> S \<Longrightarrow> T s"
"generate_topology_on S s0"
shows "T s0"
using assms(3) apply (induct rule: generate_topology_on.induct)
using assms(1) assms(2) unfolding istopology_def by auto
abbreviation\<^marker>\<open>tag unimportant\<close> topology_generated_by::"('a set set) \<Rightarrow> ('a topology)"
where "topology_generated_by S \<equiv> topology (generate_topology_on S)"
lemma openin_topology_generated_by_iff:
"openin (topology_generated_by S) s \<longleftrightarrow> generate_topology_on S s"
using topology_inverse'[OF istopology_generate_topology_on[of S]] by simp
lemma openin_topology_generated_by:
"openin (topology_generated_by S) s \<Longrightarrow> generate_topology_on S s"
using openin_topology_generated_by_iff by auto
lemma topology_generated_by_topspace [simp]:
"topspace (topology_generated_by S) = (\<Union>S)"
proof
{
fix s assume "openin (topology_generated_by S) s"
then have "generate_topology_on S s" by (rule openin_topology_generated_by)
then have "s \<subseteq> (\<Union>S)" by (induct, auto)
}
then show "topspace (topology_generated_by S) \<subseteq> (\<Union>S)"
unfolding topspace_def by auto
next
have "generate_topology_on S (\<Union>S)"
using generate_topology_on.UN[OF generate_topology_on.Basis, of S S] by simp
then show "(\<Union>S) \<subseteq> topspace (topology_generated_by S)"
unfolding topspace_def using openin_topology_generated_by_iff by auto
qed
lemma topology_generated_by_Basis:
"s \<in> S \<Longrightarrow> openin (topology_generated_by S) s"
by (simp only: openin_topology_generated_by_iff, auto simp: generate_topology_on.Basis)
lemma generate_topology_on_Inter:
"\<lbrakk>finite \<F>; \<And>K. K \<in> \<F> \<Longrightarrow> generate_topology_on \<S> K; \<F> \<noteq> {}\<rbrakk> \<Longrightarrow> generate_topology_on \<S> (\<Inter>\<F>)"
by (induction \<F> rule: finite_induct; force intro: generate_topology_on.intros)
subsection\<open>Topology bases and sub-bases\<close>
lemma istopology_base_alt:
"istopology (arbitrary union_of P) \<longleftrightarrow>
(\<forall>S T. (arbitrary union_of P) S \<and> (arbitrary union_of P) T
\<longrightarrow> (arbitrary union_of P) (S \<inter> T))"
by (simp add: istopology_def) (blast intro: arbitrary_union_of_Union)
lemma istopology_base_eq:
"istopology (arbitrary union_of P) \<longleftrightarrow>
(\<forall>S T. P S \<and> P T \<longrightarrow> (arbitrary union_of P) (S \<inter> T))"
by (simp add: istopology_base_alt arbitrary_union_of_Int_eq)
lemma istopology_base:
"(\<And>S T. \<lbrakk>P S; P T\<rbrakk> \<Longrightarrow> P(S \<inter> T)) \<Longrightarrow> istopology (arbitrary union_of P)"
by (simp add: arbitrary_def istopology_base_eq union_of_inc)
lemma openin_topology_base_unique:
"openin X = arbitrary union_of P \<longleftrightarrow>
(\<forall>V. P V \<longrightarrow> openin X V) \<and> (\<forall>U x. openin X U \<and> x \<in> U \<longrightarrow> (\<exists>V. P V \<and> x \<in> V \<and> V \<subseteq> U))"
(is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
by (auto simp: union_of_def arbitrary_def)
next
assume R: ?rhs
then have *: "\<exists>\<U>\<subseteq>Collect P. \<Union>\<U> = S" if "openin X S" for S
using that by (rule_tac x="{V. P V \<and> V \<subseteq> S}" in exI) fastforce
from R show ?lhs
by (fastforce simp add: union_of_def arbitrary_def intro: *)
qed
lemma topology_base_unique:
"\<lbrakk>\<And>S. P S \<Longrightarrow> openin X S;
\<And>U x. \<lbrakk>openin X U; x \<in> U\<rbrakk> \<Longrightarrow> \<exists>B. P B \<and> x \<in> B \<and> B \<subseteq> U\<rbrakk>
\<Longrightarrow> topology(arbitrary union_of P) = X"
by (metis openin_topology_base_unique openin_inverse [of X])
lemma topology_bases_eq_aux:
"\<lbrakk>(arbitrary union_of P) S;
\<And>U x. \<lbrakk>P U; x \<in> U\<rbrakk> \<Longrightarrow> \<exists>V. Q V \<and> x \<in> V \<and> V \<subseteq> U\<rbrakk>
\<Longrightarrow> (arbitrary union_of Q) S"
by (metis arbitrary_union_of_alt arbitrary_union_of_idempot)
lemma topology_bases_eq:
"\<lbrakk>\<And>U x. \<lbrakk>P U; x \<in> U\<rbrakk> \<Longrightarrow> \<exists>V. Q V \<and> x \<in> V \<and> V \<subseteq> U;
\<And>V x. \<lbrakk>Q V; x \<in> V\<rbrakk> \<Longrightarrow> \<exists>U. P U \<and> x \<in> U \<and> U \<subseteq> V\<rbrakk>
\<Longrightarrow> topology (arbitrary union_of P) =
topology (arbitrary union_of Q)"
by (fastforce intro: arg_cong [where f=topology] elim: topology_bases_eq_aux)
lemma istopology_subbase:
"istopology (arbitrary union_of (finite intersection_of P relative_to S))"
by (simp add: finite_intersection_of_Int istopology_base relative_to_Int)
lemma openin_subbase:
"openin (topology (arbitrary union_of (finite intersection_of B relative_to U))) S
\<longleftrightarrow> (arbitrary union_of (finite intersection_of B relative_to U)) S"
by (simp add: istopology_subbase topology_inverse')
lemma topspace_subbase [simp]:
"topspace(topology (arbitrary union_of (finite intersection_of B relative_to U))) = U" (is "?lhs = _")
proof
show "?lhs \<subseteq> U"
by (metis arbitrary_union_of_relative_to openin_subbase openin_topspace relative_to_imp_subset)
show "U \<subseteq> ?lhs"
by (metis arbitrary_union_of_inc finite_intersection_of_empty inf.orderE istopology_subbase
openin_subset relative_to_inc subset_UNIV topology_inverse')
qed
lemma minimal_topology_subbase:
"\<lbrakk>\<And>S. P S \<Longrightarrow> openin X S; openin X U;
openin(topology(arbitrary union_of (finite intersection_of P relative_to U))) S\<rbrakk>
\<Longrightarrow> openin X S"
apply (simp add: istopology_subbase topology_inverse)
apply (simp add: union_of_def intersection_of_def relative_to_def)
apply (blast intro: openin_Int_Inter)
done
lemma istopology_subbase_UNIV:
"istopology (arbitrary union_of (finite intersection_of P))"
by (simp add: istopology_base finite_intersection_of_Int)
lemma generate_topology_on_eq:
"generate_topology_on S = arbitrary union_of finite' intersection_of (\<lambda>x. x \<in> S)" (is "?lhs = ?rhs")
proof (intro ext iffI)
fix A
assume "?lhs A"
then show "?rhs A"
proof induction
case (Int a b)
then show ?case
by (metis (mono_tags, lifting) istopology_base_alt finite'_intersection_of_Int istopology_base)
next
case (UN K)
then show ?case
by (simp add: arbitrary_union_of_Union)
next
case (Basis s)
then show ?case
by (simp add: Sup_upper arbitrary_union_of_inc finite'_intersection_of_inc relative_to_subset)
qed auto
next
fix A
assume "?rhs A"
then obtain \<U> where \<U>: "\<And>T. T \<in> \<U> \<Longrightarrow> \<exists>\<F>. finite' \<F> \<and> \<F> \<subseteq> S \<and> \<Inter>\<F> = T" and eq: "A = \<Union>\<U>"
unfolding union_of_def intersection_of_def by auto
show "?lhs A"
unfolding eq
proof (rule generate_topology_on.UN)
fix T
assume "T \<in> \<U>"
with \<U> obtain \<F> where "finite' \<F>" "\<F> \<subseteq> S" "\<Inter>\<F> = T"
by blast
have "generate_topology_on S (\<Inter>\<F>)"
proof (rule generate_topology_on_Inter)
show "finite \<F>" "\<F> \<noteq> {}"
by (auto simp: \<open>finite' \<F>\<close>)
show "\<And>K. K \<in> \<F> \<Longrightarrow> generate_topology_on S K"
by (metis \<open>\<F> \<subseteq> S\<close> generate_topology_on.simps subset_iff)
qed
then show "generate_topology_on S T"
using \<open>\<Inter>\<F> = T\<close> by blast
qed
qed
lemma continuous_on_generated_topo_iff:
"continuous_map T1 (topology_generated_by S) f \<longleftrightarrow>
((\<forall>U. U \<in> S \<longrightarrow> openin T1 (f-`U \<inter> topspace(T1))) \<and> (f`(topspace T1) \<subseteq> (\<Union> S)))"
unfolding continuous_map_alt topology_generated_by_topspace
proof (auto simp add: topology_generated_by_Basis)
assume H: "\<forall>U. U \<in> S \<longrightarrow> openin T1 (f -` U \<inter> topspace T1)"
fix U assume "openin (topology_generated_by S) U"
then have "generate_topology_on S U" by (rule openin_topology_generated_by)
then show "openin T1 (f -` U \<inter> topspace T1)"
proof (induct)
fix a b
assume H: "openin T1 (f -` a \<inter> topspace T1)" "openin T1 (f -` b \<inter> topspace T1)"
have "f -` (a \<inter> b) \<inter> topspace T1 = (f-`a \<inter> topspace T1) \<inter> (f-`b \<inter> topspace T1)"
by auto
then show "openin T1 (f -` (a \<inter> b) \<inter> topspace T1)" using H by auto
next
fix K
assume H: "openin T1 (f -` k \<inter> topspace T1)" if "k\<in> K" for k
define L where "L = {f -` k \<inter> topspace T1|k. k \<in> K}"
have *: "openin T1 l" if "l \<in>L" for l using that H unfolding L_def by auto
have "openin T1 (\<Union>L)" using openin_Union[OF *] by simp
moreover have "(\<Union>L) = (f -` \<Union>K \<inter> topspace T1)" unfolding L_def by auto
ultimately show "openin T1 (f -` \<Union>K \<inter> topspace T1)" by simp
qed (auto simp add: H)
qed
lemma continuous_on_generated_topo:
assumes "\<And>U. U \<in>S \<Longrightarrow> openin T1 (f-`U \<inter> topspace(T1))"
"f`(topspace T1) \<subseteq> (\<Union> S)"
shows "continuous_map T1 (topology_generated_by S) f"
using assms continuous_on_generated_topo_iff by blast
subsection\<^marker>\<open>tag important\<close> \<open>Pullback topology\<close>
text \<open>Pulling back a topology by map gives again a topology. \<open>subtopology\<close> is
a special case of this notion, pulling back by the identity. We introduce the general notion as
we will need it to define the strong operator topology on the space of continuous linear operators,
by pulling back the product topology on the space of all functions.\<close>
text \<open>\<open>pullback_topology A f T\<close> is the pullback of the topology \<open>T\<close> by the map \<open>f\<close> on
the set \<open>A\<close>.\<close>
definition\<^marker>\<open>tag important\<close> pullback_topology::"('a set) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('b topology) \<Rightarrow> ('a topology)"
where "pullback_topology A f T = topology (\<lambda>S. \<exists>U. openin T U \<and> S = f-`U \<inter> A)"
lemma istopology_pullback_topology:
"istopology (\<lambda>S. \<exists>U. openin T U \<and> S = f-`U \<inter> A)"
unfolding istopology_def proof (auto)
fix K assume "\<forall>S\<in>K. \<exists>U. openin T U \<and> S = f -` U \<inter> A"
then have "\<exists>U. \<forall>S\<in>K. openin T (U S) \<and> S = f-`(U S) \<inter> A"
by (rule bchoice)
then obtain U where U: "\<forall>S\<in>K. openin T (U S) \<and> S = f-`(U S) \<inter> A"
by blast
define V where "V = (\<Union>S\<in>K. U S)"
have "openin T V" "\<Union>K = f -` V \<inter> A" unfolding V_def using U by auto
then show "\<exists>V. openin T V \<and> \<Union>K = f -` V \<inter> A" by auto
qed
lemma openin_pullback_topology:
"openin (pullback_topology A f T) S \<longleftrightarrow> (\<exists>U. openin T U \<and> S = f-`U \<inter> A)"
unfolding pullback_topology_def topology_inverse'[OF istopology_pullback_topology] by auto
lemma topspace_pullback_topology:
"topspace (pullback_topology A f T) = f-`(topspace T) \<inter> A"
by (auto simp add: topspace_def openin_pullback_topology)
proposition continuous_map_pullback [intro]:
assumes "continuous_map T1 T2 g"
shows "continuous_map (pullback_topology A f T1) T2 (g o f)"
unfolding continuous_map_alt
proof (auto)
fix U::"'b set" assume "openin T2 U"
then have "openin T1 (g-`U \<inter> topspace T1)"
using assms unfolding continuous_map_alt by auto
have "(g o f)-`U \<inter> topspace (pullback_topology A f T1) = (g o f)-`U \<inter> A \<inter> f-`(topspace T1)"
unfolding topspace_pullback_topology by auto
also have "... = f-`(g-`U \<inter> topspace T1) \<inter> A "
by auto
also have "openin (pullback_topology A f T1) (...)"
unfolding openin_pullback_topology using \<open>openin T1 (g-`U \<inter> topspace T1)\<close> by auto
finally show "openin (pullback_topology A f T1) ((g \<circ> f) -` U \<inter> topspace (pullback_topology A f T1))"
by auto
next
fix x assume "x \<in> topspace (pullback_topology A f T1)"
then have "f x \<in> topspace T1"
unfolding topspace_pullback_topology by auto
then show "g (f x) \<in> topspace T2"
using assms unfolding continuous_map_def by auto
qed
proposition continuous_map_pullback' [intro]:
assumes "continuous_map T1 T2 (f o g)" "topspace T1 \<subseteq> g-`A"
shows "continuous_map T1 (pullback_topology A f T2) g"
unfolding continuous_map_alt
proof (auto)
fix U assume "openin (pullback_topology A f T2) U"
then have "\<exists>V. openin T2 V \<and> U = f-`V \<inter> A"
unfolding openin_pullback_topology by auto
then obtain V where "openin T2 V" "U = f-`V \<inter> A"
by blast
then have "g -` U \<inter> topspace T1 = g-`(f-`V \<inter> A) \<inter> topspace T1"
by blast
also have "... = (f o g)-`V \<inter> (g-`A \<inter> topspace T1)"
by auto
also have "... = (f o g)-`V \<inter> topspace T1"
using assms(2) by auto
also have "openin T1 (...)"
using assms(1) \<open>openin T2 V\<close> by auto
finally show "openin T1 (g -` U \<inter> topspace T1)" by simp
next
fix x assume "x \<in> topspace T1"
have "(f o g) x \<in> topspace T2"
using assms(1) \<open>x \<in> topspace T1\<close> unfolding continuous_map_def by auto
then have "g x \<in> f-`(topspace T2)"
unfolding comp_def by blast
moreover have "g x \<in> A" using assms(2) \<open>x \<in> topspace T1\<close> by blast
ultimately show "g x \<in> topspace (pullback_topology A f T2)"
unfolding topspace_pullback_topology by blast
qed
subsection\<open>Proper maps (not a priori assumed continuous) \<close>
definition proper_map
where
"proper_map X Y f \<equiv>
closed_map X Y f \<and> (\<forall>y \<in> topspace Y. compactin X {x \<in> topspace X. f x = y})"
lemma proper_imp_closed_map:
"proper_map X Y f \<Longrightarrow> closed_map X Y f"
by (simp add: proper_map_def)
lemma proper_map_imp_subset_topspace:
"proper_map X Y f \<Longrightarrow> f ` (topspace X) \<subseteq> topspace Y"
by (simp add: closed_map_imp_subset_topspace proper_map_def)
lemma closed_injective_imp_proper_map:
assumes f: "closed_map X Y f" and inj: "inj_on f (topspace X)"
shows "proper_map X Y f"
unfolding proper_map_def
proof (clarsimp simp: f)
show "compactin X {x \<in> topspace X. f x = y}"
if "y \<in> topspace Y" for y
proof -
have "{x \<in> topspace X. f x = y} = {} \<or> (\<exists>a \<in> topspace X. {x \<in> topspace X. f x = y} = {a})"
using inj_on_eq_iff [OF inj] by auto
then show ?thesis
using that by (metis (no_types, lifting) compactin_empty compactin_sing)
qed
qed
lemma injective_imp_proper_eq_closed_map:
"inj_on f (topspace X) \<Longrightarrow> (proper_map X Y f \<longleftrightarrow> closed_map X Y f)"
using closed_injective_imp_proper_map proper_imp_closed_map by blast
lemma homeomorphic_imp_proper_map:
"homeomorphic_map X Y f \<Longrightarrow> proper_map X Y f"
by (simp add: closed_injective_imp_proper_map homeomorphic_eq_everything_map)
lemma compactin_proper_map_preimage:
assumes f: "proper_map X Y f" and "compactin Y K"
shows "compactin X {x. x \<in> topspace X \<and> f x \<in> K}"
proof -
have "f ` (topspace X) \<subseteq> topspace Y"
by (simp add: f proper_map_imp_subset_topspace)
have *: "\<And>y. y \<in> topspace Y \<Longrightarrow> compactin X {x \<in> topspace X. f x = y}"
using f by (auto simp: proper_map_def)
show ?thesis
unfolding compactin_def
proof clarsimp
show "\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> {x \<in> topspace X. f x \<in> K} \<subseteq> \<Union>\<F>"
if \<U>: "\<forall>U\<in>\<U>. openin X U" and sub: "{x \<in> topspace X. f x \<in> K} \<subseteq> \<Union>\<U>"
for \<U>
proof -
have "\<forall>y \<in> K. \<exists>\<V>. finite \<V> \<and> \<V> \<subseteq> \<U> \<and> {x \<in> topspace X. f x = y} \<subseteq> \<Union>\<V>"
proof
fix y
assume "y \<in> K"
then have "compactin X {x \<in> topspace X. f x = y}"
by (metis "*" \<open>compactin Y K\<close> compactin_subspace subsetD)
with \<open>y \<in> K\<close> show "\<exists>\<V>. finite \<V> \<and> \<V> \<subseteq> \<U> \<and> {x \<in> topspace X. f x = y} \<subseteq> \<Union>\<V>"
unfolding compactin_def using \<U> sub by fastforce
qed
then obtain \<V> where \<V>: "\<And>y. y \<in> K \<Longrightarrow> finite (\<V> y) \<and> \<V> y \<subseteq> \<U> \<and> {x \<in> topspace X. f x = y} \<subseteq> \<Union>(\<V> y)"
by (metis (full_types))
define F where "F \<equiv> \<lambda>y. topspace Y - f ` (topspace X - \<Union>(\<V> y))"
have "\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> F ` K \<and> K \<subseteq> \<Union>\<F>"
proof (rule compactinD [OF \<open>compactin Y K\<close>])
have "\<And>x. x \<in> K \<Longrightarrow> closedin Y (f ` (topspace X - \<Union>(\<V> x)))"
using f unfolding proper_map_def closed_map_def
by (meson \<U> \<V> openin_Union openin_closedin_eq subsetD)
then show "openin Y U" if "U \<in> F ` K" for U
using that by (auto simp: F_def)
show "K \<subseteq> \<Union>(F ` K)"
using \<V> \<open>compactin Y K\<close> unfolding F_def compactin_def by fastforce
qed
then obtain J where "finite J" "J \<subseteq> K" and J: "K \<subseteq> \<Union>(F ` J)"
by (auto simp: ex_finite_subset_image)
show ?thesis
unfolding F_def
proof (intro exI conjI)
show "finite (\<Union>(\<V> ` J))"
using \<V> \<open>J \<subseteq> K\<close> \<open>finite J\<close> by blast
show "\<Union>(\<V> ` J) \<subseteq> \<U>"
using \<V> \<open>J \<subseteq> K\<close> by blast
show "{x \<in> topspace X. f x \<in> K} \<subseteq> \<Union>(\<Union>(\<V> ` J))"
using J \<open>J \<subseteq> K\<close> unfolding F_def by auto
qed
qed
qed
qed
lemma compact_space_proper_map_preimage:
assumes f: "proper_map X Y f" and fim: "f ` (topspace X) = topspace Y" and "compact_space Y"
shows "compact_space X"
proof -
have eq: "topspace X = {x \<in> topspace X. f x \<in> topspace Y}"
using fim by blast
moreover have "compactin Y (topspace Y)"
using \<open>compact_space Y\<close> compact_space_def by auto
ultimately show ?thesis
unfolding compact_space_def
using eq f compactin_proper_map_preimage by fastforce
qed
lemma proper_map_alt:
"proper_map X Y f \<longleftrightarrow>
closed_map X Y f \<and> (\<forall>K. compactin Y K \<longrightarrow> compactin X {x. x \<in> topspace X \<and> f x \<in> K})"
proof (intro iffI conjI allI impI)
show "compactin X {x \<in> topspace X. f x \<in> K}"
if "proper_map X Y f" and "compactin Y K" for K
using that by (simp add: compactin_proper_map_preimage)
show "proper_map X Y f"
if f: "closed_map X Y f \<and> (\<forall>K. compactin Y K \<longrightarrow> compactin X {x \<in> topspace X. f x \<in> K})"
proof -
have "compactin X {x \<in> topspace X. f x = y}" if "y \<in> topspace Y" for y
proof -
have "compactin X {x \<in> topspace X. f x \<in> {y}}"
using f compactin_sing that by fastforce
then show ?thesis
by auto
qed
with f show ?thesis
by (auto simp: proper_map_def)
qed
qed (simp add: proper_imp_closed_map)
lemma proper_map_on_empty:
"topspace X = {} \<Longrightarrow> proper_map X Y f"
by (auto simp: proper_map_def closed_map_on_empty)
lemma proper_map_id [simp]:
"proper_map X X id"
proof (clarsimp simp: proper_map_alt closed_map_id)
fix K
assume K: "compactin X K"
then have "{a \<in> topspace X. a \<in> K} = K"
by (simp add: compactin_subspace subset_antisym subset_iff)
then show "compactin X {a \<in> topspace X. a \<in> K}"
using K by auto
qed
lemma proper_map_compose:
assumes "proper_map X Y f" "proper_map Y Z g"
shows "proper_map X Z (g \<circ> f)"
proof -
have "closed_map X Y f" and f: "\<And>K. compactin Y K \<Longrightarrow> compactin X {x \<in> topspace X. f x \<in> K}"
and "closed_map Y Z g" and g: "\<And>K. compactin Z K \<Longrightarrow> compactin Y {x \<in> topspace Y. g x \<in> K}"
using assms by (auto simp: proper_map_alt)
show ?thesis
unfolding proper_map_alt
proof (intro conjI allI impI)
show "closed_map X Z (g \<circ> f)"
using \<open>closed_map X Y f\<close> \<open>closed_map Y Z g\<close> closed_map_compose by blast
have "{x \<in> topspace X. g (f x) \<in> K} = {x \<in> topspace X. f x \<in> {b \<in> topspace Y. g b \<in> K}}" for K
using \<open>closed_map X Y f\<close> closed_map_imp_subset_topspace by blast
then show "compactin X {x \<in> topspace X. (g \<circ> f) x \<in> K}"
if "compactin Z K" for K
using f [OF g [OF that]] by auto
qed
qed
lemma proper_map_const:
"proper_map X Y (\<lambda>x. c) \<longleftrightarrow> compact_space X \<and> (topspace X = {} \<or> closedin Y {c})"
proof (cases "topspace X = {}")
case True
then show ?thesis
by (simp add: compact_space_topspace_empty proper_map_on_empty)
next
case False
have *: "compactin X {x \<in> topspace X. c = y}" if "compact_space X" for y
proof (cases "c = y")
case True
then show ?thesis
using compact_space_def \<open>compact_space X\<close> by auto
qed auto
then show ?thesis
using closed_compactin closedin_subset
by (force simp: False proper_map_def closed_map_const compact_space_def)
qed
lemma proper_map_inclusion:
"s \<subseteq> topspace X
\<Longrightarrow> proper_map (subtopology X s) X id \<longleftrightarrow> closedin X s \<and> (\<forall>k. compactin X k \<longrightarrow> compactin X (s \<inter> k))"
by (auto simp: proper_map_alt closed_map_inclusion_eq inf.absorb_iff2 Collect_conj_eq compactin_subtopology intro: closed_Int_compactin)
subsection\<open>Perfect maps (proper, continuous and surjective)\<close>
definition perfect_map
where "perfect_map X Y f \<equiv> continuous_map X Y f \<and> proper_map X Y f \<and> f ` (topspace X) = topspace Y"
lemma homeomorphic_imp_perfect_map:
"homeomorphic_map X Y f \<Longrightarrow> perfect_map X Y f"
by (simp add: homeomorphic_eq_everything_map homeomorphic_imp_proper_map perfect_map_def)
lemma perfect_imp_quotient_map:
"perfect_map X Y f \<Longrightarrow> quotient_map X Y f"
by (simp add: continuous_closed_imp_quotient_map perfect_map_def proper_map_def)
lemma homeomorphic_eq_injective_perfect_map:
"homeomorphic_map X Y f \<longleftrightarrow> perfect_map X Y f \<and> inj_on f (topspace X)"
using homeomorphic_imp_perfect_map homeomorphic_map_def perfect_imp_quotient_map by blast
lemma perfect_injective_eq_homeomorphic_map:
"perfect_map X Y f \<and> inj_on f (topspace X) \<longleftrightarrow> homeomorphic_map X Y f"
by (simp add: homeomorphic_eq_injective_perfect_map)
lemma perfect_map_id [simp]: "perfect_map X X id"
by (simp add: homeomorphic_imp_perfect_map)
lemma perfect_map_compose:
"\<lbrakk>perfect_map X Y f; perfect_map Y Z g\<rbrakk> \<Longrightarrow> perfect_map X Z (g \<circ> f)"
by (meson continuous_map_compose perfect_imp_quotient_map perfect_map_def proper_map_compose quotient_map_compose_eq quotient_map_def)
lemma perfect_imp_continuous_map:
"perfect_map X Y f \<Longrightarrow> continuous_map X Y f"
using perfect_map_def by blast
lemma perfect_imp_closed_map:
"perfect_map X Y f \<Longrightarrow> closed_map X Y f"
by (simp add: perfect_map_def proper_map_def)
lemma perfect_imp_proper_map:
"perfect_map X Y f \<Longrightarrow> proper_map X Y f"
by (simp add: perfect_map_def)
lemma perfect_imp_surjective_map:
"perfect_map X Y f \<Longrightarrow> f ` (topspace X) = topspace Y"
by (simp add: perfect_map_def)
end
diff --git a/src/HOL/Analysis/Arcwise_Connected.thy b/src/HOL/Analysis/Arcwise_Connected.thy
--- a/src/HOL/Analysis/Arcwise_Connected.thy
+++ b/src/HOL/Analysis/Arcwise_Connected.thy
@@ -1,2232 +1,2232 @@
(* Title: HOL/Analysis/Arcwise_Connected.thy
Authors: LC Paulson, based on material from HOL Light
*)
section \<open>Arcwise-Connected Sets\<close>
theory Arcwise_Connected
imports Path_Connected Ordered_Euclidean_Space "HOL-Computational_Algebra.Primes"
begin
lemma path_connected_interval [simp]:
fixes a b::"'a::ordered_euclidean_space"
shows "path_connected {a..b}"
using is_interval_cc is_interval_path_connected by blast
lemma segment_to_closest_point:
fixes S :: "'a :: euclidean_space set"
shows "\<lbrakk>closed S; S \<noteq> {}\<rbrakk> \<Longrightarrow> open_segment a (closest_point S a) \<inter> S = {}"
apply (subst disjoint_iff_not_equal)
apply (clarify dest!: dist_in_open_segment)
by (metis closest_point_le dist_commute le_less_trans less_irrefl)
lemma segment_to_point_exists:
fixes S :: "'a :: euclidean_space set"
assumes "closed S" "S \<noteq> {}"
obtains b where "b \<in> S" "open_segment a b \<inter> S = {}"
by (metis assms segment_to_closest_point closest_point_exists that)
subsection \<open>The Brouwer reduction theorem\<close>
theorem Brouwer_reduction_theorem_gen:
fixes S :: "'a::euclidean_space set"
assumes "closed S" "\<phi> S"
and \<phi>: "\<And>F. \<lbrakk>\<And>n. closed(F n); \<And>n. \<phi>(F n); \<And>n. F(Suc n) \<subseteq> F n\<rbrakk> \<Longrightarrow> \<phi>(\<Inter>(range F))"
obtains T where "T \<subseteq> S" "closed T" "\<phi> T" "\<And>U. \<lbrakk>U \<subseteq> S; closed U; \<phi> U\<rbrakk> \<Longrightarrow> \<not> (U \<subset> T)"
proof -
obtain B :: "nat \<Rightarrow> 'a set"
where "inj B" "\<And>n. open(B n)" and open_cov: "\<And>S. open S \<Longrightarrow> \<exists>K. S = \<Union>(B ` K)"
by (metis Setcompr_eq_image that univ_second_countable_sequence)
define A where "A \<equiv> rec_nat S (\<lambda>n a. if \<exists>U. U \<subseteq> a \<and> closed U \<and> \<phi> U \<and> U \<inter> (B n) = {}
then SOME U. U \<subseteq> a \<and> closed U \<and> \<phi> U \<and> U \<inter> (B n) = {}
else a)"
have [simp]: "A 0 = S"
by (simp add: A_def)
have ASuc: "A(Suc n) = (if \<exists>U. U \<subseteq> A n \<and> closed U \<and> \<phi> U \<and> U \<inter> (B n) = {}
then SOME U. U \<subseteq> A n \<and> closed U \<and> \<phi> U \<and> U \<inter> (B n) = {}
else A n)" for n
by (auto simp: A_def)
have sub: "\<And>n. A(Suc n) \<subseteq> A n"
by (auto simp: ASuc dest!: someI_ex)
have subS: "A n \<subseteq> S" for n
by (induction n) (use sub in auto)
have clo: "closed (A n) \<and> \<phi> (A n)" for n
by (induction n) (auto simp: assms ASuc dest!: someI_ex)
show ?thesis
proof
show "\<Inter>(range A) \<subseteq> S"
using \<open>\<And>n. A n \<subseteq> S\<close> by blast
show "closed (\<Inter>(A ` UNIV))"
using clo by blast
show "\<phi> (\<Inter>(A ` UNIV))"
by (simp add: clo \<phi> sub)
show "\<not> U \<subset> \<Inter>(A ` UNIV)" if "U \<subseteq> S" "closed U" "\<phi> U" for U
proof -
have "\<exists>y. x \<notin> A y" if "x \<notin> U" and Usub: "U \<subseteq> (\<Inter>x. A x)" for x
proof -
obtain e where "e > 0" and e: "ball x e \<subseteq> -U"
using \<open>closed U\<close> \<open>x \<notin> U\<close> openE [of "-U"] by blast
moreover obtain K where K: "ball x e = \<Union>(B ` K)"
using open_cov [of "ball x e"] by auto
ultimately have "\<Union>(B ` K) \<subseteq> -U"
by blast
have "K \<noteq> {}"
using \<open>0 < e\<close> \<open>ball x e = \<Union>(B ` K)\<close> by auto
then obtain n where "n \<in> K" "x \<in> B n"
by (metis K UN_E \<open>0 < e\<close> centre_in_ball)
then have "U \<inter> B n = {}"
using K e by auto
show ?thesis
proof (cases "\<exists>U\<subseteq>A n. closed U \<and> \<phi> U \<and> U \<inter> B n = {}")
case True
then show ?thesis
apply (rule_tac x="Suc n" in exI)
apply (simp add: ASuc)
apply (erule someI2_ex)
using \<open>x \<in> B n\<close> by blast
next
case False
then show ?thesis
by (meson Inf_lower Usub \<open>U \<inter> B n = {}\<close> \<open>\<phi> U\<close> \<open>closed U\<close> range_eqI subset_trans)
qed
qed
with that show ?thesis
by (meson Inter_iff psubsetE rangeI subsetI)
qed
qed
qed
corollary Brouwer_reduction_theorem:
fixes S :: "'a::euclidean_space set"
assumes "compact S" "\<phi> S" "S \<noteq> {}"
and \<phi>: "\<And>F. \<lbrakk>\<And>n. compact(F n); \<And>n. F n \<noteq> {}; \<And>n. \<phi>(F n); \<And>n. F(Suc n) \<subseteq> F n\<rbrakk> \<Longrightarrow> \<phi>(\<Inter>(range F))"
obtains T where "T \<subseteq> S" "compact T" "T \<noteq> {}" "\<phi> T"
"\<And>U. \<lbrakk>U \<subseteq> S; closed U; U \<noteq> {}; \<phi> U\<rbrakk> \<Longrightarrow> \<not> (U \<subset> T)"
proof (rule Brouwer_reduction_theorem_gen [of S "\<lambda>T. T \<noteq> {} \<and> T \<subseteq> S \<and> \<phi> T"])
fix F
assume cloF: "\<And>n. closed (F n)"
and F: "\<And>n. F n \<noteq> {} \<and> F n \<subseteq> S \<and> \<phi> (F n)" and Fsub: "\<And>n. F (Suc n) \<subseteq> F n"
show "\<Inter>(F ` UNIV) \<noteq> {} \<and> \<Inter>(F ` UNIV) \<subseteq> S \<and> \<phi> (\<Inter>(F ` UNIV))"
proof (intro conjI)
show "\<Inter>(F ` UNIV) \<noteq> {}"
apply (rule compact_nest)
apply (meson F cloF \<open>compact S\<close> seq_compact_closed_subset seq_compact_eq_compact)
apply (simp add: F)
by (meson Fsub lift_Suc_antimono_le)
show " \<Inter>(F ` UNIV) \<subseteq> S"
using F by blast
show "\<phi> (\<Inter>(F ` UNIV))"
by (metis F Fsub \<phi> \<open>compact S\<close> cloF closed_Int_compact inf.orderE)
qed
next
show "S \<noteq> {} \<and> S \<subseteq> S \<and> \<phi> S"
by (simp add: assms)
qed (meson assms compact_imp_closed seq_compact_closed_subset seq_compact_eq_compact)+
subsection\<^marker>\<open>tag unimportant\<close>\<open>Arcwise Connections\<close>(*FIX ME this subsection is empty(?) *)
subsection\<open>Density of points with dyadic rational coordinates\<close>
proposition closure_dyadic_rationals:
"closure (\<Union>k. \<Union>f \<in> Basis \<rightarrow> \<int>.
{ \<Sum>i :: 'a :: euclidean_space \<in> Basis. (f i / 2^k) *\<^sub>R i }) = UNIV"
proof -
have "x \<in> closure (\<Union>k. \<Union>f \<in> Basis \<rightarrow> \<int>. {\<Sum>i \<in> Basis. (f i / 2^k) *\<^sub>R i})" for x::'a
proof (clarsimp simp: closure_approachable)
fix e::real
assume "e > 0"
then obtain k where k: "(1/2)^k < e/DIM('a)"
by (meson DIM_positive divide_less_eq_1_pos of_nat_0_less_iff one_less_numeral_iff real_arch_pow_inv semiring_norm(76) zero_less_divide_iff zero_less_numeral)
have "dist (\<Sum>i\<in>Basis. (real_of_int \<lfloor>2^k*(x \<bullet> i)\<rfloor> / 2^k) *\<^sub>R i) x =
dist (\<Sum>i\<in>Basis. (real_of_int \<lfloor>2^k*(x \<bullet> i)\<rfloor> / 2^k) *\<^sub>R i) (\<Sum>i\<in>Basis. (x \<bullet> i) *\<^sub>R i)"
by (simp add: euclidean_representation)
also have "... = norm ((\<Sum>i\<in>Basis. (real_of_int \<lfloor>2^k*(x \<bullet> i)\<rfloor> / 2^k) *\<^sub>R i - (x \<bullet> i) *\<^sub>R i))"
by (simp add: dist_norm sum_subtractf)
also have "... \<le> DIM('a)*((1/2)^k)"
proof (rule sum_norm_bound, simp add: algebra_simps)
fix i::'a
assume "i \<in> Basis"
then have "norm ((real_of_int \<lfloor>x \<bullet> i*2^k\<rfloor> / 2^k) *\<^sub>R i - (x \<bullet> i) *\<^sub>R i) =
\<bar>real_of_int \<lfloor>x \<bullet> i*2^k\<rfloor> / 2^k - x \<bullet> i\<bar>"
by (simp add: scaleR_left_diff_distrib [symmetric])
also have "... \<le> (1/2) ^ k"
by (simp add: divide_simps) linarith
finally show "norm ((real_of_int \<lfloor>x \<bullet> i*2^k\<rfloor> / 2^k) *\<^sub>R i - (x \<bullet> i) *\<^sub>R i) \<le> (1/2) ^ k" .
qed
also have "... < DIM('a)*(e/DIM('a))"
using DIM_positive k linordered_comm_semiring_strict_class.comm_mult_strict_left_mono of_nat_0_less_iff by blast
also have "... = e"
by simp
finally have "dist (\<Sum>i\<in>Basis. (\<lfloor>2^k*(x \<bullet> i)\<rfloor> / 2^k) *\<^sub>R i) x < e" .
then
show "\<exists>k. \<exists>f \<in> Basis \<rightarrow> \<int>. dist (\<Sum>b\<in>Basis. (f b / 2^k) *\<^sub>R b) x < e"
apply (rule_tac x=k in exI)
apply (rule_tac x="\<lambda>i. of_int (floor (2^k*(x \<bullet> i)))" in bexI)
apply auto
done
qed
then show ?thesis by auto
qed
corollary closure_rational_coordinates:
"closure (\<Union>f \<in> Basis \<rightarrow> \<rat>. { \<Sum>i :: 'a :: euclidean_space \<in> Basis. f i *\<^sub>R i }) = UNIV"
proof -
have *: "(\<Union>k. \<Union>f \<in> Basis \<rightarrow> \<int>. { \<Sum>i::'a \<in> Basis. (f i / 2^k) *\<^sub>R i })
\<subseteq> (\<Union>f \<in> Basis \<rightarrow> \<rat>. { \<Sum>i \<in> Basis. f i *\<^sub>R i })"
proof clarsimp
fix k and f :: "'a \<Rightarrow> real"
assume f: "f \<in> Basis \<rightarrow> \<int>"
show "\<exists>x \<in> Basis \<rightarrow> \<rat>. (\<Sum>i \<in> Basis. (f i / 2^k) *\<^sub>R i) = (\<Sum>i \<in> Basis. x i *\<^sub>R i)"
apply (rule_tac x="\<lambda>i. f i / 2^k" in bexI)
using Ints_subset_Rats f by auto
qed
show ?thesis
using closure_dyadic_rationals closure_mono [OF *] by blast
qed
lemma closure_dyadic_rationals_in_convex_set:
"\<lbrakk>convex S; interior S \<noteq> {}\<rbrakk>
\<Longrightarrow> closure(S \<inter>
(\<Union>k. \<Union>f \<in> Basis \<rightarrow> \<int>.
{ \<Sum>i :: 'a :: euclidean_space \<in> Basis. (f i / 2^k) *\<^sub>R i })) =
closure S"
by (simp add: closure_dyadic_rationals closure_convex_Int_superset)
lemma closure_rationals_in_convex_set:
"\<lbrakk>convex S; interior S \<noteq> {}\<rbrakk>
\<Longrightarrow> closure(S \<inter> (\<Union>f \<in> Basis \<rightarrow> \<rat>. { \<Sum>i :: 'a :: euclidean_space \<in> Basis. f i *\<^sub>R i })) =
closure S"
by (simp add: closure_rational_coordinates closure_convex_Int_superset)
text\<open> Every path between distinct points contains an arc, and hence
path connection is equivalent to arcwise connection for distinct points.
The proof is based on Whyburn's "Topological Analysis".\<close>
lemma closure_dyadic_rationals_in_convex_set_pos_1:
fixes S :: "real set"
assumes "convex S" and intnz: "interior S \<noteq> {}" and pos: "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> x"
shows "closure(S \<inter> (\<Union>k m. {of_nat m / 2^k})) = closure S"
proof -
have "\<exists>m. f 1/2^k = real m / 2^k" if "(f 1) / 2^k \<in> S" "f 1 \<in> \<int>" for k and f :: "real \<Rightarrow> real"
using that by (force simp: Ints_def zero_le_divide_iff power_le_zero_eq dest: pos zero_le_imp_eq_int)
then have "S \<inter> (\<Union>k m. {real m / 2^k}) = S \<inter>
(\<Union>k. \<Union>f\<in>Basis \<rightarrow> \<int>. {\<Sum>i\<in>Basis. (f i / 2^k) *\<^sub>R i})"
by force
then show ?thesis
using closure_dyadic_rationals_in_convex_set [OF \<open>convex S\<close> intnz] by simp
qed
definition\<^marker>\<open>tag unimportant\<close> dyadics :: "'a::field_char_0 set" where "dyadics \<equiv> \<Union>k m. {of_nat m / 2^k}"
lemma real_in_dyadics [simp]: "real m \<in> dyadics"
apply (simp add: dyadics_def)
by (metis divide_numeral_1 numeral_One power_0)
lemma nat_neq_4k1: "of_nat m \<noteq> (4 * of_nat k + 1) / (2 * 2^n :: 'a::field_char_0)"
proof
assume "of_nat m = (4 * of_nat k + 1) / (2 * 2^n :: 'a)"
then have "of_nat (m * (2 * 2^n)) = (of_nat (Suc (4 * k)) :: 'a)"
by (simp add: field_split_simps)
then have "m * (2 * 2^n) = Suc (4 * k)"
using of_nat_eq_iff by blast
then have "odd (m * (2 * 2^n))"
by simp
then show False
by simp
qed
lemma nat_neq_4k3: "of_nat m \<noteq> (4 * of_nat k + 3) / (2 * 2^n :: 'a::field_char_0)"
proof
assume "of_nat m = (4 * of_nat k + 3) / (2 * 2^n :: 'a)"
then have "of_nat (m * (2 * 2^n)) = (of_nat (4 * k + 3) :: 'a)"
by (simp add: field_split_simps)
then have "m * (2 * 2^n) = (4 * k) + 3"
using of_nat_eq_iff by blast
then have "odd (m * (2 * 2^n))"
by simp
then show False
by simp
qed
lemma iff_4k:
assumes "r = real k" "odd k"
shows "(4 * real m + r) / (2 * 2^n) = (4 * real m' + r) / (2 * 2 ^ n') \<longleftrightarrow> m=m' \<and> n=n'"
proof -
{ assume "(4 * real m + r) / (2 * 2^n) = (4 * real m' + r) / (2 * 2 ^ n')"
then have "real ((4 * m + k) * (2 * 2 ^ n')) = real ((4 * m' + k) * (2 * 2^n))"
using assms by (auto simp: field_simps)
then have "(4 * m + k) * (2 * 2 ^ n') = (4 * m' + k) * (2 * 2^n)"
using of_nat_eq_iff by blast
then have "(4 * m + k) * (2 ^ n') = (4 * m' + k) * (2^n)"
by linarith
then obtain "4*m + k = 4*m' + k" "n=n'"
apply (rule prime_power_cancel2 [OF two_is_prime_nat])
using assms by auto
then have "m=m'" "n=n'"
by auto
}
then show ?thesis by blast
qed
lemma neq_4k1_k43: "(4 * real m + 1) / (2 * 2^n) \<noteq> (4 * real m' + 3) / (2 * 2 ^ n')"
proof
assume "(4 * real m + 1) / (2 * 2^n) = (4 * real m' + 3) / (2 * 2 ^ n')"
then have "real (Suc (4 * m) * (2 * 2 ^ n')) = real ((4 * m' + 3) * (2 * 2^n))"
by (auto simp: field_simps)
then have "Suc (4 * m) * (2 * 2 ^ n') = (4 * m' + 3) * (2 * 2^n)"
using of_nat_eq_iff by blast
then have "Suc (4 * m) * (2 ^ n') = (4 * m' + 3) * (2^n)"
by linarith
then have "Suc (4 * m) = (4 * m' + 3)"
by (rule prime_power_cancel2 [OF two_is_prime_nat]) auto
then have "1 + 2 * m' = 2 * m"
using \<open>Suc (4 * m) = 4 * m' + 3\<close> by linarith
then show False
using even_Suc by presburger
qed
lemma dyadic_413_cases:
obtains "(of_nat m::'a::field_char_0) / 2^k \<in> Nats"
| m' k' where "k' < k" "(of_nat m:: 'a) / 2^k = of_nat (4*m' + 1) / 2^Suc k'"
| m' k' where "k' < k" "(of_nat m:: 'a) / 2^k = of_nat (4*m' + 3) / 2^Suc k'"
proof (cases "m>0")
case False
then have "m=0" by simp
with that show ?thesis by auto
next
case True
obtain k' m' where m': "odd m'" and k': "m = m' * 2^k'"
using prime_power_canonical [OF two_is_prime_nat True] by blast
then obtain q r where q: "m' = 4*q + r" and r: "r < 4"
by (metis not_add_less2 split_div zero_neq_numeral)
show ?thesis
proof (cases "k \<le> k'")
case True
have "(of_nat m:: 'a) / 2^k = of_nat m' * (2 ^ k' / 2^k)"
using k' by (simp add: field_simps)
also have "... = (of_nat m'::'a) * 2 ^ (k'-k)"
using k' True by (simp add: power_diff)
also have "... \<in> \<nat>"
by (metis Nats_mult of_nat_in_Nats of_nat_numeral of_nat_power)
finally show ?thesis by (auto simp: that)
next
case False
then obtain kd where kd: "Suc kd = k - k'"
using Suc_diff_Suc not_less by blast
have "(of_nat m:: 'a) / 2^k = of_nat m' * (2 ^ k' / 2^k)"
using k' by (simp add: field_simps)
also have "... = (of_nat m'::'a) / 2 ^ (k-k')"
using k' False by (simp add: power_diff)
also have "... = ((of_nat r + 4 * of_nat q)::'a) / 2 ^ (k-k')"
using q by force
finally have meq: "(of_nat m:: 'a) / 2^k = (of_nat r + 4 * of_nat q) / 2 ^ (k - k')" .
have "r \<noteq> 0" "r \<noteq> 2"
using q m' by presburger+
with r consider "r = 1" | "r = 3"
by linarith
then show ?thesis
proof cases
assume "r = 1"
with meq kd that(2) [of kd q] show ?thesis
by simp
next
assume "r = 3"
with meq kd that(3) [of kd q] show ?thesis
by simp
qed
qed
qed
lemma dyadics_iff:
"(dyadics :: 'a::field_char_0 set) =
Nats \<union> (\<Union>k m. {of_nat (4*m + 1) / 2^Suc k}) \<union> (\<Union>k m. {of_nat (4*m + 3) / 2^Suc k})"
(is "_ = ?rhs")
proof
show "dyadics \<subseteq> ?rhs"
unfolding dyadics_def
apply clarify
apply (rule dyadic_413_cases, force+)
done
next
show "?rhs \<subseteq> dyadics"
apply (clarsimp simp: dyadics_def Nats_def simp del: power_Suc)
apply (intro conjI subsetI)
apply (auto simp del: power_Suc)
apply (metis divide_numeral_1 numeral_One power_0)
apply (metis of_nat_Suc of_nat_mult of_nat_numeral)
by (metis of_nat_add of_nat_mult of_nat_numeral)
qed
function\<^marker>\<open>tag unimportant\<close> (domintros) dyad_rec :: "[nat \<Rightarrow> 'a, 'a\<Rightarrow>'a, 'a\<Rightarrow>'a, real] \<Rightarrow> 'a" where
"dyad_rec b l r (real m) = b m"
| "dyad_rec b l r ((4 * real m + 1) / 2 ^ (Suc n)) = l (dyad_rec b l r ((2*m + 1) / 2^n))"
| "dyad_rec b l r ((4 * real m + 3) / 2 ^ (Suc n)) = r (dyad_rec b l r ((2*m + 1) / 2^n))"
| "x \<notin> dyadics \<Longrightarrow> dyad_rec b l r x = undefined"
using iff_4k [of _ 1] iff_4k [of _ 3]
apply (simp_all add: nat_neq_4k1 nat_neq_4k3 neq_4k1_k43, atomize_elim)
apply (fastforce simp add: dyadics_iff Nats_def field_simps)+
done
lemma dyadics_levels: "dyadics = (\<Union>K. \<Union>k<K. \<Union> m. {of_nat m / 2^k})"
unfolding dyadics_def by auto
lemma dyad_rec_level_termination:
assumes "k < K"
shows "dyad_rec_dom(b, l, r, real m / 2^k)"
using assms
proof (induction K arbitrary: k m)
case 0
then show ?case by auto
next
case (Suc K)
then consider "k = K" | "k < K"
using less_antisym by blast
then show ?case
proof cases
assume "k = K"
show ?case
proof (rule dyadic_413_cases [of m k, where 'a=real])
show "real m / 2^k \<in> \<nat> \<Longrightarrow> dyad_rec_dom (b, l, r, real m / 2^k)"
by (force simp: Nats_def nat_neq_4k1 nat_neq_4k3 intro: dyad_rec.domintros)
show ?case if "k' < k" and eq: "real m / 2^k = real (4 * m' + 1) / 2^Suc k'" for m' k'
proof -
have "dyad_rec_dom (b, l, r, (4 * real m' + 1) / 2^Suc k')"
proof (rule dyad_rec.domintros)
fix m n
assume "(4 * real m' + 1) / (2 * 2 ^ k') = (4 * real m + 1) / (2 * 2^n)"
then have "m' = m" "k' = n" using iff_4k [of _ 1]
by auto
have "dyad_rec_dom (b, l, r, real (2 * m + 1) / 2 ^ k')"
using Suc.IH \<open>k = K\<close> \<open>k' < k\<close> by blast
then show "dyad_rec_dom (b, l, r, (2 * real m + 1) / 2^n)"
using \<open>k' = n\<close> by (auto simp: algebra_simps)
next
fix m n
assume "(4 * real m' + 1) / (2 * 2 ^ k') = (4 * real m + 3) / (2 * 2^n)"
then have "False"
by (metis neq_4k1_k43)
then show "dyad_rec_dom (b, l, r, (2 * real m + 1) / 2^n)" ..
qed
then show ?case by (simp add: eq add_ac)
qed
show ?case if "k' < k" and eq: "real m / 2^k = real (4 * m' + 3) / 2^Suc k'" for m' k'
proof -
have "dyad_rec_dom (b, l, r, (4 * real m' + 3) / 2^Suc k')"
proof (rule dyad_rec.domintros)
fix m n
assume "(4 * real m' + 3) / (2 * 2 ^ k') = (4 * real m + 1) / (2 * 2^n)"
then have "False"
by (metis neq_4k1_k43)
then show "dyad_rec_dom (b, l, r, (2 * real m + 1) / 2^n)" ..
next
fix m n
assume "(4 * real m' + 3) / (2 * 2 ^ k') = (4 * real m + 3) / (2 * 2^n)"
then have "m' = m" "k' = n" using iff_4k [of _ 3]
by auto
have "dyad_rec_dom (b, l, r, real (2 * m + 1) / 2 ^ k')"
using Suc.IH \<open>k = K\<close> \<open>k' < k\<close> by blast
then show "dyad_rec_dom (b, l, r, (2 * real m + 1) / 2^n)"
using \<open>k' = n\<close> by (auto simp: algebra_simps)
qed
then show ?case by (simp add: eq add_ac)
qed
qed
next
assume "k < K"
then show ?case
using Suc.IH by blast
qed
qed
lemma dyad_rec_termination: "x \<in> dyadics \<Longrightarrow> dyad_rec_dom(b,l,r,x)"
by (auto simp: dyadics_levels intro: dyad_rec_level_termination)
lemma dyad_rec_of_nat [simp]: "dyad_rec b l r (real m) = b m"
by (simp add: dyad_rec.psimps dyad_rec_termination)
lemma dyad_rec_41 [simp]: "dyad_rec b l r ((4 * real m + 1) / 2 ^ (Suc n)) = l (dyad_rec b l r ((2*m + 1) / 2^n))"
apply (rule dyad_rec.psimps)
by (metis dyad_rec_level_termination lessI add.commute of_nat_Suc of_nat_mult of_nat_numeral)
lemma dyad_rec_43 [simp]: "dyad_rec b l r ((4 * real m + 3) / 2 ^ (Suc n)) = r (dyad_rec b l r ((2*m + 1) / 2^n))"
apply (rule dyad_rec.psimps)
by (metis dyad_rec_level_termination lessI of_nat_add of_nat_mult of_nat_numeral)
lemma dyad_rec_41_times2:
assumes "n > 0"
shows "dyad_rec b l r (2 * ((4 * real m + 1) / 2^Suc n)) = l (dyad_rec b l r (2 * (2 * real m + 1) / 2^n))"
proof -
obtain n' where n': "n = Suc n'"
using assms not0_implies_Suc by blast
have "dyad_rec b l r (2 * ((4 * real m + 1) / 2^Suc n)) = dyad_rec b l r ((2 * (4 * real m + 1)) / (2 * 2^n))"
by auto
also have "... = dyad_rec b l r ((4 * real m + 1) / 2^n)"
by (subst mult_divide_mult_cancel_left) auto
also have "... = l (dyad_rec b l r ((2 * real m + 1) / 2 ^ n'))"
by (simp add: add.commute [of 1] n' del: power_Suc)
also have "... = l (dyad_rec b l r ((2 * (2 * real m + 1)) / (2 * 2 ^ n')))"
by (subst mult_divide_mult_cancel_left) auto
also have "... = l (dyad_rec b l r (2 * (2 * real m + 1) / 2^n))"
by (simp add: add.commute n')
finally show ?thesis .
qed
lemma dyad_rec_43_times2:
assumes "n > 0"
shows "dyad_rec b l r (2 * ((4 * real m + 3) / 2^Suc n)) = r (dyad_rec b l r (2 * (2 * real m + 1) / 2^n))"
proof -
obtain n' where n': "n = Suc n'"
using assms not0_implies_Suc by blast
have "dyad_rec b l r (2 * ((4 * real m + 3) / 2^Suc n)) = dyad_rec b l r ((2 * (4 * real m + 3)) / (2 * 2^n))"
by auto
also have "... = dyad_rec b l r ((4 * real m + 3) / 2^n)"
by (subst mult_divide_mult_cancel_left) auto
also have "... = r (dyad_rec b l r ((2 * real m + 1) / 2 ^ n'))"
by (simp add: n' del: power_Suc)
also have "... = r (dyad_rec b l r ((2 * (2 * real m + 1)) / (2 * 2 ^ n')))"
by (subst mult_divide_mult_cancel_left) auto
also have "... = r (dyad_rec b l r (2 * (2 * real m + 1) / 2^n))"
by (simp add: n')
finally show ?thesis .
qed
definition\<^marker>\<open>tag unimportant\<close> dyad_rec2
where "dyad_rec2 u v lc rc x =
dyad_rec (\<lambda>z. (u,v)) (\<lambda>(a,b). (a, lc a b (midpoint a b))) (\<lambda>(a,b). (rc a b (midpoint a b), b)) (2*x)"
abbreviation\<^marker>\<open>tag unimportant\<close> leftrec where "leftrec u v lc rc x \<equiv> fst (dyad_rec2 u v lc rc x)"
abbreviation\<^marker>\<open>tag unimportant\<close> rightrec where "rightrec u v lc rc x \<equiv> snd (dyad_rec2 u v lc rc x)"
lemma leftrec_base: "leftrec u v lc rc (real m / 2) = u"
by (simp add: dyad_rec2_def)
lemma leftrec_41: "n > 0 \<Longrightarrow> leftrec u v lc rc ((4 * real m + 1) / 2 ^ (Suc n)) = leftrec u v lc rc ((2 * real m + 1) / 2^n)"
apply (simp only: dyad_rec2_def dyad_rec_41_times2)
apply (simp add: case_prod_beta)
done
lemma leftrec_43: "n > 0 \<Longrightarrow>
leftrec u v lc rc ((4 * real m + 3) / 2 ^ (Suc n)) =
rc (leftrec u v lc rc ((2 * real m + 1) / 2^n)) (rightrec u v lc rc ((2 * real m + 1) / 2^n))
(midpoint (leftrec u v lc rc ((2 * real m + 1) / 2^n)) (rightrec u v lc rc ((2 * real m + 1) / 2^n)))"
apply (simp only: dyad_rec2_def dyad_rec_43_times2)
apply (simp add: case_prod_beta)
done
lemma rightrec_base: "rightrec u v lc rc (real m / 2) = v"
by (simp add: dyad_rec2_def)
lemma rightrec_41: "n > 0 \<Longrightarrow>
rightrec u v lc rc ((4 * real m + 1) / 2 ^ (Suc n)) =
lc (leftrec u v lc rc ((2 * real m + 1) / 2^n)) (rightrec u v lc rc ((2 * real m + 1) / 2^n))
(midpoint (leftrec u v lc rc ((2 * real m + 1) / 2^n)) (rightrec u v lc rc ((2 * real m + 1) / 2^n)))"
apply (simp only: dyad_rec2_def dyad_rec_41_times2)
apply (simp add: case_prod_beta)
done
lemma rightrec_43: "n > 0 \<Longrightarrow> rightrec u v lc rc ((4 * real m + 3) / 2 ^ (Suc n)) = rightrec u v lc rc ((2 * real m + 1) / 2^n)"
apply (simp only: dyad_rec2_def dyad_rec_43_times2)
apply (simp add: case_prod_beta)
done
lemma dyadics_in_open_unit_interval:
"{0<..<1} \<inter> (\<Union>k m. {real m / 2^k}) = (\<Union>k. \<Union>m \<in> {0<..<2^k}. {real m / 2^k})"
by (auto simp: field_split_simps)
theorem homeomorphic_monotone_image_interval:
fixes f :: "real \<Rightarrow> 'a::{real_normed_vector,complete_space}"
assumes cont_f: "continuous_on {0..1} f"
and conn: "\<And>y. connected ({0..1} \<inter> f -` {y})"
and f_1not0: "f 1 \<noteq> f 0"
shows "(f ` {0..1}) homeomorphic {0..1::real}"
proof -
have "\<exists>c d. a \<le> c \<and> c \<le> m \<and> m \<le> d \<and> d \<le> b \<and>
(\<forall>x \<in> {c..d}. f x = f m) \<and>
(\<forall>x \<in> {a..<c}. (f x \<noteq> f m)) \<and>
(\<forall>x \<in> {d<..b}. (f x \<noteq> f m)) \<and>
(\<forall>x \<in> {a..<c}. \<forall>y \<in> {d<..b}. f x \<noteq> f y)"
if m: "m \<in> {a..b}" and ab01: "{a..b} \<subseteq> {0..1}" for a b m
proof -
have comp: "compact (f -` {f m} \<inter> {0..1})"
by (simp add: compact_eq_bounded_closed bounded_Int closed_vimage_Int cont_f)
obtain c0 d0 where cd0: "{0..1} \<inter> f -` {f m} = {c0..d0}"
using connected_compact_interval_1 [of "{0..1} \<inter> f -` {f m}"] conn comp
by (metis Int_commute)
with that have "m \<in> cbox c0 d0"
by auto
obtain c d where cd: "{a..b} \<inter> f -` {f m} = {c..d}"
apply (rule_tac c="max a c0" and d="min b d0" in that)
using ab01 cd0 by auto
then have cdab: "{c..d} \<subseteq> {a..b}"
by blast
show ?thesis
proof (intro exI conjI ballI)
show "a \<le> c" "d \<le> b"
using cdab cd m by auto
show "c \<le> m" "m \<le> d"
using cd m by auto
show "\<And>x. x \<in> {c..d} \<Longrightarrow> f x = f m"
using cd by blast
show "f x \<noteq> f m" if "x \<in> {a..<c}" for x
using that m cd [THEN equalityD1, THEN subsetD] \<open>c \<le> m\<close> by force
show "f x \<noteq> f m" if "x \<in> {d<..b}" for x
using that m cd [THEN equalityD1, THEN subsetD, of x] \<open>m \<le> d\<close> by force
show "f x \<noteq> f y" if "x \<in> {a..<c}" "y \<in> {d<..b}" for x y
proof (cases "f x = f m \<or> f y = f m")
case True
then show ?thesis
using \<open>\<And>x. x \<in> {a..<c} \<Longrightarrow> f x \<noteq> f m\<close> that by auto
next
case False
have False if "f x = f y"
proof -
have "x \<le> m" "m \<le> y"
using \<open>c \<le> m\<close> \<open>x \<in> {a..<c}\<close> \<open>m \<le> d\<close> \<open>y \<in> {d<..b}\<close> by auto
then have "x \<in> ({0..1} \<inter> f -` {f y})" "y \<in> ({0..1} \<inter> f -` {f y})"
using \<open>x \<in> {a..<c}\<close> \<open>y \<in> {d<..b}\<close> ab01 by (auto simp: that)
then have "m \<in> ({0..1} \<inter> f -` {f y})"
by (meson \<open>m \<le> y\<close> \<open>x \<le> m\<close> is_interval_connected_1 conn [of "f y"] is_interval_1)
with False show False by auto
qed
then show ?thesis by auto
qed
qed
qed
then obtain leftcut rightcut where LR:
"\<And>a b m. \<lbrakk>m \<in> {a..b}; {a..b} \<subseteq> {0..1}\<rbrakk> \<Longrightarrow>
(a \<le> leftcut a b m \<and> leftcut a b m \<le> m \<and> m \<le> rightcut a b m \<and> rightcut a b m \<le> b \<and>
(\<forall>x \<in> {leftcut a b m..rightcut a b m}. f x = f m) \<and>
(\<forall>x \<in> {a..<leftcut a b m}. f x \<noteq> f m) \<and>
(\<forall>x \<in> {rightcut a b m<..b}. f x \<noteq> f m) \<and>
(\<forall>x \<in> {a..<leftcut a b m}. \<forall>y \<in> {rightcut a b m<..b}. f x \<noteq> f y))"
apply atomize
apply (clarsimp simp only: imp_conjL [symmetric] choice_iff choice_iff')
apply (rule that, blast)
done
then have left_right: "\<And>a b m. \<lbrakk>m \<in> {a..b}; {a..b} \<subseteq> {0..1}\<rbrakk> \<Longrightarrow> a \<le> leftcut a b m \<and> rightcut a b m \<le> b"
and left_right_m: "\<And>a b m. \<lbrakk>m \<in> {a..b}; {a..b} \<subseteq> {0..1}\<rbrakk> \<Longrightarrow> leftcut a b m \<le> m \<and> m \<le> rightcut a b m"
by auto
have left_neq: "\<lbrakk>a \<le> x; x < leftcut a b m; a \<le> m; m \<le> b; {a..b} \<subseteq> {0..1}\<rbrakk> \<Longrightarrow> f x \<noteq> f m"
and right_neq: "\<lbrakk>rightcut a b m < x; x \<le> b; a \<le> m; m \<le> b; {a..b} \<subseteq> {0..1}\<rbrakk> \<Longrightarrow> f x \<noteq> f m"
and left_right_neq: "\<lbrakk>a \<le> x; x < leftcut a b m; rightcut a b m < y; y \<le> b; a \<le> m; m \<le> b; {a..b} \<subseteq> {0..1}\<rbrakk> \<Longrightarrow> f x \<noteq> f m"
and feqm: "\<lbrakk>leftcut a b m \<le> x; x \<le> rightcut a b m; a \<le> m; m \<le> b; {a..b} \<subseteq> {0..1}\<rbrakk>
\<Longrightarrow> f x = f m" for a b m x y
by (meson atLeastAtMost_iff greaterThanAtMost_iff atLeastLessThan_iff LR)+
have f_eqI: "\<And>a b m x y. \<lbrakk>leftcut a b m \<le> x; x \<le> rightcut a b m; leftcut a b m \<le> y; y \<le> rightcut a b m;
a \<le> m; m \<le> b; {a..b} \<subseteq> {0..1}\<rbrakk> \<Longrightarrow> f x = f y"
by (metis feqm)
define u where "u \<equiv> rightcut 0 1 0"
have lc[simp]: "leftcut 0 1 0 = 0" and u01: "0 \<le> u" "u \<le> 1"
using LR [of 0 0 1] by (auto simp: u_def)
have f0u: "\<And>x. x \<in> {0..u} \<Longrightarrow> f x = f 0"
using LR [of 0 0 1] unfolding u_def [symmetric]
by (metis \<open>leftcut 0 1 0 = 0\<close> atLeastAtMost_iff order_refl zero_le_one)
have fu1: "\<And>x. x \<in> {u<..1} \<Longrightarrow> f x \<noteq> f 0"
using LR [of 0 0 1] unfolding u_def [symmetric] by fastforce
define v where "v \<equiv> leftcut u 1 1"
have rc[simp]: "rightcut u 1 1 = 1" and v01: "u \<le> v" "v \<le> 1"
using LR [of 1 u 1] u01 by (auto simp: v_def)
have fuv: "\<And>x. x \<in> {u..<v} \<Longrightarrow> f x \<noteq> f 1"
using LR [of 1 u 1] u01 v_def by fastforce
have f0v: "\<And>x. x \<in> {0..<v} \<Longrightarrow> f x \<noteq> f 1"
by (metis f_1not0 atLeastAtMost_iff atLeastLessThan_iff f0u fuv linear)
have fv1: "\<And>x. x \<in> {v..1} \<Longrightarrow> f x = f 1"
using LR [of 1 u 1] u01 v_def by (metis atLeastAtMost_iff atLeastatMost_subset_iff order_refl rc)
define a where "a \<equiv> leftrec u v leftcut rightcut"
define b where "b \<equiv> rightrec u v leftcut rightcut"
define c where "c \<equiv> \<lambda>x. midpoint (a x) (b x)"
have a_real [simp]: "a (real j) = u" for j
using a_def leftrec_base
by (metis nonzero_mult_div_cancel_right of_nat_mult of_nat_numeral zero_neq_numeral)
have b_real [simp]: "b (real j) = v" for j
using b_def rightrec_base
by (metis nonzero_mult_div_cancel_right of_nat_mult of_nat_numeral zero_neq_numeral)
have a41: "a ((4 * real m + 1) / 2^Suc n) = a ((2 * real m + 1) / 2^n)" if "n > 0" for m n
using that a_def leftrec_41 by blast
have b41: "b ((4 * real m + 1) / 2^Suc n) =
leftcut (a ((2 * real m + 1) / 2^n))
(b ((2 * real m + 1) / 2^n))
(c ((2 * real m + 1) / 2^n))" if "n > 0" for m n
using that a_def b_def c_def rightrec_41 by blast
have a43: "a ((4 * real m + 3) / 2^Suc n) =
rightcut (a ((2 * real m + 1) / 2^n))
(b ((2 * real m + 1) / 2^n))
(c ((2 * real m + 1) / 2^n))" if "n > 0" for m n
using that a_def b_def c_def leftrec_43 by blast
have b43: "b ((4 * real m + 3) / 2^Suc n) = b ((2 * real m + 1) / 2^n)" if "n > 0" for m n
using that b_def rightrec_43 by blast
have uabv: "u \<le> a (real m / 2 ^ n) \<and> a (real m / 2 ^ n) \<le> b (real m / 2 ^ n) \<and> b (real m / 2 ^ n) \<le> v" for m n
proof (induction n arbitrary: m)
case 0
then show ?case by (simp add: v01)
next
case (Suc n p)
show ?case
proof (cases "even p")
case True
then obtain m where "p = 2*m" by (metis evenE)
then show ?thesis
by (simp add: Suc.IH)
next
case False
then obtain m where m: "p = 2*m + 1" by (metis oddE)
show ?thesis
proof (cases n)
case 0
then show ?thesis
by (simp add: a_def b_def leftrec_base rightrec_base v01)
next
case (Suc n')
then have "n > 0" by simp
have a_le_c: "a (real m / 2^n) \<le> c (real m / 2^n)" for m
unfolding c_def by (metis Suc.IH ge_midpoint_1)
have c_le_b: "c (real m / 2^n) \<le> b (real m / 2^n)" for m
unfolding c_def by (metis Suc.IH le_midpoint_1)
have c_ge_u: "c (real m / 2^n) \<ge> u" for m
using Suc.IH a_le_c order_trans by blast
have c_le_v: "c (real m / 2^n) \<le> v" for m
using Suc.IH c_le_b order_trans by blast
have a_ge_0: "0 \<le> a (real m / 2^n)" for m
using Suc.IH order_trans u01(1) by blast
have b_le_1: "b (real m / 2^n) \<le> 1" for m
using Suc.IH order_trans v01(2) by blast
have left_le: "leftcut (a ((real m) / 2^n)) (b ((real m) / 2^n)) (c ((real m) / 2^n)) \<le> c ((real m) / 2^n)" for m
by (simp add: LR a_ge_0 a_le_c b_le_1 c_le_b)
have right_ge: "rightcut (a ((real m) / 2^n)) (b ((real m) / 2^n)) (c ((real m) / 2^n)) \<ge> c ((real m) / 2^n)" for m
by (simp add: LR a_ge_0 a_le_c b_le_1 c_le_b)
show ?thesis
proof (cases "even m")
case True
then obtain r where r: "m = 2*r" by (metis evenE)
show ?thesis
using a_le_c [of "m+1"] c_le_b [of "m+1"] a_ge_0 [of "m+1"] b_le_1 [of "m+1"]
Suc.IH [of "m+1"]
apply (simp add: r m add.commute [of 1] \<open>n > 0\<close> a41 b41 del: power_Suc)
apply (auto simp: left_right [THEN conjunct1])
using order_trans [OF left_le c_le_v]
by (metis (no_types, hide_lams) add.commute mult_2 of_nat_Suc of_nat_add)
next
case False
then obtain r where r: "m = 2*r + 1" by (metis oddE)
show ?thesis
using a_le_c [of "m"] c_le_b [of "m"] a_ge_0 [of "m"] b_le_1 [of "m"]
Suc.IH [of "m+1"]
apply (simp add: r m add.commute [of 3] \<open>n > 0\<close> a43 b43 del: power_Suc)
apply (auto simp: add.commute left_right [THEN conjunct2])
using order_trans [OF c_ge_u right_ge]
apply (metis (no_types, hide_lams) mult_2 numeral_One of_nat_add of_nat_numeral)
apply (metis Suc.IH mult_2 of_nat_1 of_nat_add)
done
qed
qed
qed
qed
have a_ge_0 [simp]: "0 \<le> a(m / 2^n)" and b_le_1 [simp]: "b(m / 2^n) \<le> 1" for m::nat and n
using uabv order_trans u01 v01 by blast+
then have b_ge_0 [simp]: "0 \<le> b(m / 2^n)" and a_le_1 [simp]: "a(m / 2^n) \<le> 1" for m::nat and n
using uabv order_trans by blast+
have alec [simp]: "a(m / 2^n) \<le> c(m / 2^n)" and cleb [simp]: "c(m / 2^n) \<le> b(m / 2^n)" for m::nat and n
by (auto simp: c_def ge_midpoint_1 le_midpoint_1 uabv)
have c_ge_0 [simp]: "0 \<le> c(m / 2^n)" and c_le_1 [simp]: "c(m / 2^n) \<le> 1" for m::nat and n
using a_ge_0 alec order_trans apply blast
by (meson b_le_1 cleb order_trans)
have "\<lbrakk>d = m-n; odd j; \<bar>real i / 2^m - real j / 2^n\<bar> < 1/2 ^ n\<rbrakk>
\<Longrightarrow> (a(j / 2^n)) \<le> (c(i / 2^m)) \<and> (c(i / 2^m)) \<le> (b(j / 2^n))" for d i j m n
proof (induction d arbitrary: j n rule: less_induct)
case (less d j n)
show ?case
proof (cases "m \<le> n")
case True
have "\<bar>2^n\<bar> * \<bar>real i / 2^m - real j / 2^n\<bar> = 0"
proof (rule Ints_nonzero_abs_less1)
have "(real i * 2^n - real j * 2^m) / 2^m = (real i * 2^n) / 2^m - (real j * 2^m) / 2^m"
using diff_divide_distrib by blast
also have "... = (real i * 2 ^ (n-m)) - (real j)"
using True by (auto simp: power_diff field_simps)
also have "... \<in> \<int>"
by simp
finally have "(real i * 2^n - real j * 2^m) / 2^m \<in> \<int>" .
with True Ints_abs show "\<bar>2^n\<bar> * \<bar>real i / 2^m - real j / 2^n\<bar> \<in> \<int>"
by (fastforce simp: field_split_simps)
show "\<bar>\<bar>2^n\<bar> * \<bar>real i / 2^m - real j / 2^n\<bar>\<bar> < 1"
using less.prems by (auto simp: field_split_simps)
qed
then have "real i / 2^m = real j / 2^n"
by auto
then show ?thesis
by auto
next
case False
then have "n < m" by auto
obtain k where k: "j = Suc (2*k)"
using \<open>odd j\<close> oddE by fastforce
show ?thesis
proof (cases "n > 0")
case False
then have "a (real j / 2^n) = u"
by simp
also have "... \<le> c (real i / 2^m)"
using alec uabv by (blast intro: order_trans)
finally have ac: "a (real j / 2^n) \<le> c (real i / 2^m)" .
have "c (real i / 2^m) \<le> v"
using cleb uabv by (blast intro: order_trans)
also have "... = b (real j / 2^n)"
using False by simp
finally show ?thesis
by (auto simp: ac)
next
case True show ?thesis
proof (cases "real i / 2^m" "real j / 2^n" rule: linorder_cases)
case less
moreover have "real (4 * k + 1) / 2 ^ Suc n + 1 / (2 ^ Suc n) = real j / 2 ^ n"
using k by (force simp: field_split_simps)
moreover have "\<bar>real i / 2 ^ m - real j / 2 ^ n\<bar> < 2 / (2 ^ Suc n)"
using less.prems by simp
ultimately have closer: "\<bar>real i / 2 ^ m - real (4 * k + 1) / 2 ^ Suc n\<bar> < 1 / (2 ^ Suc n)"
using less.prems by linarith
have *: "a (real (4 * k + 1) / 2 ^ Suc n) \<le> c (real i / 2 ^ m) \<and>
c (real i / 2 ^ m) \<le> b (real (4 * k + 1) / 2 ^ Suc n)"
apply (rule less.IH [OF _ refl])
using closer \<open>n < m\<close> \<open>d = m - n\<close> apply (auto simp: field_split_simps \<open>n < m\<close> diff_less_mono2)
done
show ?thesis
using LR [of "c((2*k + 1) / 2^n)" "a((2*k + 1) / 2^n)" "b((2*k + 1) / 2^n)"]
using alec [of "2*k+1"] cleb [of "2*k+1"] a_ge_0 [of "2*k+1"] b_le_1 [of "2*k+1"]
using k a41 b41 * \<open>0 < n\<close>
apply (simp add: add.commute)
done
next
case equal then show ?thesis by simp
next
case greater
moreover have "real (4 * k + 3) / 2 ^ Suc n - 1 / (2 ^ Suc n) = real j / 2 ^ n"
using k by (force simp: field_split_simps)
moreover have "\<bar>real i / 2 ^ m - real j / 2 ^ n\<bar> < 2 * 1 / (2 ^ Suc n)"
using less.prems by simp
ultimately have closer: "\<bar>real i / 2 ^ m - real (4 * k + 3) / 2 ^ Suc n\<bar> < 1 / (2 ^ Suc n)"
using less.prems by linarith
have *: "a (real (4 * k + 3) / 2 ^ Suc n) \<le> c (real i / 2 ^ m) \<and>
c (real i / 2 ^ m) \<le> b (real (4 * k + 3) / 2 ^ Suc n)"
apply (rule less.IH [OF _ refl])
using closer \<open>n < m\<close> \<open>d = m - n\<close> apply (auto simp: field_split_simps \<open>n < m\<close> diff_less_mono2)
done
show ?thesis
using LR [of "c((2*k + 1) / 2^n)" "a((2*k + 1) / 2^n)" "b((2*k + 1) / 2^n)"]
using alec [of "2*k+1"] cleb [of "2*k+1"] a_ge_0 [of "2*k+1"] b_le_1 [of "2*k+1"]
using k a43 b43 * \<open>0 < n\<close>
apply (simp add: add.commute)
done
qed
qed
qed
qed
then have aj_le_ci: "a (real j / 2 ^ n) \<le> c (real i / 2 ^ m)"
and ci_le_bj: "c (real i / 2 ^ m) \<le> b (real j / 2 ^ n)" if "odd j" "\<bar>real i / 2^m - real j / 2^n\<bar> < 1/2 ^ n" for i j m n
using that by blast+
have close_ab: "odd m \<Longrightarrow> \<bar>a (real m / 2 ^ n) - b (real m / 2 ^ n)\<bar> \<le> 2 / 2^n" for m n
proof (induction n arbitrary: m)
case 0
with u01 v01 show ?case by auto
next
case (Suc n m)
with oddE obtain k where k: "m = Suc (2*k)" by fastforce
show ?case
proof (cases "n > 0")
case False
with u01 v01 show ?thesis
by (simp add: a_def b_def leftrec_base rightrec_base)
next
case True
show ?thesis
proof (cases "even k")
case True
then obtain j where j: "k = 2*j" by (metis evenE)
have "\<bar>a ((2 * real j + 1) / 2 ^ n) - (b ((2 * real j + 1) / 2 ^ n))\<bar> \<le> 2/2 ^ n"
proof -
have "odd (Suc k)"
using True by auto
then show ?thesis
by (metis (no_types) Groups.add_ac(2) Suc.IH j of_nat_Suc of_nat_mult of_nat_numeral)
qed
moreover have "a ((2 * real j + 1) / 2 ^ n) \<le>
leftcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n))"
using alec [of "2*j+1"] cleb [of "2*j+1"] a_ge_0 [of "2*j+1"] b_le_1 [of "2*j+1"]
by (auto simp: add.commute left_right)
moreover have "leftcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n)) \<le>
c ((2 * real j + 1) / 2 ^ n)"
using alec [of "2*j+1"] cleb [of "2*j+1"] a_ge_0 [of "2*j+1"] b_le_1 [of "2*j+1"]
by (auto simp: add.commute left_right_m)
ultimately have "\<bar>a ((2 * real j + 1) / 2 ^ n) -
leftcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n))\<bar>
\<le> 2/2 ^ Suc n"
by (simp add: c_def midpoint_def)
with j k \<open>n > 0\<close> show ?thesis
by (simp add: add.commute [of 1] a41 b41 del: power_Suc)
next
case False
then obtain j where j: "k = 2*j + 1" by (metis oddE)
have "\<bar>a ((2 * real j + 1) / 2 ^ n) - (b ((2 * real j + 1) / 2 ^ n))\<bar> \<le> 2/2 ^ n"
using Suc.IH [OF False] j by (auto simp: algebra_simps)
moreover have "c ((2 * real j + 1) / 2 ^ n) \<le>
rightcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n))"
using alec [of "2*j+1"] cleb [of "2*j+1"] a_ge_0 [of "2*j+1"] b_le_1 [of "2*j+1"]
by (auto simp: add.commute left_right_m)
moreover have "rightcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n)) \<le>
b ((2 * real j + 1) / 2 ^ n)"
using alec [of "2*j+1"] cleb [of "2*j+1"] a_ge_0 [of "2*j+1"] b_le_1 [of "2*j+1"]
by (auto simp: add.commute left_right)
ultimately have "\<bar>rightcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n)) -
b ((2 * real j + 1) / 2 ^ n)\<bar> \<le> 2/2 ^ Suc n"
by (simp add: c_def midpoint_def)
with j k \<open>n > 0\<close> show ?thesis
by (simp add: add.commute [of 3] a43 b43 del: power_Suc)
qed
qed
qed
have m1_to_3: "4 * real k - 1 = real (4 * (k-1)) + 3" if "0 < k" for k
using that by auto
have fb_eq_fa: "\<lbrakk>0 < j; 2*j < 2 ^ n\<rbrakk> \<Longrightarrow> f(b((2 * real j - 1) / 2^n)) = f(a((2 * real j + 1) / 2^n))" for n j
proof (induction n arbitrary: j)
case 0
then show ?case by auto
next
case (Suc n j) show ?case
proof (cases "n > 0")
case False
with Suc.prems show ?thesis by auto
next
case True
show ?thesis proof (cases "even j")
case True
then obtain k where k: "j = 2*k" by (metis evenE)
with \<open>0 < j\<close> have "k > 0" "2 * k < 2 ^ n"
using Suc.prems(2) k by auto
with k \<open>0 < n\<close> Suc.IH [of k] show ?thesis
apply (simp add: m1_to_3 a41 b43 del: power_Suc)
apply (subst of_nat_diff, auto)
done
next
case False
then obtain k where k: "j = 2*k + 1" by (metis oddE)
have "f (leftcut (a ((2 * k + 1) / 2^n)) (b ((2 * k + 1) / 2^n)) (c ((2 * k + 1) / 2^n)))
= f (c ((2 * k + 1) / 2^n))"
"f (c ((2 * k + 1) / 2^n))
= f (rightcut (a ((2 * k + 1) / 2^n)) (b ((2 * k + 1) / 2^n)) (c ((2 * k + 1) / 2^n)))"
using alec [of "2*k+1" n] cleb [of "2*k+1" n] a_ge_0 [of "2*k+1" n] b_le_1 [of "2*k+1" n] k
using left_right_m [of "c((2*k + 1) / 2^n)" "a((2*k + 1) / 2^n)" "b((2*k + 1) / 2^n)"]
apply (auto simp: add.commute feqm [OF order_refl] feqm [OF _ order_refl, symmetric])
done
then
show ?thesis
by (simp add: k add.commute [of 1] add.commute [of 3] a43 b41\<open>0 < n\<close> del: power_Suc)
qed
qed
qed
have f_eq_fc: "\<lbrakk>0 < j; j < 2 ^ n\<rbrakk>
\<Longrightarrow> f(b((2*j - 1) / 2 ^ (Suc n))) = f(c(j / 2^n)) \<and>
f(a((2*j + 1) / 2 ^ (Suc n))) = f(c(j / 2^n))" for n and j::nat
proof (induction n arbitrary: j)
case 0
then show ?case by auto
next
case (Suc n)
show ?case
proof (cases "even j")
case True
then obtain k where k: "j = 2*k" by (metis evenE)
then have less2n: "k < 2 ^ n"
using Suc.prems(2) by auto
have "0 < k" using \<open>0 < j\<close> k by linarith
then have m1_to_3: "real (4 * k - Suc 0) = real (4 * (k-1)) + 3"
by auto
then show ?thesis
using Suc.IH [of k] k \<open>0 < k\<close>
apply (simp add: less2n add.commute [of 1] m1_to_3 a41 b43 del: power_Suc)
apply (auto simp: of_nat_diff)
done
next
case False
then obtain k where k: "j = 2*k + 1" by (metis oddE)
with Suc.prems have "k < 2^n" by auto
show ?thesis
using alec [of "2*k+1" "Suc n"] cleb [of "2*k+1" "Suc n"] a_ge_0 [of "2*k+1" "Suc n"] b_le_1 [of "2*k+1" "Suc n"] k
using left_right_m [of "c((2*k + 1) / 2 ^ Suc n)" "a((2*k + 1) / 2 ^ Suc n)" "b((2*k + 1) / 2 ^ Suc n)"]
apply (simp add: add.commute [of 1] add.commute [of 3] m1_to_3 b41 a43 del: power_Suc)
apply (force intro: feqm)
done
qed
qed
define D01 where "D01 \<equiv> {0<..<1} \<inter> (\<Union>k m. {real m / 2^k})"
have cloD01 [simp]: "closure D01 = {0..1}"
unfolding D01_def
by (subst closure_dyadic_rationals_in_convex_set_pos_1) auto
have "uniformly_continuous_on D01 (f \<circ> c)"
proof (clarsimp simp: uniformly_continuous_on_def)
fix e::real
assume "0 < e"
have ucontf: "uniformly_continuous_on {0..1} f"
by (simp add: compact_uniformly_continuous [OF cont_f])
then obtain d where "0 < d" and d: "\<And>x x'. \<lbrakk>x \<in> {0..1}; x' \<in> {0..1}; norm (x' - x) < d\<rbrakk> \<Longrightarrow> norm (f x' - f x) < e/2"
unfolding uniformly_continuous_on_def dist_norm
by (metis \<open>0 < e\<close> less_divide_eq_numeral1(1) mult_zero_left)
obtain n where n: "1/2^n < min d 1"
by (metis \<open>0 < d\<close> divide_less_eq_1 less_numeral_extra(1) min_def one_less_numeral_iff power_one_over real_arch_pow_inv semiring_norm(76) zero_less_numeral)
with gr0I have "n > 0"
by (force simp: field_split_simps)
show "\<exists>d>0. \<forall>x\<in>D01. \<forall>x'\<in>D01. dist x' x < d \<longrightarrow> dist (f (c x')) (f (c x)) < e"
proof (intro exI ballI impI conjI)
show "(0::real) < 1/2^n" by auto
next
have dist_fc_close: "dist (f(c(real i / 2^m))) (f(c(real j / 2^n))) < e/2"
if i: "0 < i" "i < 2 ^ m" and j: "0 < j" "j < 2 ^ n" and clo: "abs(i / 2^m - j / 2^n) < 1/2 ^ n" for i j m
proof -
have abs3: "\<bar>x - a\<bar> < e \<Longrightarrow> x = a \<or> \<bar>x - (a - e/2)\<bar> < e/2 \<or> \<bar>x - (a + e/2)\<bar> < e/2" for x a e::real
by linarith
consider "i / 2 ^ m = j / 2 ^ n"
| "\<bar>i / 2 ^ m - (2 * j - 1) / 2 ^ Suc n\<bar> < 1/2 ^ Suc n"
| "\<bar>i / 2 ^ m - (2 * j + 1) / 2 ^ Suc n\<bar> < 1/2 ^ Suc n"
using abs3 [OF clo] j by (auto simp: field_simps of_nat_diff)
then show ?thesis
proof cases
case 1 with \<open>0 < e\<close> show ?thesis by auto
next
case 2
have *: "abs(a - b) \<le> 1/2 ^ n \<and> 1/2 ^ n < d \<and> a \<le> c \<and> c \<le> b \<Longrightarrow> b - c < d" for a b c
by auto
have "norm (c (real i / 2 ^ m) - b (real (2 * j - 1) / 2 ^ Suc n)) < d"
using 2 j n close_ab [of "2*j-1" "Suc n"]
using b_ge_0 [of "2*j-1" "Suc n"] b_le_1 [of "2*j-1" "Suc n"]
using aj_le_ci [of "2*j-1" i m "Suc n"]
using ci_le_bj [of "2*j-1" i m "Suc n"]
apply (simp add: divide_simps of_nat_diff del: power_Suc)
apply (auto simp: divide_simps intro!: *)
done
moreover have "f(c(j / 2^n)) = f(b ((2*j - 1) / 2 ^ (Suc n)))"
using f_eq_fc [OF j] by metis
ultimately show ?thesis
by (metis dist_norm atLeastAtMost_iff b_ge_0 b_le_1 c_ge_0 c_le_1 d)
next
case 3
have *: "abs(a - b) \<le> 1/2 ^ n \<and> 1/2 ^ n < d \<and> a \<le> c \<and> c \<le> b \<Longrightarrow> c - a < d" for a b c
by auto
have "norm (c (real i / 2 ^ m) - a (real (2 * j + 1) / 2 ^ Suc n)) < d"
using 3 j n close_ab [of "2*j+1" "Suc n"]
using b_ge_0 [of "2*j+1" "Suc n"] b_le_1 [of "2*j+1" "Suc n"]
using aj_le_ci [of "2*j+1" i m "Suc n"]
using ci_le_bj [of "2*j+1" i m "Suc n"]
apply (simp add: divide_simps of_nat_diff del: power_Suc)
apply (auto simp: divide_simps intro!: *)
done
moreover have "f(c(j / 2^n)) = f(a ((2*j + 1) / 2 ^ (Suc n)))"
using f_eq_fc [OF j] by metis
ultimately show ?thesis
by (metis dist_norm a_ge_0 atLeastAtMost_iff a_ge_0 a_le_1 c_ge_0 c_le_1 d)
qed
qed
show "dist (f (c x')) (f (c x)) < e"
if "x \<in> D01" "x' \<in> D01" "dist x' x < 1/2^n" for x x'
using that unfolding D01_def dyadics_in_open_unit_interval
proof clarsimp
fix i k::nat and m p
assume i: "0 < i" "i < 2 ^ m" and k: "0<k" "k < 2 ^ p"
assume clo: "dist (real k / 2 ^ p) (real i / 2 ^ m) < 1/2 ^ n"
obtain j::nat where "0 < j" "j < 2 ^ n"
and clo_ij: "abs(i / 2^m - j / 2^n) < 1/2 ^ n"
and clo_kj: "abs(k / 2^p - j / 2^n) < 1/2 ^ n"
proof -
have "max (2^n * i / 2^m) (2^n * k / 2^p) \<ge> 0"
by (auto simp: le_max_iff_disj)
then obtain j where "floor (max (2^n*i / 2^m) (2^n*k / 2^p)) = int j"
using zero_le_floor zero_le_imp_eq_int by blast
then have j_le: "real j \<le> max (2^n * i / 2^m) (2^n * k / 2^p)"
and less_j1: "max (2^n * i / 2^m) (2^n * k / 2^p) < real j + 1"
using floor_correct [of "max (2^n * i / 2^m) (2^n * k / 2^p)"] by linarith+
show thesis
proof (cases "j = 0")
case True
show thesis
proof
show "(1::nat) < 2 ^ n"
apply (subst one_less_power)
using \<open>n > 0\<close> by auto
show "\<bar>real i / 2 ^ m - real 1/2 ^ n\<bar> < 1/2 ^ n"
using i less_j1 by (simp add: dist_norm field_simps True)
show "\<bar>real k / 2 ^ p - real 1/2 ^ n\<bar> < 1/2 ^ n"
using k less_j1 by (simp add: dist_norm field_simps True)
qed simp
next
case False
have 1: "real j * 2 ^ m < real i * 2 ^ n"
if j: "real j * 2 ^ p \<le> real k * 2 ^ n" and k: "real k * 2 ^ m < real i * 2 ^ p"
for i k m p
proof -
have "real j * 2 ^ p * 2 ^ m \<le> real k * 2 ^ n * 2 ^ m"
using j by simp
moreover have "real k * 2 ^ m * 2 ^ n < real i * 2 ^ p * 2 ^ n"
using k by simp
ultimately have "real j * 2 ^ p * 2 ^ m < real i * 2 ^ p * 2 ^ n"
by (simp only: mult_ac)
then show ?thesis
by simp
qed
have 2: "real j * 2 ^ m < 2 ^ m + real i * 2 ^ n"
if j: "real j * 2 ^ p \<le> real k * 2 ^ n" and k: "real k * (2 ^ m * 2 ^ n) < 2 ^ m * 2 ^ p + real i * (2 ^ n * 2 ^ p)"
for i k m p
proof -
have "real j * 2 ^ p * 2 ^ m \<le> real k * (2 ^ m * 2 ^ n)"
using j by simp
also have "... < 2 ^ m * 2 ^ p + real i * (2 ^ n * 2 ^ p)"
by (rule k)
finally have "(real j * 2 ^ m) * 2 ^ p < (2 ^ m + real i * 2 ^ n) * 2 ^ p"
by (simp add: algebra_simps)
then show ?thesis
by simp
qed
have 3: "real j * 2 ^ p < 2 ^ p + real k * 2 ^ n"
if j: "real j * 2 ^ m \<le> real i * 2 ^ n" and i: "real i * 2 ^ p \<le> real k * 2 ^ m"
proof -
have "real j * 2 ^ m * 2 ^ p \<le> real i * 2 ^ n * 2 ^ p"
using j by simp
moreover have "real i * 2 ^ p * 2 ^ n \<le> real k * 2 ^ m * 2 ^ n"
using i by simp
ultimately have "real j * 2 ^ m * 2 ^ p \<le> real k * 2 ^ m * 2 ^ n"
by (simp only: mult_ac)
then have "real j * 2 ^ p \<le> real k * 2 ^ n"
by simp
also have "... < 2 ^ p + real k * 2 ^ n"
by auto
finally show ?thesis by simp
qed
show ?thesis
proof
have "real j < 2 ^ n"
using j_le i k
apply (auto simp: le_max_iff_disj simp del: of_nat_less_numeral_power_cancel_iff
elim!: le_less_trans)
apply (auto simp: field_simps)
done
then show "j < 2 ^ n"
by auto
show "\<bar>real i / 2 ^ m - real j / 2 ^ n\<bar> < 1/2 ^ n"
using clo less_j1 j_le
apply (auto simp: le_max_iff_disj field_split_simps dist_norm)
apply (auto simp: algebra_simps abs_if split: if_split_asm dest: 1 2)
done
show "\<bar>real k / 2 ^ p - real j / 2 ^ n\<bar> < 1/2 ^ n"
using clo less_j1 j_le
apply (auto simp: le_max_iff_disj field_split_simps dist_norm)
apply (auto simp: algebra_simps not_less abs_if split: if_split_asm dest: 3 2)
done
qed (use False in simp)
qed
qed
show "dist (f (c (real k / 2 ^ p))) (f (c (real i / 2 ^ m))) < e"
proof (rule dist_triangle_half_l)
show "dist (f (c (real k / 2 ^ p))) (f(c(j / 2^n))) < e/2"
apply (rule dist_fc_close)
using \<open>0 < j\<close> \<open>j < 2 ^ n\<close> k clo_kj by auto
show "dist (f (c (real i / 2 ^ m))) (f (c (real j / 2 ^ n))) < e/2"
apply (rule dist_fc_close)
using \<open>0 < j\<close> \<open>j < 2 ^ n\<close> i clo_ij by auto
qed
qed
qed
qed
then obtain h where ucont_h: "uniformly_continuous_on {0..1} h"
and fc_eq: "\<And>x. x \<in> D01 \<Longrightarrow> (f \<circ> c) x = h x"
proof (rule uniformly_continuous_on_extension_on_closure [of D01 "f \<circ> c"])
qed (use closure_subset [of D01] in \<open>auto intro!: that\<close>)
then have cont_h: "continuous_on {0..1} h"
using uniformly_continuous_imp_continuous by blast
have h_eq: "h (real k / 2 ^ m) = f (c (real k / 2 ^ m))" if "0 < k" "k < 2^m" for k m
using fc_eq that by (force simp: D01_def)
have "h ` {0..1} = f ` {0..1}"
proof
have "h ` (closure D01) \<subseteq> f ` {0..1}"
proof (rule image_closure_subset)
show "continuous_on (closure D01) h"
using cont_h by simp
show "closed (f ` {0..1})"
using compact_continuous_image [OF cont_f] compact_imp_closed by blast
show "h ` D01 \<subseteq> f ` {0..1}"
by (force simp: dyadics_in_open_unit_interval D01_def h_eq)
qed
with cloD01 show "h ` {0..1} \<subseteq> f ` {0..1}" by simp
have a12 [simp]: "a (1/2) = u"
by (metis a_def leftrec_base numeral_One of_nat_numeral)
have b12 [simp]: "b (1/2) = v"
by (metis b_def rightrec_base numeral_One of_nat_numeral)
have "f ` {0..1} \<subseteq> closure(h ` D01)"
proof (clarsimp simp: closure_approachable dyadics_in_open_unit_interval D01_def)
fix x e::real
assume "0 \<le> x" "x \<le> 1" "0 < e"
have ucont_f: "uniformly_continuous_on {0..1} f"
using compact_uniformly_continuous cont_f by blast
then obtain \<delta> where "\<delta> > 0"
and \<delta>: "\<And>x x'. \<lbrakk>x \<in> {0..1}; x' \<in> {0..1}; dist x' x < \<delta>\<rbrakk> \<Longrightarrow> norm (f x' - f x) < e"
using \<open>0 < e\<close> by (auto simp: uniformly_continuous_on_def dist_norm)
have *: "\<exists>m::nat. \<exists>y. odd m \<and> 0 < m \<and> m < 2 ^ n \<and> y \<in> {a(m / 2^n) .. b(m / 2^n)} \<and> f y = f x"
if "n \<noteq> 0" for n
using that
proof (induction n)
case 0 then show ?case by auto
next
case (Suc n)
show ?case
proof (cases "n=0")
case True
consider "x \<in> {0..u}" | "x \<in> {u..v}" | "x \<in> {v..1}"
using \<open>0 \<le> x\<close> \<open>x \<le> 1\<close> by force
then have "\<exists>y\<ge>a (real 1/2). y \<le> b (real 1/2) \<and> f y = f x"
proof cases
case 1
then show ?thesis
apply (rule_tac x=u in exI)
using uabv [of 1 1] f0u [of u] f0u [of x] by auto
next
case 2
then show ?thesis
by (rule_tac x=x in exI) auto
next
case 3
then show ?thesis
apply (rule_tac x=v in exI)
using uabv [of 1 1] fv1 [of v] fv1 [of x] by auto
qed
with \<open>n=0\<close> show ?thesis
by (rule_tac x=1 in exI) auto
next
case False
with Suc obtain m y
where "odd m" "0 < m" and mless: "m < 2 ^ n"
and y: "y \<in> {a (real m / 2 ^ n)..b (real m / 2 ^ n)}" and feq: "f y = f x"
by metis
then obtain j where j: "m = 2*j + 1" by (metis oddE)
consider "y \<in> {a((2*j + 1) / 2^n) .. b((4*j + 1) / 2 ^ (Suc n))}"
| "y \<in> {b((4*j + 1) / 2 ^ (Suc n)) .. a((4*j + 3) / 2 ^ (Suc n))}"
| "y \<in> {a((4*j + 3) / 2 ^ (Suc n)) .. b((2*j + 1) / 2^n)}"
using y j by force
then show ?thesis
proof cases
case 1
then show ?thesis
apply (rule_tac x="4*j + 1" in exI)
apply (rule_tac x=y in exI)
using mless j \<open>n \<noteq> 0\<close>
apply (simp add: feq a41 b41 add.commute [of 1] del: power_Suc)
apply (simp add: algebra_simps)
done
next
case 2
show ?thesis
apply (rule_tac x="4*j + 1" in exI)
apply (rule_tac x="b((4*j + 1) / 2 ^ (Suc n))" in exI)
using mless \<open>n \<noteq> 0\<close> 2 j
using alec [of "2*j+1" n] cleb [of "2*j+1" n] a_ge_0 [of "2*j+1" n] b_le_1 [of "2*j+1" n]
using left_right [of "c((2*j + 1) / 2^n)" "a((2*j + 1) / 2^n)" "b((2*j + 1) / 2^n)"]
apply (simp add: a41 b41 a43 b43 add.commute [of 1] add.commute [of 3] del: power_Suc)
apply (auto simp: feq [symmetric] intro: f_eqI)
done
next
case 3
then show ?thesis
apply (rule_tac x="4*j + 3" in exI)
apply (rule_tac x=y in exI)
using mless j \<open>n \<noteq> 0\<close>
apply (simp add: feq a43 b43 del: power_Suc)
apply (simp add: algebra_simps)
done
qed
qed
qed
obtain n where n: "1/2^n < min (\<delta> / 2) 1"
by (metis \<open>0 < \<delta>\<close> divide_less_eq_1 less_numeral_extra(1) min_less_iff_conj one_less_numeral_iff power_one_over real_arch_pow_inv semiring_norm(76) zero_less_divide_iff zero_less_numeral)
with gr0I have "n \<noteq> 0"
by fastforce
with * obtain m::nat and y
where "odd m" "0 < m" and mless: "m < 2 ^ n"
and y: "y \<in> {a(m / 2^n) .. b(m / 2^n)}" and feq: "f x = f y"
by metis
then have "0 \<le> y" "y \<le> 1"
by (metis atLeastAtMost_iff a_ge_0 b_le_1 order.trans)+
moreover have "y < \<delta> + c (real m / 2 ^ n)" "c (real m / 2 ^ n) < \<delta> + y"
using y apply simp_all
using alec [of m n] cleb [of m n] n field_sum_of_halves close_ab [OF \<open>odd m\<close>, of n]
by linarith+
moreover note \<open>0 < m\<close> mless \<open>0 \<le> x\<close> \<open>x \<le> 1\<close>
ultimately show "\<exists>k. \<exists>m\<in>{0<..<2 ^ k}. dist (h (real m / 2 ^ k)) (f x) < e"
apply (rule_tac x=n in exI)
apply (rule_tac x=m in bexI)
apply (auto simp: dist_norm h_eq feq \<delta>)
done
qed
also have "... \<subseteq> h ` {0..1}"
apply (rule closure_minimal)
using compact_continuous_image [OF cont_h] compact_imp_closed by (auto simp: D01_def)
finally show "f ` {0..1} \<subseteq> h ` {0..1}" .
qed
moreover have "inj_on h {0..1}"
proof -
have "u < v"
by (metis atLeastAtMost_iff f0u f_1not0 fv1 order.not_eq_order_implies_strict u01(1) u01(2) v01(1))
have f_not_fu: "\<And>x. \<lbrakk>u < x; x \<le> v\<rbrakk> \<Longrightarrow> f x \<noteq> f u"
by (metis atLeastAtMost_iff f0u fu1 greaterThanAtMost_iff order_refl order_trans u01(1) v01(2))
have f_not_fv: "\<And>x. \<lbrakk>u \<le> x; x < v\<rbrakk> \<Longrightarrow> f x \<noteq> f v"
by (metis atLeastAtMost_iff order_refl order_trans v01(2) atLeastLessThan_iff fuv fv1)
have a_less_b:
"a(j / 2^n) < b(j / 2^n) \<and>
(\<forall>x. a(j / 2^n) < x \<longrightarrow> x \<le> b(j / 2^n) \<longrightarrow> f x \<noteq> f(a(j / 2^n))) \<and>
(\<forall>x. a(j / 2^n) \<le> x \<longrightarrow> x < b(j / 2^n) \<longrightarrow> f x \<noteq> f(b(j / 2^n)))" for n and j::nat
proof (induction n arbitrary: j)
case 0 then show ?case
by (simp add: \<open>u < v\<close> f_not_fu f_not_fv)
next
case (Suc n j) show ?case
proof (cases "n > 0")
case False then show ?thesis
by (auto simp: a_def b_def leftrec_base rightrec_base \<open>u < v\<close> f_not_fu f_not_fv)
next
case True show ?thesis
proof (cases "even j")
case True
with \<open>0 < n\<close> Suc.IH show ?thesis
by (auto elim!: evenE)
next
case False
then obtain k where k: "j = 2*k + 1" by (metis oddE)
then show ?thesis
proof (cases "even k")
case True
then obtain m where m: "k = 2*m" by (metis evenE)
have fleft: "f (leftcut (a ((2*m + 1) / 2^n)) (b ((2*m + 1) / 2^n)) (c ((2*m + 1) / 2^n))) =
f (c((2*m + 1) / 2^n))"
using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n]
using left_right_m [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"]
by (auto intro: f_eqI)
show ?thesis
proof (intro conjI impI notI allI)
have False if "b (real j / 2 ^ Suc n) \<le> a (real j / 2 ^ Suc n)"
proof -
have "f (c ((1 + real m * 2) / 2 ^ n)) = f (a ((1 + real m * 2) / 2 ^ n))"
using k m \<open>0 < n\<close> fleft that a41 [of n m] b41 [of n m]
using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n]
using left_right [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"]
by (auto simp: algebra_simps)
moreover have "a (real (1 + m * 2) / 2 ^ n) < c (real (1 + m * 2) / 2 ^ n)"
using Suc.IH [of "1 + m * 2"] by (simp add: c_def midpoint_def)
moreover have "c (real (1 + m * 2) / 2 ^ n) \<le> b (real (1 + m * 2) / 2 ^ n)"
using cleb by blast
ultimately show ?thesis
using Suc.IH [of "1 + m * 2"] by force
qed
then show "a (real j / 2 ^ Suc n) < b (real j / 2 ^ Suc n)" by force
next
fix x
assume "a (real j / 2 ^ Suc n) < x" "x \<le> b (real j / 2 ^ Suc n)" "f x = f (a (real j / 2 ^ Suc n))"
then show False
using Suc.IH [of "1 + m * 2", THEN conjunct2, THEN conjunct1]
using k m \<open>0 < n\<close> a41 [of n m] b41 [of n m]
using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n]
using left_right_m [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"]
by (auto simp: algebra_simps)
next
fix x
assume "a (real j / 2 ^ Suc n) \<le> x" "x < b (real j / 2 ^ Suc n)" "f x = f (b (real j / 2 ^ Suc n))"
then show False
using k m \<open>0 < n\<close> a41 [of n m] b41 [of n m] fleft left_neq
using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n]
by (auto simp: algebra_simps)
qed
next
case False
with oddE obtain m where m: "k = Suc (2*m)" by fastforce
have fright: "f (rightcut (a ((2*m + 1) / 2^n)) (b ((2*m + 1) / 2^n)) (c ((2*m + 1) / 2^n))) = f (c((2*m + 1) / 2^n))"
using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n]
using left_right_m [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"]
by (auto intro: f_eqI [OF _ order_refl])
show ?thesis
proof (intro conjI impI notI allI)
have False if "b (real j / 2 ^ Suc n) \<le> a (real j / 2 ^ Suc n)"
proof -
have "f (c ((1 + real m * 2) / 2 ^ n)) = f (b ((1 + real m * 2) / 2 ^ n))"
using k m \<open>0 < n\<close> fright that a43 [of n m] b43 [of n m]
using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n]
using left_right [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"]
by (auto simp: algebra_simps)
moreover have "a (real (1 + m * 2) / 2 ^ n) \<le> c (real (1 + m * 2) / 2 ^ n)"
using alec by blast
moreover have "c (real (1 + m * 2) / 2 ^ n) < b (real (1 + m * 2) / 2 ^ n)"
using Suc.IH [of "1 + m * 2"] by (simp add: c_def midpoint_def)
ultimately show ?thesis
using Suc.IH [of "1 + m * 2"] by force
qed
then show "a (real j / 2 ^ Suc n) < b (real j / 2 ^ Suc n)" by force
next
fix x
assume "a (real j / 2 ^ Suc n) < x" "x \<le> b (real j / 2 ^ Suc n)" "f x = f (a (real j / 2 ^ Suc n))"
then show False
using k m \<open>0 < n\<close> a43 [of n m] b43 [of n m] fright right_neq
using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n]
by (auto simp: algebra_simps)
next
fix x
assume "a (real j / 2 ^ Suc n) \<le> x" "x < b (real j / 2 ^ Suc n)" "f x = f (b (real j / 2 ^ Suc n))"
then show False
using Suc.IH [of "1 + m * 2", THEN conjunct2, THEN conjunct2]
using k m \<open>0 < n\<close> a43 [of n m] b43 [of n m]
using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n]
using left_right_m [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"]
by (auto simp: algebra_simps fright simp del: power_Suc)
qed
qed
qed
qed
qed
have c_gt_0 [simp]: "0 < c(m / 2^n)" and c_less_1 [simp]: "c(m / 2^n) < 1" for m::nat and n
using a_less_b [of m n] apply (simp_all add: c_def midpoint_def)
using a_ge_0 [of m n] b_le_1 [of m n] apply linarith+
done
have approx: "\<exists>j n. odd j \<and> n \<noteq> 0 \<and>
real i / 2^m \<le> real j / 2^n \<and>
real j / 2^n \<le> real k / 2^p \<and>
\<bar>real i / 2 ^ m - real j / 2 ^ n\<bar> < 1/2^n \<and>
\<bar>real k / 2 ^ p - real j / 2 ^ n\<bar> < 1/2^n"
if "0 < i" "i < 2 ^ m" "0 < k" "k < 2 ^ p" "i / 2^m < k / 2^p" "m + p = N" for N m p i k
using that
proof (induction N arbitrary: m p i k rule: less_induct)
case (less N)
then consider "i / 2^m \<le> 1/2" "1/2 \<le> k / 2^p" | "k / 2^p < 1/2" | "k / 2^p \<ge> 1/2" "1/2 < i / 2^m"
by linarith
then show ?case
proof cases
case 1
with less.prems show ?thesis
by (rule_tac x=1 in exI)+ (fastforce simp: field_split_simps)
next
case 2 show ?thesis
proof (cases m)
case 0 with less.prems show ?thesis
by auto
next
case (Suc m') show ?thesis
proof (cases p)
case 0 with less.prems show ?thesis by auto
next
case (Suc p')
have False if "real i * 2 ^ p' < real k * 2 ^ m'" "k < 2 ^ p'" "2 ^ m' \<le> i"
proof -
have "real k * 2 ^ m' < 2 ^ p' * 2 ^ m'"
using that by simp
then have "real i * 2 ^ p' < 2 ^ p' * 2 ^ m'"
using that by linarith
with that show ?thesis by simp
qed
then show ?thesis
using less.IH [of "m'+p'" i m' k p'] less.prems \<open>m = Suc m'\<close> 2 Suc
apply atomize
apply (force simp: field_split_simps)
done
qed
qed
next
case 3 show ?thesis
proof (cases m)
case 0 with less.prems show ?thesis
by auto
next
case (Suc m') show ?thesis
proof (cases p)
case 0 with less.prems show ?thesis by auto
next
case (Suc p')
then show ?thesis
using less.IH [of "m'+p'" "i - 2^m'" m' "k - 2 ^ p'" p'] less.prems \<open>m = Suc m'\<close> Suc 3
apply atomize
apply (auto simp: field_simps of_nat_diff)
apply (rule_tac x="2 ^ n + j" in exI, simp)
apply (rule_tac x="Suc n" in exI)
apply (auto simp: field_simps)
done
qed
qed
qed
qed
have clec: "c(real i / 2^m) \<le> c(real j / 2^n)"
if i: "0 < i" "i < 2 ^ m" and j: "0 < j" "j < 2 ^ n" and ij: "i / 2^m < j / 2^n" for m i n j
proof -
obtain j' n' where "odd j'" "n' \<noteq> 0"
and i_le_j: "real i / 2 ^ m \<le> real j' / 2 ^ n'"
and j_le_j: "real j' / 2 ^ n' \<le> real j / 2 ^ n"
and clo_ij: "\<bar>real i / 2 ^ m - real j' / 2 ^ n'\<bar> < 1/2 ^ n'"
and clo_jj: "\<bar>real j / 2 ^ n - real j' / 2 ^ n'\<bar> < 1/2 ^ n'"
using approx [of i m j n "m+n"] that i j ij by auto
with oddE obtain q where q: "j' = Suc (2*q)" by fastforce
have "c (real i / 2 ^ m) \<le> c((2*q + 1) / 2^n')"
proof (cases "i / 2^m = (2*q + 1) / 2^n'")
case True then show ?thesis by simp
next
case False
with i_le_j q have less: "i / 2^m < (2*q + 1) / 2^n'"
by auto
have *: "\<lbrakk>i < q; abs(i - q) < s*2; q = r + s\<rbrakk> \<Longrightarrow> abs(i - r) < s" for i q s r::real
by auto
have "c(i / 2^m) \<le> b(real(4 * q + 1) / 2 ^ (Suc n'))"
apply (rule ci_le_bj, force)
apply (rule * [OF less])
using i_le_j clo_ij q apply (auto simp: field_split_simps)
done
then show ?thesis
using alec [of "2*q+1" n'] cleb [of "2*q+1" n'] a_ge_0 [of "2*q+1" n'] b_le_1 [of "2*q+1" n'] b41 [of n' q] \<open>n' \<noteq> 0\<close>
using left_right_m [of "c((2*q + 1) / 2^n')" "a((2*q + 1) / 2^n')" "b((2*q + 1) / 2^n')"]
by (auto simp: algebra_simps)
qed
also have "... \<le> c(real j / 2^n)"
proof (cases "j / 2^n = (2*q + 1) / 2^n'")
case True
then show ?thesis by simp
next
case False
with j_le_j q have less: "(2*q + 1) / 2^n' < j / 2^n"
by auto
have *: "\<lbrakk>q < i; abs(i - q) < s*2; r = q + s\<rbrakk> \<Longrightarrow> abs(i - r) < s" for i q s r::real
by auto
have "a(real(4*q + 3) / 2 ^ (Suc n')) \<le> c(j / 2^n)"
apply (rule aj_le_ci, force)
apply (rule * [OF less])
using j_le_j clo_jj q apply (auto simp: field_split_simps)
done
then show ?thesis
using alec [of "2*q+1" n'] cleb [of "2*q+1" n'] a_ge_0 [of "2*q+1" n'] b_le_1 [of "2*q+1" n'] a43 [of n' q] \<open>n' \<noteq> 0\<close>
using left_right_m [of "c((2*q + 1) / 2^n')" "a((2*q + 1) / 2^n')" "b((2*q + 1) / 2^n')"]
by (auto simp: algebra_simps)
qed
finally show ?thesis .
qed
have "x = y" if "0 \<le> x" "x \<le> 1" "0 \<le> y" "y \<le> 1" "h x = h y" for x y
using that
proof (induction x y rule: linorder_class.linorder_less_wlog)
case (less x1 x2)
obtain m n where m: "0 < m" "m < 2 ^ n"
and x12: "x1 < m / 2^n" "m / 2^n < x2"
and neq: "h x1 \<noteq> h (real m / 2^n)"
proof -
have "(x1 + x2) / 2 \<in> closure D01"
using cloD01 less.hyps less.prems by auto
with less obtain y where "y \<in> D01" and dist_y: "dist y ((x1 + x2) / 2) < (x2 - x1) / 64"
unfolding closure_approachable
by (metis diff_gt_0_iff_gt less_divide_eq_numeral1(1) mult_zero_left)
obtain m n where m: "0 < m" "m < 2 ^ n"
and clo: "\<bar>real m / 2 ^ n - (x1 + x2) / 2\<bar> < (x2 - x1) / 64"
and n: "1/2^n < (x2 - x1) / 128"
proof -
have "min 1 ((x2 - x1) / 128) > 0" "1/2 < (1::real)"
using less by auto
then obtain N where N: "1/2^N < min 1 ((x2 - x1) / 128)"
by (metis power_one_over real_arch_pow_inv)
then have "N > 0"
using less_divide_eq_1 by force
obtain p q where p: "p < 2 ^ q" "p \<noteq> 0" and yeq: "y = real p / 2 ^ q"
using \<open>y \<in> D01\<close> by (auto simp: zero_less_divide_iff D01_def)
show ?thesis
proof
show "0 < 2^N * p"
using p by auto
show "2 ^ N * p < 2 ^ (N+q)"
by (simp add: p power_add)
have "\<bar>real (2 ^ N * p) / 2 ^ (N + q) - (x1 + x2) / 2\<bar> = \<bar>real p / 2 ^ q - (x1 + x2) / 2\<bar>"
by (simp add: power_add)
also have "... = \<bar>y - (x1 + x2) / 2\<bar>"
by (simp add: yeq)
also have "... < (x2 - x1) / 64"
using dist_y by (simp add: dist_norm)
finally show "\<bar>real (2 ^ N * p) / 2 ^ (N + q) - (x1 + x2) / 2\<bar> < (x2 - x1) / 64" .
have "(1::real) / 2 ^ (N + q) \<le> 1/2^N"
by (simp add: field_simps)
also have "... < (x2 - x1) / 128"
using N by force
finally show "1/2 ^ (N + q) < (x2 - x1) / 128" .
qed
qed
obtain m' n' m'' n'' where "0 < m'" "m' < 2 ^ n'" "x1 < m' / 2^n'" "m' / 2^n' < x2"
and "0 < m''" "m'' < 2 ^ n''" "x1 < m'' / 2^n''" "m'' / 2^n'' < x2"
and neq: "h (real m'' / 2^n'') \<noteq> h (real m' / 2^n')"
proof
show "0 < Suc (2*m)"
by simp
show m21: "Suc (2*m) < 2 ^ Suc n"
using m by auto
show "x1 < real (Suc (2 * m)) / 2 ^ Suc n"
using clo by (simp add: field_simps abs_if split: if_split_asm)
show "real (Suc (2 * m)) / 2 ^ Suc n < x2"
using n clo by (simp add: field_simps abs_if split: if_split_asm)
show "0 < 4*m + 3"
by simp
have "m+1 \<le> 2 ^ n"
using m by simp
then have "4 * (m+1) \<le> 4 * (2 ^ n)"
by simp
then show m43: "4*m + 3 < 2 ^ (n+2)"
by (simp add: algebra_simps)
show "x1 < real (4 * m + 3) / 2 ^ (n + 2)"
using clo by (simp add: field_simps abs_if split: if_split_asm)
show "real (4 * m + 3) / 2 ^ (n + 2) < x2"
using n clo by (simp add: field_simps abs_if split: if_split_asm)
have c_fold: "midpoint (a ((2 * real m + 1) / 2 ^ Suc n)) (b ((2 * real m + 1) / 2 ^ Suc n)) = c ((2 * real m + 1) / 2 ^ Suc n)"
by (simp add: c_def)
define R where "R \<equiv> rightcut (a ((2 * real m + 1) / 2 ^ Suc n)) (b ((2 * real m + 1) / 2 ^ Suc n)) (c ((2 * real m + 1) / 2 ^ Suc n))"
have "R < b ((2 * real m + 1) / 2 ^ Suc n)"
unfolding R_def using a_less_b [of "4*m + 3" "n+2"] a43 [of "Suc n" m] b43 [of "Suc n" m]
by simp
then have Rless: "R < midpoint R (b ((2 * real m + 1) / 2 ^ Suc n))"
by (simp add: midpoint_def)
have midR_le: "midpoint R (b ((2 * real m + 1) / 2 ^ Suc n)) \<le> b ((2 * real m + 1) / (2 * 2 ^ n))"
using \<open>R < b ((2 * real m + 1) / 2 ^ Suc n)\<close>
by (simp add: midpoint_def)
have "(real (Suc (2 * m)) / 2 ^ Suc n) \<in> D01" "real (4 * m + 3) / 2 ^ (n + 2) \<in> D01"
by (simp_all add: D01_def m21 m43 del: power_Suc of_nat_Suc of_nat_add add_2_eq_Suc') blast+
then show "h (real (4 * m + 3) / 2 ^ (n + 2)) \<noteq> h (real (Suc (2 * m)) / 2 ^ Suc n)"
using a_less_b [of "4*m + 3" "n+2", THEN conjunct1]
using a43 [of "Suc n" m] b43 [of "Suc n" m]
using alec [of "2*m+1" "Suc n"] cleb [of "2*m+1" "Suc n"] a_ge_0 [of "2*m+1" "Suc n"] b_le_1 [of "2*m+1" "Suc n"]
apply (simp add: fc_eq [symmetric] c_def del: power_Suc)
apply (simp only: add.commute [of 1] c_fold R_def [symmetric])
apply (rule right_neq)
using Rless apply (simp add: R_def)
apply (rule midR_le, auto)
done
qed
then show ?thesis by (metis that)
qed
have m_div: "0 < m / 2^n" "m / 2^n < 1"
using m by (auto simp: field_split_simps)
have closure0m: "{0..m / 2^n} = closure ({0<..< m / 2^n} \<inter> (\<Union>k m. {real m / 2 ^ k}))"
by (subst closure_dyadic_rationals_in_convex_set_pos_1, simp_all add: not_le m)
have closurem1: "{m / 2^n .. 1} = closure ({m / 2^n <..< 1} \<inter> (\<Union>k m. {real m / 2 ^ k}))"
apply (subst closure_dyadic_rationals_in_convex_set_pos_1; simp add: not_le m)
using \<open>0 < real m / 2 ^ n\<close> by linarith
have cont_h': "continuous_on (closure ({u<..<v} \<inter> (\<Union>k m. {real m / 2 ^ k}))) h"
if "0 \<le> u" "v \<le> 1" for u v
apply (rule continuous_on_subset [OF cont_h])
apply (rule closure_minimal [OF subsetI])
using that apply auto
done
have closed_f': "closed (f ` {u..v})" if "0 \<le> u" "v \<le> 1" for u v
by (metis compact_continuous_image cont_f compact_interval atLeastatMost_subset_iff
compact_imp_closed continuous_on_subset that)
have less_2I: "\<And>k i. real i / 2 ^ k < 1 \<Longrightarrow> i < 2 ^ k"
by simp
have "h ` ({0<..<m / 2 ^ n} \<inter> (\<Union>q p. {real p / 2 ^ q})) \<subseteq> f ` {0..c (m / 2 ^ n)}"
proof clarsimp
fix p q
assume p: "0 < real p / 2 ^ q" "real p / 2 ^ q < real m / 2 ^ n"
then have [simp]: "0 < p" "p < 2 ^ q"
apply (simp add: field_split_simps)
apply (blast intro: p less_2I m_div less_trans)
done
have "f (c (real p / 2 ^ q)) \<in> f ` {0..c (real m / 2 ^ n)}"
by (auto simp: clec p m)
then show "h (real p / 2 ^ q) \<in> f ` {0..c (real m / 2 ^ n)}"
by (simp add: h_eq)
qed
then have "h ` {0 .. m / 2^n} \<subseteq> f ` {0 .. c(m / 2^n)}"
apply (subst closure0m)
apply (rule image_closure_subset [OF cont_h' closed_f'])
using m_div apply auto
done
then have hx1: "h x1 \<in> f ` {0 .. c(m / 2^n)}"
using x12 less.prems(1) by auto
then obtain t1 where t1: "h x1 = f t1" "0 \<le> t1" "t1 \<le> c (m / 2 ^ n)"
by auto
have "h ` ({m / 2 ^ n<..<1} \<inter> (\<Union>q p. {real p / 2 ^ q})) \<subseteq> f ` {c (m / 2 ^ n)..1}"
proof clarsimp
fix p q
assume p: "real m / 2 ^ n < real p / 2 ^ q" and [simp]: "p < 2 ^ q"
then have [simp]: "0 < p"
using gr_zeroI m_div by fastforce
have "f (c (real p / 2 ^ q)) \<in> f ` {c (m / 2 ^ n)..1}"
by (auto simp: clec p m)
then show "h (real p / 2 ^ q) \<in> f ` {c (real m / 2 ^ n)..1}"
by (simp add: h_eq)
qed
then have "h ` {m / 2^n .. 1} \<subseteq> f ` {c(m / 2^n) .. 1}"
apply (subst closurem1)
apply (rule image_closure_subset [OF cont_h' closed_f'])
using m apply auto
done
then have hx2: "h x2 \<in> f ` {c(m / 2^n)..1}"
using x12 less.prems by auto
then obtain t2 where t2: "h x2 = f t2" "c (m / 2 ^ n) \<le> t2" "t2 \<le> 1"
by auto
with t1 less neq have False
using conn [of "h x2", unfolded is_interval_connected_1 [symmetric] is_interval_1, rule_format, of t1 t2 "c(m / 2^n)"]
by (simp add: h_eq m)
then show ?case by blast
qed auto
then show ?thesis
by (auto simp: inj_on_def)
qed
ultimately have "{0..1::real} homeomorphic f ` {0..1}"
using homeomorphic_compact [OF _ cont_h] by blast
then show ?thesis
using homeomorphic_sym by blast
qed
theorem path_contains_arc:
fixes p :: "real \<Rightarrow> 'a::{complete_space,real_normed_vector}"
assumes "path p" and a: "pathstart p = a" and b: "pathfinish p = b" and "a \<noteq> b"
obtains q where "arc q" "path_image q \<subseteq> path_image p" "pathstart q = a" "pathfinish q = b"
proof -
have ucont_p: "uniformly_continuous_on {0..1} p"
using \<open>path p\<close> unfolding path_def
by (metis compact_Icc compact_uniformly_continuous)
define \<phi> where "\<phi> \<equiv> \<lambda>S. S \<subseteq> {0..1} \<and> 0 \<in> S \<and> 1 \<in> S \<and>
(\<forall>x \<in> S. \<forall>y \<in> S. open_segment x y \<inter> S = {} \<longrightarrow> p x = p y)"
obtain T where "closed T" "\<phi> T" and T: "\<And>U. \<lbrakk>closed U; \<phi> U\<rbrakk> \<Longrightarrow> \<not> (U \<subset> T)"
proof (rule Brouwer_reduction_theorem_gen [of "{0..1}" \<phi>])
have *: "{x<..<y} \<inter> {0..1} = {x<..<y}" if "0 \<le> x" "y \<le> 1" "x \<le> y" for x y::real
using that by auto
show "\<phi> {0..1}"
by (auto simp: \<phi>_def open_segment_eq_real_ivl *)
show "\<phi> (\<Inter>(F ` UNIV))"
if "\<And>n. closed (F n)" and \<phi>: "\<And>n. \<phi> (F n)" and Fsub: "\<And>n. F (Suc n) \<subseteq> F n" for F
proof -
have F01: "\<And>n. F n \<subseteq> {0..1} \<and> 0 \<in> F n \<and> 1 \<in> F n"
and peq: "\<And>n x y. \<lbrakk>x \<in> F n; y \<in> F n; open_segment x y \<inter> F n = {}\<rbrakk> \<Longrightarrow> p x = p y"
by (metis \<phi> \<phi>_def)+
have pqF: False if "\<forall>u. x \<in> F u" "\<forall>x. y \<in> F x" "open_segment x y \<inter> (\<Inter>x. F x) = {}" and neg: "p x \<noteq> p y"
for x y
using that
proof (induction x y rule: linorder_class.linorder_less_wlog)
case (less x y)
have xy: "x \<in> {0..1}" "y \<in> {0..1}"
by (metis less.prems subsetCE F01)+
have "norm(p x - p y) / 2 > 0"
using less by auto
then obtain e where "e > 0"
and e: "\<And>u v. \<lbrakk>u \<in> {0..1}; v \<in> {0..1}; dist v u < e\<rbrakk> \<Longrightarrow> dist (p v) (p u) < norm(p x - p y) / 2"
by (metis uniformly_continuous_onE [OF ucont_p])
have minxy: "min e (y - x) < (y - x) * (3 / 2)"
by (subst min_less_iff_disj) (simp add: less)
obtain w z where "w < z" and w: "w \<in> {x<..<y}" and z: "z \<in> {x<..<y}"
and wxe: "norm(w - x) < e" and zye: "norm(z - y) < e"
apply (rule_tac w = "x + (min e (y - x) / 3)" and z = "y - (min e (y - x) / 3)" in that)
using minxy \<open>0 < e\<close> less by simp_all
have Fclo: "\<And>T. T \<in> range F \<Longrightarrow> closed T"
by (metis \<open>\<And>n. closed (F n)\<close> image_iff)
have eq: "{w..z} \<inter> \<Inter>(F ` UNIV) = {}"
using less w z apply (auto simp: open_segment_eq_real_ivl)
by (metis (no_types, hide_lams) INT_I IntI empty_iff greaterThanLessThan_iff not_le order.trans)
then obtain K where "finite K" and K: "{w..z} \<inter> (\<Inter> (F ` K)) = {}"
by (metis finite_subset_image compact_imp_fip [OF compact_interval Fclo])
then have "K \<noteq> {}"
using \<open>w < z\<close> \<open>{w..z} \<inter> \<Inter>(F ` K) = {}\<close> by auto
define n where "n \<equiv> Max K"
have "n \<in> K" unfolding n_def by (metis \<open>K \<noteq> {}\<close> \<open>finite K\<close> Max_in)
have "F n \<subseteq> \<Inter> (F ` K)"
unfolding n_def by (metis Fsub Max_ge \<open>K \<noteq> {}\<close> \<open>finite K\<close> cINF_greatest lift_Suc_antimono_le)
with K have wzF_null: "{w..z} \<inter> F n = {}"
by (metis disjoint_iff_not_equal subset_eq)
obtain u where u: "u \<in> F n" "u \<in> {x..w}" "({u..w} - {u}) \<inter> F n = {}"
proof (cases "w \<in> F n")
case True
then show ?thesis
by (metis wzF_null \<open>w < z\<close> atLeastAtMost_iff disjoint_iff_not_equal less_eq_real_def)
next
case False
obtain u where "u \<in> F n" "u \<in> {x..w}" "{u<..<w} \<inter> F n = {}"
proof (rule segment_to_point_exists [of "F n \<inter> {x..w}" w])
show "closed (F n \<inter> {x..w})"
by (metis \<open>\<And>n. closed (F n)\<close> closed_Int closed_real_atLeastAtMost)
show "F n \<inter> {x..w} \<noteq> {}"
by (metis atLeastAtMost_iff disjoint_iff_not_equal greaterThanLessThan_iff less.prems(1) less_eq_real_def w)
qed (auto simp: open_segment_eq_real_ivl intro!: that)
with False show thesis
apply (auto simp: disjoint_iff_not_equal intro!: that)
by (metis greaterThanLessThan_iff less_eq_real_def)
qed
obtain v where v: "v \<in> F n" "v \<in> {z..y}" "({z..v} - {v}) \<inter> F n = {}"
proof (cases "z \<in> F n")
case True
have "z \<in> {w..z}"
using \<open>w < z\<close> by auto
then show ?thesis
by (metis wzF_null Int_iff True empty_iff)
next
case False
show ?thesis
proof (rule segment_to_point_exists [of "F n \<inter> {z..y}" z])
show "closed (F n \<inter> {z..y})"
by (metis \<open>\<And>n. closed (F n)\<close> closed_Int closed_atLeastAtMost)
show "F n \<inter> {z..y} \<noteq> {}"
by (metis atLeastAtMost_iff disjoint_iff_not_equal greaterThanLessThan_iff less.prems(2) less_eq_real_def z)
show "\<And>b. \<lbrakk>b \<in> F n \<inter> {z..y}; open_segment z b \<inter> (F n \<inter> {z..y}) = {}\<rbrakk> \<Longrightarrow> thesis"
apply (rule that)
apply (auto simp: open_segment_eq_real_ivl)
by (metis DiffI Int_iff atLeastAtMost_diff_ends atLeastAtMost_iff atLeastatMost_empty_iff empty_iff insert_iff False)
qed
qed
obtain u v where "u \<in> {0..1}" "v \<in> {0..1}" "norm(u - x) < e" "norm(v - y) < e" "p u = p v"
proof
show "u \<in> {0..1}" "v \<in> {0..1}"
by (metis F01 \<open>u \<in> F n\<close> \<open>v \<in> F n\<close> subsetD)+
show "norm(u - x) < e" "norm (v - y) < e"
using \<open>u \<in> {x..w}\<close> \<open>v \<in> {z..y}\<close> atLeastAtMost_iff real_norm_def wxe zye by auto
show "p u = p v"
proof (rule peq)
show "u \<in> F n" "v \<in> F n"
by (auto simp: u v)
have "False" if "\<xi> \<in> F n" "u < \<xi>" "\<xi> < v" for \<xi>
proof -
have "\<xi> \<notin> {z..v}"
- by (metis DiffI disjoint_iff_not_equal less_irrefl singletonD that v(3))
+ by (metis DiffI disjoint_iff_not_equal less_irrefl singletonD that(1,3) v(3))
moreover have "\<xi> \<notin> {w..z} \<inter> F n"
by (metis equals0D wzF_null)
ultimately have "\<xi> \<in> {u..w}"
using that by auto
then show ?thesis
- by (metis DiffI disjoint_iff_not_equal less_eq_real_def not_le singletonD that u(3))
+ by (metis DiffI disjoint_iff_not_equal less_eq_real_def not_le singletonD that(1,2) u(3))
qed
moreover
have "\<lbrakk>\<xi> \<in> F n; v < \<xi>; \<xi> < u\<rbrakk> \<Longrightarrow> False" for \<xi>
using \<open>u \<in> {x..w}\<close> \<open>v \<in> {z..y}\<close> \<open>w < z\<close> by simp
ultimately
show "open_segment u v \<inter> F n = {}"
by (force simp: open_segment_eq_real_ivl)
qed
qed
then show ?case
using e [of x u] e [of y v] xy
apply (simp add: open_segment_eq_real_ivl dist_norm del: divide_const_simps)
by (metis dist_norm dist_triangle_half_r less_irrefl)
qed (auto simp: open_segment_commute)
show ?thesis
- unfolding \<phi>_def by (metis (no_types, hide_lams) INT_I Inf_lower2 rangeI that F01 subsetCE pqF)
+ unfolding \<phi>_def by (metis (no_types, hide_lams) INT_I Inf_lower2 rangeI that(3) F01 subsetCE pqF)
qed
show "closed {0..1::real}" by auto
qed (meson \<phi>_def)
then have "T \<subseteq> {0..1}" "0 \<in> T" "1 \<in> T"
and peq: "\<And>x y. \<lbrakk>x \<in> T; y \<in> T; open_segment x y \<inter> T = {}\<rbrakk> \<Longrightarrow> p x = p y"
unfolding \<phi>_def by metis+
then have "T \<noteq> {}" by auto
define h where "h \<equiv> \<lambda>x. p(SOME y. y \<in> T \<and> open_segment x y \<inter> T = {})"
have "p y = p z" if "y \<in> T" "z \<in> T" and xyT: "open_segment x y \<inter> T = {}" and xzT: "open_segment x z \<inter> T = {}"
for x y z
proof (cases "x \<in> T")
case True
with that show ?thesis by (metis \<open>\<phi> T\<close> \<phi>_def)
next
case False
have "insert x (open_segment x y \<union> open_segment x z) \<inter> T = {}"
by (metis False Int_Un_distrib2 Int_insert_left Un_empty_right xyT xzT)
moreover have "open_segment y z \<inter> T \<subseteq> insert x (open_segment x y \<union> open_segment x z) \<inter> T"
apply auto
by (metis greaterThanLessThan_iff less_eq_real_def less_le_trans linorder_neqE_linordered_idom open_segment_eq_real_ivl)
ultimately have "open_segment y z \<inter> T = {}"
by blast
with that peq show ?thesis by metis
qed
then have h_eq_p_gen: "h x = p y" if "y \<in> T" "open_segment x y \<inter> T = {}" for x y
using that unfolding h_def
by (metis (mono_tags, lifting) some_eq_ex)
then have h_eq_p: "\<And>x. x \<in> T \<Longrightarrow> h x = p x"
by simp
have disjoint: "\<And>x. \<exists>y. y \<in> T \<and> open_segment x y \<inter> T = {}"
by (meson \<open>T \<noteq> {}\<close> \<open>closed T\<close> segment_to_point_exists)
have heq: "h x = h x'" if "open_segment x x' \<inter> T = {}" for x x'
proof (cases "x \<in> T \<or> x' \<in> T")
case True
then show ?thesis
by (metis h_eq_p h_eq_p_gen open_segment_commute that)
next
case False
obtain y y' where "y \<in> T" "open_segment x y \<inter> T = {}" "h x = p y"
"y' \<in> T" "open_segment x' y' \<inter> T = {}" "h x' = p y'"
by (meson disjoint h_eq_p_gen)
moreover have "open_segment y y' \<subseteq> (insert x (insert x' (open_segment x y \<union> open_segment x' y' \<union> open_segment x x')))"
by (auto simp: open_segment_eq_real_ivl)
ultimately show ?thesis
using False that by (fastforce simp add: h_eq_p intro!: peq)
qed
have "h ` {0..1} homeomorphic {0..1::real}"
proof (rule homeomorphic_monotone_image_interval)
show "continuous_on {0..1} h"
proof (clarsimp simp add: continuous_on_iff)
fix u \<epsilon>::real
assume "0 < \<epsilon>" "0 \<le> u" "u \<le> 1"
then obtain \<delta> where "\<delta> > 0" and \<delta>: "\<And>v. v \<in> {0..1} \<Longrightarrow> dist v u < \<delta> \<longrightarrow> dist (p v) (p u) < \<epsilon> / 2"
using ucont_p [unfolded uniformly_continuous_on_def]
by (metis atLeastAtMost_iff half_gt_zero_iff)
then have "dist (h v) (h u) < \<epsilon>" if "v \<in> {0..1}" "dist v u < \<delta>" for v
proof (cases "open_segment u v \<inter> T = {}")
case True
then show ?thesis
using \<open>0 < \<epsilon>\<close> heq by auto
next
case False
have uvT: "closed (closed_segment u v \<inter> T)" "closed_segment u v \<inter> T \<noteq> {}"
using False open_closed_segment by (auto simp: \<open>closed T\<close> closed_Int)
obtain w where "w \<in> T" and w: "w \<in> closed_segment u v" "open_segment u w \<inter> T = {}"
apply (rule segment_to_point_exists [OF uvT, of u])
by (metis IntD1 Int_commute Int_left_commute ends_in_segment(1) inf.orderE subset_oc_segment)
then have puw: "dist (p u) (p w) < \<epsilon> / 2"
by (metis (no_types) \<open>T \<subseteq> {0..1}\<close> \<open>dist v u < \<delta>\<close> \<delta> dist_commute dist_in_closed_segment le_less_trans subsetCE)
obtain z where "z \<in> T" and z: "z \<in> closed_segment u v" "open_segment v z \<inter> T = {}"
apply (rule segment_to_point_exists [OF uvT, of v])
by (metis IntD2 Int_commute Int_left_commute ends_in_segment(2) inf.orderE subset_oc_segment)
then have "dist (p u) (p z) < \<epsilon> / 2"
by (metis \<open>T \<subseteq> {0..1}\<close> \<open>dist v u < \<delta>\<close> \<delta> dist_commute dist_in_closed_segment le_less_trans subsetCE)
then show ?thesis
using puw by (metis (no_types) \<open>w \<in> T\<close> \<open>z \<in> T\<close> dist_commute dist_triangle_half_l h_eq_p_gen w(2) z(2))
qed
with \<open>0 < \<delta>\<close> show "\<exists>\<delta>>0. \<forall>v\<in>{0..1}. dist v u < \<delta> \<longrightarrow> dist (h v) (h u) < \<epsilon>" by blast
qed
show "connected ({0..1} \<inter> h -` {z})" for z
proof (clarsimp simp add: connected_iff_connected_component)
fix u v
assume huv_eq: "h v = h u" and uv: "0 \<le> u" "u \<le> 1" "0 \<le> v" "v \<le> 1"
have "\<exists>T. connected T \<and> T \<subseteq> {0..1} \<and> T \<subseteq> h -` {h u} \<and> u \<in> T \<and> v \<in> T"
proof (intro exI conjI)
show "connected (closed_segment u v)"
by simp
show "closed_segment u v \<subseteq> {0..1}"
by (simp add: uv closed_segment_eq_real_ivl)
have pxy: "p x = p y"
if "T \<subseteq> {0..1}" "0 \<in> T" "1 \<in> T" "x \<in> T" "y \<in> T"
and disjT: "open_segment x y \<inter> (T - open_segment u v) = {}"
and xynot: "x \<notin> open_segment u v" "y \<notin> open_segment u v"
for x y
proof (cases "open_segment x y \<inter> open_segment u v = {}")
case True
then show ?thesis
by (metis Diff_Int_distrib Diff_empty peq disjT \<open>x \<in> T\<close> \<open>y \<in> T\<close>)
next
case False
then have "open_segment x u \<union> open_segment y v \<subseteq> open_segment x y - open_segment u v \<or>
open_segment y u \<union> open_segment x v \<subseteq> open_segment x y - open_segment u v" (is "?xuyv \<or> ?yuxv")
using xynot by (fastforce simp add: open_segment_eq_real_ivl not_le not_less split: if_split_asm)
then show "p x = p y"
proof
assume "?xuyv"
then have "open_segment x u \<inter> T = {}" "open_segment y v \<inter> T = {}"
using disjT by auto
then have "h x = h y"
using heq huv_eq by auto
then show ?thesis
using h_eq_p \<open>x \<in> T\<close> \<open>y \<in> T\<close> by auto
next
assume "?yuxv"
then have "open_segment y u \<inter> T = {}" "open_segment x v \<inter> T = {}"
using disjT by auto
then have "h x = h y"
using heq [of y u] heq [of x v] huv_eq by auto
then show ?thesis
using h_eq_p \<open>x \<in> T\<close> \<open>y \<in> T\<close> by auto
qed
qed
have "\<not> T - open_segment u v \<subset> T"
proof (rule T)
show "closed (T - open_segment u v)"
by (simp add: closed_Diff [OF \<open>closed T\<close>] open_segment_eq_real_ivl)
have "0 \<notin> open_segment u v" "1 \<notin> open_segment u v"
using open_segment_eq_real_ivl uv by auto
then show "\<phi> (T - open_segment u v)"
using \<open>T \<subseteq> {0..1}\<close> \<open>0 \<in> T\<close> \<open>1 \<in> T\<close>
by (auto simp: \<phi>_def) (meson peq pxy)
qed
then have "open_segment u v \<inter> T = {}"
by blast
then show "closed_segment u v \<subseteq> h -` {h u}"
by (force intro: heq simp: open_segment_eq_real_ivl closed_segment_eq_real_ivl split: if_split_asm)+
qed auto
then show "connected_component ({0..1} \<inter> h -` {h u}) u v"
by (simp add: connected_component_def)
qed
show "h 1 \<noteq> h 0"
by (metis \<open>\<phi> T\<close> \<phi>_def a \<open>a \<noteq> b\<close> b h_eq_p pathfinish_def pathstart_def)
qed
then obtain f and g :: "real \<Rightarrow> 'a"
where gfeq: "(\<forall>x\<in>h ` {0..1}. (g(f x) = x))" and fhim: "f ` h ` {0..1} = {0..1}" and contf: "continuous_on (h ` {0..1}) f"
and fgeq: "(\<forall>y\<in>{0..1}. (f(g y) = y))" and pag: "path_image g = h ` {0..1}" and contg: "continuous_on {0..1} g"
by (auto simp: homeomorphic_def homeomorphism_def path_image_def)
then have "arc g"
by (metis arc_def path_def inj_on_def)
obtain u v where "u \<in> {0..1}" "a = g u" "v \<in> {0..1}" "b = g v"
by (metis (mono_tags, hide_lams) \<open>\<phi> T\<close> \<phi>_def a b fhim gfeq h_eq_p imageI path_image_def pathfinish_def pathfinish_in_path_image pathstart_def pathstart_in_path_image)
then have "a \<in> path_image g" "b \<in> path_image g"
using path_image_def by blast+
have ph: "path_image h \<subseteq> path_image p"
by (metis image_mono image_subset_iff path_image_def disjoint h_eq_p_gen \<open>T \<subseteq> {0..1}\<close>)
show ?thesis
proof
show "pathstart (subpath u v g) = a" "pathfinish (subpath u v g) = b"
by (simp_all add: \<open>a = g u\<close> \<open>b = g v\<close>)
show "path_image (subpath u v g) \<subseteq> path_image p"
- by (metis \<open>arc g\<close> \<open>u \<in> {0..1}\<close> \<open>v \<in> {0..1}\<close> arc_imp_path order_trans pag path_image_def path_image_subpath_subset ph)
+ by (metis \<open>u \<in> {0..1}\<close> \<open>v \<in> {0..1}\<close> order_trans pag path_image_def path_image_subpath_subset ph)
show "arc (subpath u v g)"
using \<open>arc g\<close> \<open>a = g u\<close> \<open>b = g v\<close> \<open>u \<in> {0..1}\<close> \<open>v \<in> {0..1}\<close> arc_subpath_arc \<open>a \<noteq> b\<close> by blast
qed
qed
corollary path_connected_arcwise:
fixes S :: "'a::{complete_space,real_normed_vector} set"
shows "path_connected S \<longleftrightarrow>
(\<forall>x \<in> S. \<forall>y \<in> S. x \<noteq> y \<longrightarrow> (\<exists>g. arc g \<and> path_image g \<subseteq> S \<and> pathstart g = x \<and> pathfinish g = y))"
(is "?lhs = ?rhs")
proof (intro iffI impI ballI)
fix x y
assume "path_connected S" "x \<in> S" "y \<in> S" "x \<noteq> y"
then obtain p where p: "path p" "path_image p \<subseteq> S" "pathstart p = x" "pathfinish p = y"
by (force simp: path_connected_def)
then show "\<exists>g. arc g \<and> path_image g \<subseteq> S \<and> pathstart g = x \<and> pathfinish g = y"
by (metis \<open>x \<noteq> y\<close> order_trans path_contains_arc)
next
assume R [rule_format]: ?rhs
show ?lhs
unfolding path_connected_def
proof (intro ballI)
fix x y
assume "x \<in> S" "y \<in> S"
show "\<exists>g. path g \<and> path_image g \<subseteq> S \<and> pathstart g = x \<and> pathfinish g = y"
proof (cases "x = y")
case True with \<open>x \<in> S\<close> path_component_def path_component_refl show ?thesis
by blast
next
case False with R [OF \<open>x \<in> S\<close> \<open>y \<in> S\<close>] show ?thesis
by (auto intro: arc_imp_path)
qed
qed
qed
corollary arc_connected_trans:
fixes g :: "real \<Rightarrow> 'a::{complete_space,real_normed_vector}"
assumes "arc g" "arc h" "pathfinish g = pathstart h" "pathstart g \<noteq> pathfinish h"
obtains i where "arc i" "path_image i \<subseteq> path_image g \<union> path_image h"
"pathstart i = pathstart g" "pathfinish i = pathfinish h"
by (metis (no_types, hide_lams) arc_imp_path assms path_contains_arc path_image_join path_join pathfinish_join pathstart_join)
subsection\<open>Accessibility of frontier points\<close>
lemma dense_accessible_frontier_points:
fixes S :: "'a::{complete_space,real_normed_vector} set"
assumes "open S" and opeSV: "openin (top_of_set (frontier S)) V" and "V \<noteq> {}"
obtains g where "arc g" "g ` {0..<1} \<subseteq> S" "pathstart g \<in> S" "pathfinish g \<in> V"
proof -
obtain z where "z \<in> V"
using \<open>V \<noteq> {}\<close> by auto
then obtain r where "r > 0" and r: "ball z r \<inter> frontier S \<subseteq> V"
by (metis openin_contains_ball opeSV)
then have "z \<in> frontier S"
using \<open>z \<in> V\<close> opeSV openin_contains_ball by blast
then have "z \<in> closure S" "z \<notin> S"
by (simp_all add: frontier_def assms interior_open)
with \<open>r > 0\<close> have "infinite (S \<inter> ball z r)"
by (auto simp: closure_def islimpt_eq_infinite_ball)
then obtain y where "y \<in> S" and y: "y \<in> ball z r"
using infinite_imp_nonempty by force
then have "y \<notin> frontier S"
by (meson \<open>open S\<close> disjoint_iff_not_equal frontier_disjoint_eq)
have "y \<noteq> z"
using \<open>y \<in> S\<close> \<open>z \<notin> S\<close> by blast
have "path_connected(ball z r)"
by (simp add: convex_imp_path_connected)
with y \<open>r > 0\<close> obtain g where "arc g" and pig: "path_image g \<subseteq> ball z r"
and g: "pathstart g = y" "pathfinish g = z"
using \<open>y \<noteq> z\<close> by (force simp: path_connected_arcwise)
have "compact (g -` frontier S \<inter> {0..1})"
- apply (simp add: compact_eq_bounded_closed bounded_Int bounded_closed_interval)
+ apply (simp add: compact_eq_bounded_closed bounded_Int)
apply (rule closed_vimage_Int)
using \<open>arc g\<close> apply (auto simp: arc_def path_def)
done
moreover have "g -` frontier S \<inter> {0..1} \<noteq> {}"
proof -
have "\<exists>r. r \<in> g -` frontier S \<and> r \<in> {0..1}"
by (metis \<open>z \<in> frontier S\<close> g(2) imageE path_image_def pathfinish_in_path_image vimageI2)
then show ?thesis
by blast
qed
ultimately obtain t where gt: "g t \<in> frontier S" and "0 \<le> t" "t \<le> 1"
and t: "\<And>u. \<lbrakk>g u \<in> frontier S; 0 \<le> u; u \<le> 1\<rbrakk> \<Longrightarrow> t \<le> u"
by (force simp: dest!: compact_attains_inf)
moreover have "t \<noteq> 0"
by (metis \<open>y \<notin> frontier S\<close> g(1) gt pathstart_def)
ultimately have t01: "0 < t" "t \<le> 1"
by auto
have "V \<subseteq> frontier S"
using opeSV openin_contains_ball by blast
show ?thesis
proof
show "arc (subpath 0 t g)"
by (simp add: \<open>0 \<le> t\<close> \<open>t \<le> 1\<close> \<open>arc g\<close> \<open>t \<noteq> 0\<close> arc_subpath_arc)
have "g 0 \<in> S"
by (metis \<open>y \<in> S\<close> g(1) pathstart_def)
then show "pathstart (subpath 0 t g) \<in> S"
by auto
have "g t \<in> V"
by (metis IntI atLeastAtMost_iff gt image_eqI path_image_def pig r subsetCE \<open>0 \<le> t\<close> \<open>t \<le> 1\<close>)
then show "pathfinish (subpath 0 t g) \<in> V"
by auto
then have "inj_on (subpath 0 t g) {0..1}"
using t01
apply (clarsimp simp: inj_on_def subpath_def)
apply (drule inj_onD [OF arc_imp_inj_on [OF \<open>arc g\<close>]])
using mult_le_one apply auto
done
then have "subpath 0 t g ` {0..<1} \<subseteq> subpath 0 t g ` {0..1} - {subpath 0 t g 1}"
by (force simp: dest: inj_onD)
moreover have False if "subpath 0 t g ` ({0..<1}) - S \<noteq> {}"
proof -
have contg: "continuous_on {0..1} g"
using \<open>arc g\<close> by (auto simp: arc_def path_def)
have "subpath 0 t g ` {0..<1} \<inter> frontier S \<noteq> {}"
proof (rule connected_Int_frontier [OF _ _ that])
show "connected (subpath 0 t g ` {0..<1})"
apply (rule connected_continuous_image)
apply (simp add: subpath_def)
apply (intro continuous_intros continuous_on_compose2 [OF contg])
apply (auto simp: \<open>0 \<le> t\<close> \<open>t \<le> 1\<close> mult_le_one)
done
show "subpath 0 t g ` {0..<1} \<inter> S \<noteq> {}"
using \<open>y \<in> S\<close> g(1) by (force simp: subpath_def image_def pathstart_def)
qed
then obtain x where "x \<in> subpath 0 t g ` {0..<1}" "x \<in> frontier S"
by blast
with t01 \<open>0 \<le> t\<close> mult_le_one t show False
by (fastforce simp: subpath_def)
qed
then have "subpath 0 t g ` {0..1} - {subpath 0 t g 1} \<subseteq> S"
using subsetD by fastforce
ultimately show "subpath 0 t g ` {0..<1} \<subseteq> S"
by auto
qed
qed
lemma dense_accessible_frontier_points_connected:
fixes S :: "'a::{complete_space,real_normed_vector} set"
assumes "open S" "connected S" "x \<in> S" "V \<noteq> {}"
and ope: "openin (top_of_set (frontier S)) V"
obtains g where "arc g" "g ` {0..<1} \<subseteq> S" "pathstart g = x" "pathfinish g \<in> V"
proof -
have "V \<subseteq> frontier S"
using ope openin_imp_subset by blast
with \<open>open S\<close> \<open>x \<in> S\<close> have "x \<notin> V"
using interior_open by (auto simp: frontier_def)
obtain g where "arc g" and g: "g ` {0..<1} \<subseteq> S" "pathstart g \<in> S" "pathfinish g \<in> V"
by (metis dense_accessible_frontier_points [OF \<open>open S\<close> ope \<open>V \<noteq> {}\<close>])
then have "path_connected S"
by (simp add: assms connected_open_path_connected)
with \<open>pathstart g \<in> S\<close> \<open>x \<in> S\<close> have "path_component S x (pathstart g)"
by (simp add: path_connected_component)
then obtain f where "path f" and f: "path_image f \<subseteq> S" "pathstart f = x" "pathfinish f = pathstart g"
by (auto simp: path_component_def)
then have "path (f +++ g)"
by (simp add: \<open>arc g\<close> arc_imp_path)
then obtain h where "arc h"
and h: "path_image h \<subseteq> path_image (f +++ g)" "pathstart h = x" "pathfinish h = pathfinish g"
apply (rule path_contains_arc [of "f +++ g" x "pathfinish g"])
using f \<open>x \<notin> V\<close> \<open>pathfinish g \<in> V\<close> by auto
have "h ` {0..1} - {h 1} \<subseteq> S"
using f g h apply (clarsimp simp: path_image_join)
apply (simp add: path_image_def pathfinish_def subset_iff image_def Bex_def)
by (metis le_less)
then have "h ` {0..<1} \<subseteq> S"
using \<open>arc h\<close> by (force simp: arc_def dest: inj_onD)
then show thesis
apply (rule that [OF \<open>arc h\<close>])
using h \<open>pathfinish g \<in> V\<close> by auto
qed
lemma dense_access_fp_aux:
fixes S :: "'a::{complete_space,real_normed_vector} set"
assumes S: "open S" "connected S"
and opeSU: "openin (top_of_set (frontier S)) U"
and opeSV: "openin (top_of_set (frontier S)) V"
and "V \<noteq> {}" "\<not> U \<subseteq> V"
obtains g where "arc g" "pathstart g \<in> U" "pathfinish g \<in> V" "g ` {0<..<1} \<subseteq> S"
proof -
have "S \<noteq> {}"
using opeSV \<open>V \<noteq> {}\<close> by (metis frontier_empty openin_subtopology_empty)
then obtain x where "x \<in> S" by auto
obtain g where "arc g" and g: "g ` {0..<1} \<subseteq> S" "pathstart g = x" "pathfinish g \<in> V"
using dense_accessible_frontier_points_connected [OF S \<open>x \<in> S\<close> \<open>V \<noteq> {}\<close> opeSV] by blast
obtain h where "arc h" and h: "h ` {0..<1} \<subseteq> S" "pathstart h = x" "pathfinish h \<in> U - {pathfinish g}"
proof (rule dense_accessible_frontier_points_connected [OF S \<open>x \<in> S\<close>])
show "U - {pathfinish g} \<noteq> {}"
using \<open>pathfinish g \<in> V\<close> \<open>\<not> U \<subseteq> V\<close> by blast
show "openin (top_of_set (frontier S)) (U - {pathfinish g})"
by (simp add: opeSU openin_delete)
qed auto
obtain \<gamma> where "arc \<gamma>"
and \<gamma>: "path_image \<gamma> \<subseteq> path_image (reversepath h +++ g)"
"pathstart \<gamma> = pathfinish h" "pathfinish \<gamma> = pathfinish g"
proof (rule path_contains_arc [of "(reversepath h +++ g)" "pathfinish h" "pathfinish g"])
show "path (reversepath h +++ g)"
by (simp add: \<open>arc g\<close> \<open>arc h\<close> \<open>pathstart g = x\<close> \<open>pathstart h = x\<close> arc_imp_path)
show "pathstart (reversepath h +++ g) = pathfinish h"
"pathfinish (reversepath h +++ g) = pathfinish g"
by auto
show "pathfinish h \<noteq> pathfinish g"
using \<open>pathfinish h \<in> U - {pathfinish g}\<close> by auto
qed auto
show ?thesis
proof
show "arc \<gamma>" "pathstart \<gamma> \<in> U" "pathfinish \<gamma> \<in> V"
using \<gamma> \<open>arc \<gamma>\<close> \<open>pathfinish h \<in> U - {pathfinish g}\<close> \<open>pathfinish g \<in> V\<close> by auto
have "\<gamma> ` {0..1} - {\<gamma> 0, \<gamma> 1} \<subseteq> S"
using \<gamma> g h
apply (simp add: path_image_join)
apply (simp add: path_image_def pathstart_def pathfinish_def subset_iff image_def Bex_def)
by (metis linorder_neqE_linordered_idom not_less)
then show "\<gamma> ` {0<..<1} \<subseteq> S"
using \<open>arc h\<close> \<open>arc \<gamma>\<close>
by (metis arc_imp_simple_path path_image_def pathfinish_def pathstart_def simple_path_endless)
qed
qed
lemma dense_accessible_frontier_point_pairs:
fixes S :: "'a::{complete_space,real_normed_vector} set"
assumes S: "open S" "connected S"
and opeSU: "openin (top_of_set (frontier S)) U"
and opeSV: "openin (top_of_set (frontier S)) V"
and "U \<noteq> {}" "V \<noteq> {}" "U \<noteq> V"
obtains g where "arc g" "pathstart g \<in> U" "pathfinish g \<in> V" "g ` {0<..<1} \<subseteq> S"
proof -
consider "\<not> U \<subseteq> V" | "\<not> V \<subseteq> U"
using \<open>U \<noteq> V\<close> by blast
then show ?thesis
proof cases
case 1 then show ?thesis
using assms dense_access_fp_aux [OF S opeSU opeSV] that by blast
next
case 2
obtain g where "arc g" and g: "pathstart g \<in> V" "pathfinish g \<in> U" "g ` {0<..<1} \<subseteq> S"
using assms dense_access_fp_aux [OF S opeSV opeSU] "2" by blast
show ?thesis
proof
show "arc (reversepath g)"
by (simp add: \<open>arc g\<close> arc_reversepath)
show "pathstart (reversepath g) \<in> U" "pathfinish (reversepath g) \<in> V"
using g by auto
show "reversepath g ` {0<..<1} \<subseteq> S"
using g by (auto simp: reversepath_def)
qed
qed
qed
end
diff --git a/src/HOL/Analysis/Ball_Volume.thy b/src/HOL/Analysis/Ball_Volume.thy
--- a/src/HOL/Analysis/Ball_Volume.thy
+++ b/src/HOL/Analysis/Ball_Volume.thy
@@ -1,333 +1,333 @@
(*
File: HOL/Analysis/Ball_Volume.thy
Author: Manuel Eberl, TU München
*)
section \<open>The Volume of an \<open>n\<close>-Dimensional Ball\<close>
theory Ball_Volume
imports Gamma_Function Lebesgue_Integral_Substitution
begin
text \<open>
We define the volume of the unit ball in terms of the Gamma function. Note that the
dimension need not be an integer; we also allow fractional dimensions, although we do
not use this case or prove anything about it for now.
\<close>
definition\<^marker>\<open>tag important\<close> unit_ball_vol :: "real \<Rightarrow> real" where
"unit_ball_vol n = pi powr (n / 2) / Gamma (n / 2 + 1)"
lemma unit_ball_vol_pos [simp]: "n \<ge> 0 \<Longrightarrow> unit_ball_vol n > 0"
by (force simp: unit_ball_vol_def intro: divide_nonneg_pos)
lemma unit_ball_vol_nonneg [simp]: "n \<ge> 0 \<Longrightarrow> unit_ball_vol n \<ge> 0"
by (simp add: dual_order.strict_implies_order)
text \<open>
We first need the value of the following integral, which is at the core of
computing the measure of an \<open>n + 1\<close>-dimensional ball in terms of the measure of an
\<open>n\<close>-dimensional one.
\<close>
lemma emeasure_cball_aux_integral:
"(\<integral>\<^sup>+x. indicator {-1..1} x * sqrt (1 - x\<^sup>2) ^ n \<partial>lborel) =
ennreal (Beta (1 / 2) (real n / 2 + 1))"
proof -
have "((\<lambda>t. t powr (-1 / 2) * (1 - t) powr (real n / 2)) has_integral
Beta (1 / 2) (real n / 2 + 1)) {0..1}"
using has_integral_Beta_real[of "1/2" "n / 2 + 1"] by simp
from nn_integral_has_integral_lebesgue[OF _ this] have
"ennreal (Beta (1 / 2) (real n / 2 + 1)) =
nn_integral lborel (\<lambda>t. ennreal (t powr (-1 / 2) * (1 - t) powr (real n / 2) *
indicator {0^2..1^2} t))"
by (simp add: mult_ac ennreal_mult' ennreal_indicator)
also have "\<dots> = (\<integral>\<^sup>+ x. ennreal (x\<^sup>2 powr - (1 / 2) * (1 - x\<^sup>2) powr (real n / 2) * (2 * x) *
indicator {0..1} x) \<partial>lborel)"
by (subst nn_integral_substitution[where g = "\<lambda>x. x ^ 2" and g' = "\<lambda>x. 2 * x"])
(auto intro!: derivative_eq_intros continuous_intros simp: set_borel_measurable_def)
also have "\<dots> = (\<integral>\<^sup>+ x. 2 * ennreal ((1 - x\<^sup>2) powr (real n / 2) * indicator {0..1} x) \<partial>lborel)"
by (intro nn_integral_cong_AE AE_I[of _ _ "{0}"])
- (auto simp: indicator_def powr_minus powr_half_sqrt field_split_simps ennreal_mult' mult_ac)
+ (auto simp: indicator_def powr_minus powr_half_sqrt field_split_simps ennreal_mult')
also have "\<dots> = (\<integral>\<^sup>+ x. ennreal ((1 - x\<^sup>2) powr (real n / 2) * indicator {0..1} x) \<partial>lborel) +
(\<integral>\<^sup>+ x. ennreal ((1 - x\<^sup>2) powr (real n / 2) * indicator {0..1} x) \<partial>lborel)"
(is "_ = ?I + _") by (simp add: mult_2 nn_integral_add)
also have "?I = (\<integral>\<^sup>+ x. ennreal ((1 - x\<^sup>2) powr (real n / 2) * indicator {-1..0} x) \<partial>lborel)"
by (subst nn_integral_real_affine[of _ "-1" 0])
(auto simp: indicator_def intro!: nn_integral_cong)
hence "?I + ?I = \<dots> + ?I" by simp
also have "\<dots> = (\<integral>\<^sup>+ x. ennreal ((1 - x\<^sup>2) powr (real n / 2) *
(indicator {-1..0} x + indicator{0..1} x)) \<partial>lborel)"
by (subst nn_integral_add [symmetric]) (auto simp: algebra_simps)
also have "\<dots> = (\<integral>\<^sup>+ x. ennreal ((1 - x\<^sup>2) powr (real n / 2) * indicator {-1..1} x) \<partial>lborel)"
by (intro nn_integral_cong_AE AE_I[of _ _ "{0}"]) (auto simp: indicator_def)
also have "\<dots> = (\<integral>\<^sup>+ x. ennreal (indicator {-1..1} x * sqrt (1 - x\<^sup>2) ^ n) \<partial>lborel)"
by (intro nn_integral_cong_AE AE_I[of _ _ "{1, -1}"])
(auto simp: powr_half_sqrt [symmetric] indicator_def abs_square_le_1
abs_square_eq_1 powr_def exp_of_nat_mult [symmetric] emeasure_lborel_countable)
finally show ?thesis ..
qed
lemma real_sqrt_le_iff': "x \<ge> 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> sqrt x \<le> y \<longleftrightarrow> x \<le> y ^ 2"
using real_le_lsqrt sqrt_le_D by blast
lemma power2_le_iff_abs_le: "y \<ge> 0 \<Longrightarrow> (x::real) ^ 2 \<le> y ^ 2 \<longleftrightarrow> abs x \<le> y"
by (subst real_sqrt_le_iff' [symmetric]) auto
text \<open>
Isabelle's type system makes it very difficult to do an induction over the dimension
of a Euclidean space type, because the type would change in the inductive step. To avoid
this problem, we instead formulate the problem in a more concrete way by unfolding the
definition of the Euclidean norm.
\<close>
lemma emeasure_cball_aux:
assumes "finite A" "r > 0"
shows "emeasure (Pi\<^sub>M A (\<lambda>_. lborel))
({f. sqrt (\<Sum>i\<in>A. (f i)\<^sup>2) \<le> r} \<inter> space (Pi\<^sub>M A (\<lambda>_. lborel))) =
ennreal (unit_ball_vol (real (card A)) * r ^ card A)"
using assms
proof (induction arbitrary: r)
case (empty r)
thus ?case
by (simp add: unit_ball_vol_def space_PiM)
next
case (insert i A r)
interpret product_sigma_finite "\<lambda>_. lborel"
by standard
have "emeasure (Pi\<^sub>M (insert i A) (\<lambda>_. lborel))
({f. sqrt (\<Sum>i\<in>insert i A. (f i)\<^sup>2) \<le> r} \<inter> space (Pi\<^sub>M (insert i A) (\<lambda>_. lborel))) =
nn_integral (Pi\<^sub>M (insert i A) (\<lambda>_. lborel))
(indicator ({f. sqrt (\<Sum>i\<in>insert i A. (f i)\<^sup>2) \<le> r} \<inter>
space (Pi\<^sub>M (insert i A) (\<lambda>_. lborel))))"
by (subst nn_integral_indicator) auto
also have "\<dots> = (\<integral>\<^sup>+ y. \<integral>\<^sup>+ x. indicator ({f. sqrt ((f i)\<^sup>2 + (\<Sum>i\<in>A. (f i)\<^sup>2)) \<le> r} \<inter>
space (Pi\<^sub>M (insert i A) (\<lambda>_. lborel))) (x(i := y))
\<partial>Pi\<^sub>M A (\<lambda>_. lborel) \<partial>lborel)"
using insert.prems insert.hyps by (subst product_nn_integral_insert_rev) auto
also have "\<dots> = (\<integral>\<^sup>+ (y::real). \<integral>\<^sup>+ x. indicator {-r..r} y * indicator ({f. sqrt ((\<Sum>i\<in>A. (f i)\<^sup>2)) \<le>
sqrt (r ^ 2 - y ^ 2)} \<inter> space (Pi\<^sub>M A (\<lambda>_. lborel))) x \<partial>Pi\<^sub>M A (\<lambda>_. lborel) \<partial>lborel)"
proof (intro nn_integral_cong, goal_cases)
case (1 y f)
have *: "y \<in> {-r..r}" if "y ^ 2 + c \<le> r ^ 2" "c \<ge> 0" for c
proof -
have "y ^ 2 \<le> y ^ 2 + c" using that by simp
also have "\<dots> \<le> r ^ 2" by fact
finally show ?thesis
using \<open>r > 0\<close> by (simp add: power2_le_iff_abs_le abs_if split: if_splits)
qed
have "(\<Sum>x\<in>A. (if x = i then y else f x)\<^sup>2) = (\<Sum>x\<in>A. (f x)\<^sup>2)"
using insert.hyps by (intro sum.cong) auto
thus ?case using 1 \<open>r > 0\<close>
by (auto simp: sum_nonneg real_sqrt_le_iff' indicator_def PiE_def space_PiM dest!: *)
qed
also have "\<dots> = (\<integral>\<^sup>+ (y::real). indicator {-r..r} y * (\<integral>\<^sup>+ x. indicator ({f. sqrt ((\<Sum>i\<in>A. (f i)\<^sup>2))
\<le> sqrt (r ^ 2 - y ^ 2)} \<inter> space (Pi\<^sub>M A (\<lambda>_. lborel))) x
\<partial>Pi\<^sub>M A (\<lambda>_. lborel)) \<partial>lborel)" by (subst nn_integral_cmult) auto
also have "\<dots> = (\<integral>\<^sup>+ (y::real). indicator {-r..r} y * emeasure (PiM A (\<lambda>_. lborel))
({f. sqrt ((\<Sum>i\<in>A. (f i)\<^sup>2)) \<le> sqrt (r ^ 2 - y ^ 2)} \<inter> space (Pi\<^sub>M A (\<lambda>_. lborel))) \<partial>lborel)"
using \<open>finite A\<close> by (intro nn_integral_cong, subst nn_integral_indicator) auto
also have "\<dots> = (\<integral>\<^sup>+ (y::real). indicator {-r..r} y * ennreal (unit_ball_vol (real (card A)) *
(sqrt (r ^ 2 - y ^ 2)) ^ card A) \<partial>lborel)"
proof (intro nn_integral_cong_AE, goal_cases)
case 1
have "AE y in lborel. y \<notin> {-r,r}"
by (intro AE_not_in countable_imp_null_set_lborel) auto
thus ?case
proof eventually_elim
case (elim y)
show ?case
proof (cases "y \<in> {-r<..<r}")
case True
hence "y\<^sup>2 < r\<^sup>2" by (subst real_sqrt_less_iff [symmetric]) auto
- thus ?thesis by (subst insert.IH) (auto simp: real_sqrt_less_iff)
+ thus ?thesis by (subst insert.IH) (auto)
qed (insert elim, auto)
qed
qed
also have "\<dots> = ennreal (unit_ball_vol (real (card A))) *
(\<integral>\<^sup>+ (y::real). indicator {-r..r} y * (sqrt (r ^ 2 - y ^ 2)) ^ card A \<partial>lborel)"
by (subst nn_integral_cmult [symmetric])
(auto simp: mult_ac ennreal_mult' [symmetric] indicator_def intro!: nn_integral_cong)
also have "(\<integral>\<^sup>+ (y::real). indicator {-r..r} y * (sqrt (r ^ 2 - y ^ 2)) ^ card A \<partial>lborel) =
(\<integral>\<^sup>+ (y::real). r ^ card A * indicator {-1..1} y * (sqrt (1 - y ^ 2)) ^ card A
\<partial>(distr lborel borel ((*) (1/r))))" using \<open>r > 0\<close>
by (subst nn_integral_distr)
(auto simp: indicator_def field_simps real_sqrt_divide intro!: nn_integral_cong)
also have "\<dots> = (\<integral>\<^sup>+ x. ennreal (r ^ Suc (card A)) *
(indicator {- 1..1} x * sqrt (1 - x\<^sup>2) ^ card A) \<partial>lborel)" using \<open>r > 0\<close>
by (subst lborel_distr_mult) (auto simp: nn_integral_density ennreal_mult' [symmetric] mult_ac)
also have "\<dots> = ennreal (r ^ Suc (card A)) * (\<integral>\<^sup>+ x. indicator {- 1..1} x *
sqrt (1 - x\<^sup>2) ^ card A \<partial>lborel)"
by (subst nn_integral_cmult) auto
also note emeasure_cball_aux_integral
also have "ennreal (unit_ball_vol (real (card A))) * (ennreal (r ^ Suc (card A)) *
ennreal (Beta (1/2) (card A / 2 + 1))) =
ennreal (unit_ball_vol (card A) * Beta (1/2) (card A / 2 + 1) * r ^ Suc (card A))"
using \<open>r > 0\<close> by (simp add: ennreal_mult' [symmetric] mult_ac)
also have "unit_ball_vol (card A) * Beta (1/2) (card A / 2 + 1) = unit_ball_vol (Suc (card A))"
by (auto simp: unit_ball_vol_def Beta_def Gamma_eq_zero_iff field_simps
Gamma_one_half_real powr_half_sqrt [symmetric] powr_add [symmetric])
also have "Suc (card A) = card (insert i A)" using insert.hyps by simp
finally show ?case .
qed
text \<open>
We now get the main theorem very easily by just applying the above lemma.
\<close>
context
fixes c :: "'a :: euclidean_space" and r :: real
assumes r: "r \<ge> 0"
begin
theorem\<^marker>\<open>tag unimportant\<close> emeasure_cball:
"emeasure lborel (cball c r) = ennreal (unit_ball_vol (DIM('a)) * r ^ DIM('a))"
proof (cases "r = 0")
case False
with r have r: "r > 0" by simp
have "(lborel :: 'a measure) =
distr (Pi\<^sub>M Basis (\<lambda>_. lborel)) borel (\<lambda>f. \<Sum>b\<in>Basis. f b *\<^sub>R b)"
by (rule lborel_eq)
also have "emeasure \<dots> (cball 0 r) =
emeasure (Pi\<^sub>M Basis (\<lambda>_. lborel))
({y. dist 0 (\<Sum>b\<in>Basis. y b *\<^sub>R b :: 'a) \<le> r} \<inter> space (Pi\<^sub>M Basis (\<lambda>_. lborel)))"
by (subst emeasure_distr) (auto simp: cball_def)
also have "{f. dist 0 (\<Sum>b\<in>Basis. f b *\<^sub>R b :: 'a) \<le> r} = {f. sqrt (\<Sum>i\<in>Basis. (f i)\<^sup>2) \<le> r}"
by (subst euclidean_dist_l2) (auto simp: L2_set_def)
also have "emeasure (Pi\<^sub>M Basis (\<lambda>_. lborel)) (\<dots> \<inter> space (Pi\<^sub>M Basis (\<lambda>_. lborel))) =
ennreal (unit_ball_vol (real DIM('a)) * r ^ DIM('a))"
using r by (subst emeasure_cball_aux) simp_all
also have "emeasure lborel (cball 0 r :: 'a set) =
emeasure (distr lborel borel (\<lambda>x. c + x)) (cball c r)"
by (subst emeasure_distr) (auto simp: cball_def dist_norm norm_minus_commute)
also have "distr lborel borel (\<lambda>x. c + x) = lborel"
using lborel_affine[of 1 c] by (simp add: density_1)
finally show ?thesis .
qed auto
corollary\<^marker>\<open>tag unimportant\<close> content_cball:
"content (cball c r) = unit_ball_vol (DIM('a)) * r ^ DIM('a)"
by (simp add: measure_def emeasure_cball r)
corollary\<^marker>\<open>tag unimportant\<close> emeasure_ball:
"emeasure lborel (ball c r) = ennreal (unit_ball_vol (DIM('a)) * r ^ DIM('a))"
proof -
from negligible_sphere[of c r] have "sphere c r \<in> null_sets lborel"
by (auto simp: null_sets_completion_iff negligible_iff_null_sets negligible_convex_frontier)
hence "emeasure lborel (ball c r \<union> sphere c r :: 'a set) = emeasure lborel (ball c r :: 'a set)"
by (intro emeasure_Un_null_set) auto
also have "ball c r \<union> sphere c r = (cball c r :: 'a set)" by auto
also have "emeasure lborel \<dots> = ennreal (unit_ball_vol (real DIM('a)) * r ^ DIM('a))"
by (rule emeasure_cball)
finally show ?thesis ..
qed
corollary\<^marker>\<open>tag important\<close> content_ball:
"content (ball c r) = unit_ball_vol (DIM('a)) * r ^ DIM('a)"
by (simp add: measure_def r emeasure_ball)
end
text \<open>
Lastly, we now prove some nicer explicit formulas for the volume of the unit balls in
the cases of even and odd integer dimensions.
\<close>
lemma unit_ball_vol_even:
"unit_ball_vol (real (2 * n)) = pi ^ n / fact n"
by (simp add: unit_ball_vol_def add_ac powr_realpow Gamma_fact)
lemma unit_ball_vol_odd':
"unit_ball_vol (real (2 * n + 1)) = pi ^ n / pochhammer (1 / 2) (Suc n)"
and unit_ball_vol_odd:
"unit_ball_vol (real (2 * n + 1)) =
(2 ^ (2 * Suc n) * fact (Suc n)) / fact (2 * Suc n) * pi ^ n"
proof -
have "unit_ball_vol (real (2 * n + 1)) =
pi powr (real n + 1 / 2) / Gamma (1 / 2 + real (Suc n))"
by (simp add: unit_ball_vol_def field_simps)
also have "pochhammer (1 / 2) (Suc n) = Gamma (1 / 2 + real (Suc n)) / Gamma (1 / 2)"
by (intro pochhammer_Gamma) auto
hence "Gamma (1 / 2 + real (Suc n)) = sqrt pi * pochhammer (1 / 2) (Suc n)"
by (simp add: Gamma_one_half_real)
also have "pi powr (real n + 1 / 2) / \<dots> = pi ^ n / pochhammer (1 / 2) (Suc n)"
by (simp add: powr_add powr_half_sqrt powr_realpow)
finally show "unit_ball_vol (real (2 * n + 1)) = \<dots>" .
also have "pochhammer (1 / 2 :: real) (Suc n) =
fact (2 * Suc n) / (2 ^ (2 * Suc n) * fact (Suc n))"
using fact_double[of "Suc n", where ?'a = real] by (simp add: divide_simps mult_ac)
also have "pi ^n / \<dots> = (2 ^ (2 * Suc n) * fact (Suc n)) / fact (2 * Suc n) * pi ^ n"
by simp
finally show "unit_ball_vol (real (2 * n + 1)) = \<dots>" .
qed
lemma unit_ball_vol_numeral:
"unit_ball_vol (numeral (Num.Bit0 n)) = pi ^ numeral n / fact (numeral n)" (is ?th1)
"unit_ball_vol (numeral (Num.Bit1 n)) = 2 ^ (2 * Suc (numeral n)) * fact (Suc (numeral n)) /
fact (2 * Suc (numeral n)) * pi ^ numeral n" (is ?th2)
proof -
have "numeral (Num.Bit0 n) = (2 * numeral n :: nat)"
by (simp only: numeral_Bit0 mult_2 ring_distribs)
also have "unit_ball_vol \<dots> = pi ^ numeral n / fact (numeral n)"
by (rule unit_ball_vol_even)
finally show ?th1 by simp
next
have "numeral (Num.Bit1 n) = (2 * numeral n + 1 :: nat)"
by (simp only: numeral_Bit1 mult_2)
also have "unit_ball_vol \<dots> = 2 ^ (2 * Suc (numeral n)) * fact (Suc (numeral n)) /
fact (2 * Suc (numeral n)) * pi ^ numeral n"
by (rule unit_ball_vol_odd)
finally show ?th2 by simp
qed
lemmas eval_unit_ball_vol = unit_ball_vol_numeral fact_numeral
text \<open>
Just for fun, we compute the volume of unit balls for a few dimensions.
\<close>
lemma unit_ball_vol_0 [simp]: "unit_ball_vol 0 = 1"
using unit_ball_vol_even[of 0] by simp
lemma unit_ball_vol_1 [simp]: "unit_ball_vol 1 = 2"
using unit_ball_vol_odd[of 0] by simp
corollary\<^marker>\<open>tag unimportant\<close>
unit_ball_vol_2: "unit_ball_vol 2 = pi"
and unit_ball_vol_3: "unit_ball_vol 3 = 4 / 3 * pi"
and unit_ball_vol_4: "unit_ball_vol 4 = pi\<^sup>2 / 2"
and unit_ball_vol_5: "unit_ball_vol 5 = 8 / 15 * pi\<^sup>2"
by (simp_all add: eval_unit_ball_vol)
corollary\<^marker>\<open>tag unimportant\<close> circle_area:
"r \<ge> 0 \<Longrightarrow> content (ball c r :: (real ^ 2) set) = r ^ 2 * pi"
by (simp add: content_ball unit_ball_vol_2)
corollary\<^marker>\<open>tag unimportant\<close> sphere_volume:
"r \<ge> 0 \<Longrightarrow> content (ball c r :: (real ^ 3) set) = 4 / 3 * r ^ 3 * pi"
by (simp add: content_ball unit_ball_vol_3)
text \<open>
Useful equivalent forms
\<close>
corollary\<^marker>\<open>tag unimportant\<close> content_ball_eq_0_iff [simp]: "content (ball c r) = 0 \<longleftrightarrow> r \<le> 0"
proof -
have "r > 0 \<Longrightarrow> content (ball c r) > 0"
by (simp add: content_ball unit_ball_vol_def)
then show ?thesis
by (fastforce simp: ball_empty)
qed
corollary\<^marker>\<open>tag unimportant\<close> content_ball_gt_0_iff [simp]: "0 < content (ball z r) \<longleftrightarrow> 0 < r"
by (auto simp: zero_less_measure_iff)
corollary\<^marker>\<open>tag unimportant\<close> content_cball_eq_0_iff [simp]: "content (cball c r) = 0 \<longleftrightarrow> r \<le> 0"
proof (cases "r = 0")
case False
moreover have "r > 0 \<Longrightarrow> content (cball c r) > 0"
by (simp add: content_cball unit_ball_vol_def)
ultimately show ?thesis
by fastforce
qed auto
corollary\<^marker>\<open>tag unimportant\<close> content_cball_gt_0_iff [simp]: "0 < content (cball z r) \<longleftrightarrow> 0 < r"
by (auto simp: zero_less_measure_iff)
end
diff --git a/src/HOL/Analysis/Bochner_Integration.thy b/src/HOL/Analysis/Bochner_Integration.thy
--- a/src/HOL/Analysis/Bochner_Integration.thy
+++ b/src/HOL/Analysis/Bochner_Integration.thy
@@ -1,3364 +1,3364 @@
(* 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>
proposition 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 sum_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 sum_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!: sum.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)
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)
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 sum_nonneg split: split_indicator split_indicator_asm
simp del: sum_mult_indicator simp: sum_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\<^marker>\<open>tag important\<close> 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)"
proposition 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!: sum.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: sum.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!: sum.cong simp: scaleR_sum_left)
also have "\<dots> = ?r"
by (subst sum.swap)
(auto intro!: sum.cong simp: sum.If_cases scaleR_sum_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: sum.distrib[symmetric] scaleR_add_right)
qed
lemma simple_bochner_integral_linear:
assumes "linear f"
assumes g: "simple_bochner_integrable M g"
shows "simple_bochner_integral M (\<lambda>x. f (g x)) = f (simple_bochner_integral M g)"
proof -
interpret linear f by fact
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
+ (auto simp: simple_bochner_integrable_compose2[where p="\<lambda>x y. f x"]
elim: simple_bochner_integrable.cases)
also have "\<dots> = f (simple_bochner_integral M g)"
by (simp add: simple_bochner_integral_def sum scale)
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 -
from linear_uminus 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_sum)
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 (force simp add: simple_bochner_integral_def intro: sum_nonneg)
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!: sum.cong ennreal_cong_mult
simp: ac_simps ennreal_mult
simp flip: sum_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 "(-)" 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 flip: 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_simp nn_integral_cong_simp)
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 flip: ennreal_plus)
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_onI 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: simple_bochner_integral_linear T.linear_axioms)
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_sum:
"(\<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
proposition 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 flip: ennreal_plus)
(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
proposition 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 flip: ennreal_plus)
(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 flip: 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_left_commute)
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!: sum.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\<^marker>\<open>tag important\<close> 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 integrable_cong_AE_imp:
"integrable M g \<Longrightarrow> f \<in> borel_measurable M \<Longrightarrow> (AE x in M. g x = f x) \<Longrightarrow> integrable M f"
using integrable_cong_AE[of f M g] by (auto simp: eq_commute)
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_sum[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_sum 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_sum: "(\<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_sum has_bochner_integral_integrable)
lemma integral_sum'[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_sum)
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 flip: ennreal_plus)
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
proposition 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 flip: ennreal_0)
qed
qed (insert bnd w_nonneg, auto)
then show ?thesis by simp
qed
proposition 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])
+ by (auto)
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 [simp]:
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 integrable_mult_right_iff [simp]:
fixes f :: "'a \<Rightarrow> real"
shows "integrable M (\<lambda>x. f x * c) \<longleftrightarrow> c = 0 \<or> integrable M f"
using integrable_mult_left_iff [of M c f] by (simp add: mult.commute)
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: integrable_mult_left_iff)
+ qed (auto)
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
proposition 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)
+ by (auto simp add: 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 nn_integral_eq_integrable:
assumes f: "f \<in> M \<rightarrow>\<^sub>M borel" "AE x in M. 0 \<le> f x" and "0 \<le> x"
shows "(\<integral>\<^sup>+x. f x \<partial>M) = ennreal x \<longleftrightarrow> (integrable M f \<and> integral\<^sup>L M f = x)"
proof (safe intro!: nn_integral_eq_integral assms)
assume *: "(\<integral>\<^sup>+x. f x \<partial>M) = ennreal x"
with integrableI_nn_integral_finite[OF f this] nn_integral_eq_integral[of M f, OF _ f(2)]
show "integrable M f" "integral\<^sup>L M f = x"
by (simp_all add: * assms integral_nonneg_AE)
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_sum)
also have "\<dots> \<le> (\<Sum>i. norm (f i x))"
using sum_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
proposition integral_norm_bound [simp]:
fixes f :: "_ \<Rightarrow> 'a :: {banach, second_countable_topology}"
shows "norm (integral\<^sup>L M f) \<le> (\<integral>x. norm (f x) \<partial>M)"
proof (cases "integrable M f")
case True then show ?thesis
using nn_integral_eq_integral[of M "\<lambda>x. norm (f x)"] integral_norm_bound_ennreal[of M f]
by (simp add: integral_nonneg_AE)
next
case False
then have "norm (integral\<^sup>L M f) = 0" by (simp add: not_integrable_integral_eq)
moreover have "(\<integral>x. norm (f x) \<partial>M) \<ge> 0" by auto
ultimately show ?thesis by simp
qed
proposition integral_abs_bound [simp]:
fixes f :: "'a \<Rightarrow> real" shows "abs (\<integral>x. f x \<partial>M) \<le> (\<integral>x. \<bar>f x\<bar> \<partial>M)"
using integral_norm_bound[of M f] by auto
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)
proposition 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 sum = 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="(*\<^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!: sum 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 integral_norm_bound_integral:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::{banach,second_countable_topology}"
assumes "integrable M f" "integrable M g" "\<And>x. x \<in> space M \<Longrightarrow> norm(f x) \<le> g x"
shows "norm (\<integral>x. f x \<partial>M) \<le> (\<integral>x. g x \<partial>M)"
proof -
have "norm (\<integral>x. f x \<partial>M) \<le> (\<integral>x. norm (f x) \<partial>M)"
by (rule integral_norm_bound)
also have "... \<le> (\<integral>x. g x \<partial>M)"
using assms integrable_norm integral_mono by blast
finally show ?thesis .
qed
lemma integral_abs_bound_integral:
fixes f :: "'a::euclidean_space \<Rightarrow> real"
assumes "integrable M f" "integrable M g" "\<And>x. x \<in> space M \<Longrightarrow> \<bar>f x\<bar> \<le> g x"
shows "\<bar>\<integral>x. f x \<partial>M\<bar> \<le> (\<integral>x. g x \<partial>M)"
by (metis integral_norm_bound_integral assms real_norm_def)
text \<open>The next two statements are useful to bound Lebesgue integrals, as they avoid one
integrability assumption. The price to pay is that the upper function has to be nonnegative,
but this is often true and easy to check in computations.\<close>
lemma integral_mono_AE':
fixes f::"_ \<Rightarrow> real"
assumes "integrable M f" "AE x in M. g x \<le> f x" "AE x in M. 0 \<le> f x"
shows "(\<integral>x. g x \<partial>M) \<le> (\<integral>x. f x \<partial>M)"
proof (cases "integrable M g")
case True
show ?thesis by (rule integral_mono_AE, auto simp add: assms True)
next
case False
then have "(\<integral>x. g x \<partial>M) = 0" by (simp add: not_integrable_integral_eq)
also have "... \<le> (\<integral>x. f x \<partial>M)" by (simp add: integral_nonneg_AE[OF assms(3)])
finally show ?thesis by simp
qed
lemma integral_mono':
fixes f::"_ \<Rightarrow> real"
assumes "integrable M f" "\<And>x. x \<in> space M \<Longrightarrow> g x \<le> f x" "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
shows "(\<integral>x. g x \<partial>M) \<le> (\<integral>x. f x \<partial>M)"
by (rule integral_mono_AE', insert assms, 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 nn_integral_nonneg_infinite:
fixes f::"'a \<Rightarrow> real"
assumes "f \<in> borel_measurable M" "\<not> integrable M f" "AE x in M. f x \<ge> 0"
shows "(\<integral>\<^sup>+x. f x \<partial>M) = \<infinity>"
using assms integrableI_real_bounded less_top by auto
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
lemma integrable_MIN:
fixes f:: "_ \<Rightarrow> _ \<Rightarrow> real"
shows "\<lbrakk> finite I; I \<noteq> {}; \<And>i. i \<in> I \<Longrightarrow> integrable M (f i) \<rbrakk>
\<Longrightarrow> integrable M (\<lambda>x. MIN i\<in>I. f i x)"
by (induct rule: finite_ne_induct) simp+
lemma integrable_MAX:
fixes f:: "_ \<Rightarrow> _ \<Rightarrow> real"
shows "\<lbrakk> finite I; I \<noteq> {}; \<And>i. i \<in> I \<Longrightarrow> integrable M (f i) \<rbrakk>
\<Longrightarrow> integrable M (\<lambda>x. MAX i\<in>I. f i x)"
by (induct rule: finite_ne_induct) simp+
theorem integral_Markov_inequality:
assumes [measurable]: "integrable M u" and "AE x in M. 0 \<le> u x" "0 < (c::real)"
shows "(emeasure M) {x\<in>space M. u x \<ge> c} \<le> (1/c) * (\<integral>x. u x \<partial>M)"
proof -
have "(\<integral>\<^sup>+ x. ennreal(u x) * indicator (space M) x \<partial>M) \<le> (\<integral>\<^sup>+ x. u x \<partial>M)"
by (rule nn_integral_mono_AE, auto simp add: \<open>c>0\<close> less_eq_real_def)
also have "... = (\<integral>x. u x \<partial>M)"
by (rule nn_integral_eq_integral, auto simp add: assms)
finally have *: "(\<integral>\<^sup>+ x. ennreal(u x) * indicator (space M) x \<partial>M) \<le> (\<integral>x. u x \<partial>M)"
by simp
have "{x \<in> space M. u x \<ge> c} = {x \<in> space M. ennreal(1/c) * u x \<ge> 1} \<inter> (space M)"
using \<open>c>0\<close> by (auto simp: ennreal_mult'[symmetric])
then have "emeasure M {x \<in> space M. u x \<ge> c} = emeasure M ({x \<in> space M. ennreal(1/c) * u x \<ge> 1} \<inter> (space M))"
by simp
also have "... \<le> ennreal(1/c) * (\<integral>\<^sup>+ x. ennreal(u x) * indicator (space M) x \<partial>M)"
by (rule nn_integral_Markov_inequality) (auto simp add: assms)
also have "... \<le> ennreal(1/c) * (\<integral>x. u x \<partial>M)"
apply (rule mult_left_mono) using * \<open>c>0\<close> by auto
finally show ?thesis
using \<open>0<c\<close> by (simp add: ennreal_mult'[symmetric])
qed
lemma integral_ineq_eq_0_then_AE:
fixes f::"_ \<Rightarrow> real"
assumes "AE x in M. f x \<le> g x" "integrable M f" "integrable M g"
"(\<integral>x. f x \<partial>M) = (\<integral>x. g x \<partial>M)"
shows "AE x in M. f x = g x"
proof -
define h where "h = (\<lambda>x. g x - f x)"
have "AE x in M. h x = 0"
apply (subst integral_nonneg_eq_0_iff_AE[symmetric])
unfolding h_def using assms by auto
then show ?thesis unfolding h_def by auto
qed
lemma not_AE_zero_int_E:
fixes f::"'a \<Rightarrow> real"
assumes "AE x in M. f x \<ge> 0" "(\<integral>x. f x \<partial>M) > 0"
and [measurable]: "f \<in> borel_measurable M"
shows "\<exists>A e. A \<in> sets M \<and> e>0 \<and> emeasure M A > 0 \<and> (\<forall>x \<in> A. f x \<ge> e)"
proof (rule not_AE_zero_E, auto simp add: assms)
assume *: "AE x in M. f x = 0"
have "(\<integral>x. f x \<partial>M) = (\<integral>x. 0 \<partial>M)" by (rule integral_cong_AE, auto simp add: *)
then have "(\<integral>x. f x \<partial>M) = 0" by simp
then show False using assms(2) by simp
qed
proposition tendsto_L1_int:
fixes u :: "_ \<Rightarrow> _ \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes [measurable]: "\<And>n. integrable M (u n)" "integrable M f"
and "((\<lambda>n. (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)) \<longlongrightarrow> 0) F"
shows "((\<lambda>n. (\<integral>x. u n x \<partial>M)) \<longlongrightarrow> (\<integral>x. f x \<partial>M)) F"
proof -
have "((\<lambda>n. norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<longlongrightarrow> (0::ennreal)) F"
proof (rule tendsto_sandwich[of "\<lambda>_. 0", where ?h = "\<lambda>n. (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)"], auto simp add: assms)
{
fix n
have "(\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M) = (\<integral>x. u n x - f x \<partial>M)"
apply (rule Bochner_Integration.integral_diff[symmetric]) using assms by auto
then have "norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M)) = norm (\<integral>x. u n x - f x \<partial>M)"
by auto
also have "... \<le> (\<integral>x. norm(u n x - f x) \<partial>M)"
by (rule integral_norm_bound)
finally have "ennreal(norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<le> (\<integral>x. norm(u n x - f x) \<partial>M)"
by simp
also have "... = (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)"
apply (rule nn_integral_eq_integral[symmetric]) using assms by auto
finally have "norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M)) \<le> (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)" by simp
}
then show "eventually (\<lambda>n. norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M)) \<le> (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)) F"
by auto
qed
then have "((\<lambda>n. norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<longlongrightarrow> 0) F"
by (simp flip: ennreal_0)
then have "((\<lambda>n. ((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<longlongrightarrow> 0) F" using tendsto_norm_zero_iff by blast
then show ?thesis using Lim_null by auto
qed
text \<open>The next lemma asserts that, if a sequence of functions converges in \<open>L\<^sup>1\<close>, then
it admits a subsequence that converges almost everywhere.\<close>
proposition tendsto_L1_AE_subseq:
fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes [measurable]: "\<And>n. integrable M (u n)"
and "(\<lambda>n. (\<integral>x. norm(u n x) \<partial>M)) \<longlonglongrightarrow> 0"
shows "\<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> (AE x in M. (\<lambda>n. u (r n) x) \<longlonglongrightarrow> 0)"
proof -
{
fix k
have "eventually (\<lambda>n. (\<integral>x. norm(u n x) \<partial>M) < (1/2)^k) sequentially"
using order_tendstoD(2)[OF assms(2)] by auto
with eventually_elim2[OF eventually_gt_at_top[of k] this]
have "\<exists>n>k. (\<integral>x. norm(u n x) \<partial>M) < (1/2)^k"
by (metis eventually_False_sequentially)
}
then have "\<exists>r. \<forall>n. True \<and> (r (Suc n) > r n \<and> (\<integral>x. norm(u (r (Suc n)) x) \<partial>M) < (1/2)^(r n))"
by (intro dependent_nat_choice, auto)
then obtain r0 where r0: "strict_mono r0" "\<And>n. (\<integral>x. norm(u (r0 (Suc n)) x) \<partial>M) < (1/2)^(r0 n)"
by (auto simp: strict_mono_Suc_iff)
define r where "r = (\<lambda>n. r0(n+1))"
have "strict_mono r" unfolding r_def using r0(1) by (simp add: strict_mono_Suc_iff)
have I: "(\<integral>\<^sup>+x. norm(u (r n) x) \<partial>M) < ennreal((1/2)^n)" for n
proof -
have "r0 n \<ge> n" using \<open>strict_mono r0\<close> by (simp add: seq_suble)
have "(1/2::real)^(r0 n) \<le> (1/2)^n" by (rule power_decreasing[OF \<open>r0 n \<ge> n\<close>], auto)
then have "(\<integral>x. norm(u (r0 (Suc n)) x) \<partial>M) < (1/2)^n"
using r0(2) less_le_trans by blast
then have "(\<integral>x. norm(u (r n) x) \<partial>M) < (1/2)^n"
unfolding r_def by auto
moreover have "(\<integral>\<^sup>+x. norm(u (r n) x) \<partial>M) = (\<integral>x. norm(u (r n) x) \<partial>M)"
by (rule nn_integral_eq_integral, auto simp add: integrable_norm[OF assms(1)[of "r n"]])
ultimately show ?thesis by (auto intro: ennreal_lessI)
qed
have "AE x in M. limsup (\<lambda>n. ennreal (norm(u (r n) x))) \<le> 0"
proof (rule AE_upper_bound_inf_ennreal)
fix e::real assume "e > 0"
define A where "A = (\<lambda>n. {x \<in> space M. norm(u (r n) x) \<ge> e})"
have A_meas [measurable]: "\<And>n. A n \<in> sets M" unfolding A_def by auto
have A_bound: "emeasure M (A n) < (1/e) * ennreal((1/2)^n)" for n
proof -
have *: "indicator (A n) x \<le> (1/e) * ennreal(norm(u (r n) x))" for x
apply (cases "x \<in> A n") unfolding A_def using \<open>0 < e\<close> by (auto simp: ennreal_mult[symmetric])
have "emeasure M (A n) = (\<integral>\<^sup>+x. indicator (A n) x \<partial>M)" by auto
also have "... \<le> (\<integral>\<^sup>+x. (1/e) * ennreal(norm(u (r n) x)) \<partial>M)"
apply (rule nn_integral_mono) using * by auto
also have "... = (1/e) * (\<integral>\<^sup>+x. norm(u (r n) x) \<partial>M)"
apply (rule nn_integral_cmult) using \<open>e > 0\<close> by auto
also have "... < (1/e) * ennreal((1/2)^n)"
using I[of n] \<open>e > 0\<close> by (intro ennreal_mult_strict_left_mono) auto
finally show ?thesis by simp
qed
have A_fin: "emeasure M (A n) < \<infinity>" for n
using \<open>e > 0\<close> A_bound[of n]
by (auto simp add: ennreal_mult less_top[symmetric])
have A_sum: "summable (\<lambda>n. measure M (A n))"
proof (rule summable_comparison_test'[of "\<lambda>n. (1/e) * (1/2)^n" 0])
have "summable (\<lambda>n. (1/(2::real))^n)" by (simp add: summable_geometric)
then show "summable (\<lambda>n. (1/e) * (1/2)^n)" using summable_mult by blast
fix n::nat assume "n \<ge> 0"
have "norm(measure M (A n)) = measure M (A n)" by simp
also have "... = enn2real(emeasure M (A n))" unfolding measure_def by simp
also have "... < enn2real((1/e) * (1/2)^n)"
using A_bound[of n] \<open>emeasure M (A n) < \<infinity>\<close> \<open>0 < e\<close>
by (auto simp: emeasure_eq_ennreal_measure ennreal_mult[symmetric] ennreal_less_iff)
also have "... = (1/e) * (1/2)^n"
using \<open>0<e\<close> by auto
finally show "norm(measure M (A n)) \<le> (1/e) * (1/2)^n" by simp
qed
have "AE x in M. eventually (\<lambda>n. x \<in> space M - A n) sequentially"
by (rule borel_cantelli_AE1[OF A_meas A_fin A_sum])
moreover
{
fix x assume "eventually (\<lambda>n. x \<in> space M - A n) sequentially"
moreover have "norm(u (r n) x) \<le> ennreal e" if "x \<in> space M - A n" for n
using that unfolding A_def by (auto intro: ennreal_leI)
ultimately have "eventually (\<lambda>n. norm(u (r n) x) \<le> ennreal e) sequentially"
by (simp add: eventually_mono)
then have "limsup (\<lambda>n. ennreal (norm(u (r n) x))) \<le> e"
by (simp add: Limsup_bounded)
}
ultimately show "AE x in M. limsup (\<lambda>n. ennreal (norm(u (r n) x))) \<le> 0 + ennreal e" by auto
qed
moreover
{
fix x assume "limsup (\<lambda>n. ennreal (norm(u (r n) x))) \<le> 0"
moreover then have "liminf (\<lambda>n. ennreal (norm(u (r n) x))) \<le> 0"
by (rule order_trans[rotated]) (auto intro: Liminf_le_Limsup)
ultimately have "(\<lambda>n. ennreal (norm(u (r n) x))) \<longlonglongrightarrow> 0"
using tendsto_Limsup[of sequentially "\<lambda>n. ennreal (norm(u (r n) x))"] by auto
then have "(\<lambda>n. norm(u (r n) x)) \<longlonglongrightarrow> 0"
by (simp flip: ennreal_0)
then have "(\<lambda>n. u (r n) x) \<longlonglongrightarrow> 0"
by (simp add: tendsto_norm_zero_iff)
}
ultimately have "AE x in M. (\<lambda>n. u (r n) x) \<longlonglongrightarrow> 0" by auto
then show ?thesis using \<open>strict_mono r\<close> by auto
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="(=)"] 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>\<open>count_space\<close>\<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 sum.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_sum) (auto intro!: sum.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!: sum.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)
lemma integrable_bij_count_space:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes g: "bij_betw g A B"
shows "integrable (count_space A) (\<lambda>x. f (g x)) \<longleftrightarrow> integrable (count_space B) f"
unfolding integrable_iff_bounded by (subst nn_integral_bij_count_space[OF g]) auto
lemma integral_bij_count_space:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes g: "bij_betw g A B"
shows "integral\<^sup>L (count_space A) (\<lambda>x. f (g x)) = integral\<^sup>L (count_space B) f"
using g[THEN bij_betw_imp_funcset]
apply (subst distr_bij_count_space[OF g, symmetric])
apply (intro integral_distr[symmetric])
apply auto
done
lemma has_bochner_integral_count_space_nat:
fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}"
shows "has_bochner_integral (count_space UNIV) f x \<Longrightarrow> f sums x"
unfolding has_bochner_integral_iff by (auto intro!: sums_integral_count_space_nat)
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)
proposition 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>\<open>null_measure\<close>\<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>
theorem 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
theorem 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 "0 \<le> r" "0 \<le> q"
"(\<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 (in finite_measure) integral_bounded_eq_bound_then_AE:
assumes "AE x in M. f x \<le> (c::real)"
"integrable M f" "(\<integral>x. f x \<partial>M) = c * measure M (space M)"
shows "AE x in M. f x = c"
apply (rule integral_ineq_eq_0_then_AE) using assms by auto
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_sum) (auto simp add: integrable_mult_left_iff)
+ using finite by (subst integral_sum) (auto)
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)
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
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 tendsto_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 (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
proposition (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 sum.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 unfolding integrable_iff_bounded by simp
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
theorem (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
proposition (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
proposition (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_prod:
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: prod_norm prod_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_prod) auto
also have "\<dots> < \<infinity>"
using integrable by (simp add: less_top[symmetric] ennreal_prod_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_prod:
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_prod 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!: prod.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
diff --git a/src/HOL/Analysis/Borel_Space.thy b/src/HOL/Analysis/Borel_Space.thy
--- a/src/HOL/Analysis/Borel_Space.thy
+++ b/src/HOL/Analysis/Borel_Space.thy
@@ -1,2080 +1,2080 @@
(* Title: HOL/Analysis/Borel_Space.thy
Author: Johannes Hölzl, TU München
Author: Armin Heller, TU München
*)
section \<open>Borel Space\<close>
theory Borel_Space
imports
Measurable Derivative Ordered_Euclidean_Space Extended_Real_Limits
begin
lemma is_interval_real_ereal_oo: "is_interval (real_of_ereal ` {N<..<M::ereal})"
by (auto simp: real_atLeastGreaterThan_eq)
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)
proposition 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
proposition 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_lowerbound)
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 ((+) (-x) ` interior A)" by (rule open_translation)
also have "((+) (-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
proposition 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)
using assms differentiable_at_withinI real_differentiable_def by blast
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\<^marker>\<open>tag important\<close> (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 sets_borel_eq_count_space: "sets (borel :: 'a::{countable, t2_space} measure) = count_space UNIV"
proof -
have "(\<Union>a\<in>A. {a}) \<in> sets borel" for A :: "'a set"
by (intro sets.countable_UN') auto
then show ?thesis
by auto
qed
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
theorem 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
proposition 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 "open" B using assms by (rule countable_basis_openI)
fix X::"'a set" assume "open X"
from open_countable_basisE[OF this] obtain B' where B': "B' \<subseteq> B" "X = \<Union> B'" .
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_onI: "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_onI[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\<in>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\<in>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\<in>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\<in>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\<in>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\<in>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\<in>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\<in>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 simp add: image_comp)
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 simp add: image_comp)
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_sum[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_onI 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_prod[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 borel_measurable_uminus_eq [simp]:
fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}"
shows "(\<lambda>x. - f x) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r")
proof
assume ?l from borel_measurable_uminus[OF this] show ?r by simp
qed auto
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_onI 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_onI continuous_intros)
lemma borel_measurable_sqrt [measurable]: "sqrt \<in> borel_measurable borel"
by (intro borel_measurable_continuous_onI 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_onI continuous_intros)
lemma borel_measurable_Im [measurable]: "Im \<in> borel_measurable borel"
by (intro borel_measurable_continuous_onI continuous_intros)
lemma borel_measurable_of_real [measurable]: "(of_real :: _ \<Rightarrow> (_::real_normed_algebra)) \<in> borel_measurable borel"
by (intro borel_measurable_continuous_onI continuous_intros)
lemma borel_measurable_sin [measurable]: "(sin :: _ \<Rightarrow> (_::{real_normed_field,banach})) \<in> borel_measurable borel"
by (intro borel_measurable_continuous_onI continuous_intros)
lemma borel_measurable_cos [measurable]: "(cos :: _ \<Rightarrow> (_::{real_normed_field,banach})) \<in> borel_measurable borel"
by (intro borel_measurable_continuous_onI continuous_intros)
lemma borel_measurable_arctan [measurable]: "arctan \<in> borel_measurable borel"
by (intro borel_measurable_continuous_onI continuous_intros)
lemma\<^marker>\<open>tag important\<close> 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
lemma powr_real_measurable [measurable]:
assumes "f \<in> measurable M borel" "g \<in> measurable M borel"
shows "(\<lambda>x. f x powr g x :: real) \<in> measurable M borel"
using assms by (simp_all add: powr_def)
lemma measurable_of_bool[measurable]: "of_bool \<in> count_space UNIV \<rightarrow>\<^sub>M borel"
by simp
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
theorem 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_sum[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_prod[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>\<open>ennreal\<close> 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_onI continuous_on_enn2ereal)
lemma measurable_e2ennreal[measurable]: "e2ennreal \<in> borel \<rightarrow>\<^sub>M borel"
by (intro borel_measurable_continuous_onI 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\<^marker>\<open>tag important\<close> [simp]: "is_borel f M \<longleftrightarrow> f \<in> borel_measurable M"
lemma is_borel_transfer[transfer_rule]: "rel_fun (rel_fun (=) pcr_ennreal) (=) 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_prod_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)
+ by (simp add: field_simps)
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
text \<open>The next lemmas hold in any second countable linorder (including ennreal or ereal for instance),
but in the current state they are restricted to reals.\<close>
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_piecewise_mono:
fixes f::"real \<Rightarrow> real" and C::"real set set"
assumes "countable C" "\<And>c. c \<in> C \<Longrightarrow> c \<in> sets borel" "\<And>c. c \<in> C \<Longrightarrow> mono_on f c" "(\<Union>C) = UNIV"
shows "f \<in> borel_measurable borel"
by (rule measurable_piecewise_restrict[of C], auto intro: borel_measurable_mono_on_fnc simp: assms)
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)
lemma measurable_bdd_below_real[measurable (raw)]:
fixes F :: "'a \<Rightarrow> 'i \<Rightarrow> real"
assumes [simp]: "countable I" and [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> M \<rightarrow>\<^sub>M borel"
shows "Measurable.pred M (\<lambda>x. bdd_below ((\<lambda>i. F i x)`I))"
proof (subst measurable_cong)
show "bdd_below ((\<lambda>i. F i x)`I) \<longleftrightarrow> (\<exists>q\<in>\<int>. \<forall>i\<in>I. q \<le> F i x)" for x
by (auto simp: bdd_below_def intro!: bexI[of _ "of_int (floor _)"] intro: order_trans of_int_floor_le)
show "Measurable.pred M (\<lambda>w. \<exists>q\<in>\<int>. \<forall>i\<in>I. q \<le> F i w)"
using countable_int by measurable
qed
lemma borel_measurable_cINF_real[measurable (raw)]:
fixes F :: "_ \<Rightarrow> _ \<Rightarrow> real"
assumes [simp]: "countable I"
assumes F[measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M"
shows "(\<lambda>x. INF i\<in>I. F i x) \<in> borel_measurable M"
proof (rule measurable_piecewise_restrict)
let ?\<Omega> = "{x\<in>space M. bdd_below ((\<lambda>i. F i x)`I)}"
show "countable {?\<Omega>, - ?\<Omega>}" "space M \<subseteq> \<Union>{?\<Omega>, - ?\<Omega>}" "\<And>X. X \<in> {?\<Omega>, - ?\<Omega>} \<Longrightarrow> X \<inter> space M \<in> sets M"
by auto
fix X assume "X \<in> {?\<Omega>, - ?\<Omega>}" then show "(\<lambda>x. INF i\<in>I. F i x) \<in> borel_measurable (restrict_space M X)"
proof safe
show "(\<lambda>x. INF i\<in>I. F i x) \<in> borel_measurable (restrict_space M ?\<Omega>)"
by (intro borel_measurable_cINF measurable_restrict_space1 F)
(auto simp: space_restrict_space)
show "(\<lambda>x. INF i\<in>I. F i x) \<in> borel_measurable (restrict_space M (-?\<Omega>))"
proof (subst measurable_cong)
fix x assume "x \<in> space (restrict_space M (-?\<Omega>))"
then have "\<not> (\<forall>i\<in>I. - F i x \<le> y)" for y
by (auto simp: space_restrict_space bdd_above_def bdd_above_uminus[symmetric])
then show "(INF i\<in>I. F i x) = - (THE x. False)"
by (auto simp: space_restrict_space Inf_real_def Sup_real_def Least_def simp del: Set.ball_simps(10))
qed simp
qed
qed
lemma borel_Ici: "borel = sigma UNIV (range (\<lambda>x::real. {x ..}))"
proof (safe intro!: borel_eq_sigmaI1[OF borel_Iio])
fix x :: real
have eq: "{..<x} = space (sigma UNIV (range atLeast)) - {x ..}"
by auto
show "{..<x} \<in> sets (sigma UNIV (range atLeast))"
unfolding eq by (intro sets.compl_sets) auto
qed auto
lemma borel_measurable_pred_less[measurable (raw)]:
fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, linorder_topology}"
shows "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> Measurable.pred M (\<lambda>w. f w < g w)"
unfolding Measurable.pred_def by (rule borel_measurable_less)
no_notation
eucl_less (infix "<e" 50)
lemma borel_measurable_Max2[measurable (raw)]:
fixes f::"_ \<Rightarrow> _ \<Rightarrow> 'a::{second_countable_topology, dense_linorder, linorder_topology}"
assumes "finite I"
and [measurable]: "\<And>i. f i \<in> borel_measurable M"
shows "(\<lambda>x. Max{f i x |i. i \<in> I}) \<in> borel_measurable M"
by (simp add: borel_measurable_Max[OF assms(1), where ?f=f and ?M=M] Setcompr_eq_image)
lemma measurable_compose_n [measurable (raw)]:
assumes "T \<in> measurable M M"
shows "(T^^n) \<in> measurable M M"
by (induction n, auto simp add: measurable_compose[OF _ assms])
lemma measurable_real_imp_nat:
fixes f::"'a \<Rightarrow> nat"
assumes [measurable]: "(\<lambda>x. real(f x)) \<in> borel_measurable M"
shows "f \<in> measurable M (count_space UNIV)"
proof -
let ?g = "(\<lambda>x. real(f x))"
have "\<And>(n::nat). ?g-`({real n}) \<inter> space M = f-`{n} \<inter> space M" by auto
moreover have "\<And>(n::nat). ?g-`({real n}) \<inter> space M \<in> sets M" using assms by measurable
ultimately have "\<And>(n::nat). f-`{n} \<inter> space M \<in> sets M" by simp
then show ?thesis using measurable_count_space_eq2_countable by blast
qed
lemma measurable_equality_set [measurable]:
fixes f g::"_\<Rightarrow> 'a::{second_countable_topology, t2_space}"
assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
shows "{x \<in> space M. f x = g x} \<in> sets M"
proof -
define A where "A = {x \<in> space M. f x = g x}"
define B where "B = {y. \<exists>x::'a. y = (x,x)}"
have "A = (\<lambda>x. (f x, g x))-`B \<inter> space M" unfolding A_def B_def by auto
moreover have "(\<lambda>x. (f x, g x)) \<in> borel_measurable M" by simp
moreover have "B \<in> sets borel" unfolding B_def by (simp add: closed_diagonal)
ultimately have "A \<in> sets M" by simp
then show ?thesis unfolding A_def by simp
qed
lemma measurable_inequality_set [measurable]:
fixes f g::"_ \<Rightarrow> 'a::{second_countable_topology, linorder_topology}"
assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
shows "{x \<in> space M. f x \<le> g x} \<in> sets M"
"{x \<in> space M. f x < g x} \<in> sets M"
"{x \<in> space M. f x \<ge> g x} \<in> sets M"
"{x \<in> space M. f x > g x} \<in> sets M"
proof -
define F where "F = (\<lambda>x. (f x, g x))"
have * [measurable]: "F \<in> borel_measurable M" unfolding F_def by simp
have "{x \<in> space M. f x \<le> g x} = F-`{(x, y) | x y. x \<le> y} \<inter> space M" unfolding F_def by auto
moreover have "{(x, y) | x y. x \<le> (y::'a)} \<in> sets borel" using closed_subdiagonal borel_closed by blast
ultimately show "{x \<in> space M. f x \<le> g x} \<in> sets M" using * by (metis (mono_tags, lifting) measurable_sets)
have "{x \<in> space M. f x < g x} = F-`{(x, y) | x y. x < y} \<inter> space M" unfolding F_def by auto
moreover have "{(x, y) | x y. x < (y::'a)} \<in> sets borel" using open_subdiagonal borel_open by blast
ultimately show "{x \<in> space M. f x < g x} \<in> sets M" using * by (metis (mono_tags, lifting) measurable_sets)
have "{x \<in> space M. f x \<ge> g x} = F-`{(x, y) | x y. x \<ge> y} \<inter> space M" unfolding F_def by auto
moreover have "{(x, y) | x y. x \<ge> (y::'a)} \<in> sets borel" using closed_superdiagonal borel_closed by blast
ultimately show "{x \<in> space M. f x \<ge> g x} \<in> sets M" using * by (metis (mono_tags, lifting) measurable_sets)
have "{x \<in> space M. f x > g x} = F-`{(x, y) | x y. x > y} \<inter> space M" unfolding F_def by auto
moreover have "{(x, y) | x y. x > (y::'a)} \<in> sets borel" using open_superdiagonal borel_open by blast
ultimately show "{x \<in> space M. f x > g x} \<in> sets M" using * by (metis (mono_tags, lifting) measurable_sets)
qed
proposition measurable_limit [measurable]:
fixes f::"nat \<Rightarrow> 'a \<Rightarrow> 'b::first_countable_topology"
assumes [measurable]: "\<And>n::nat. f n \<in> borel_measurable M"
shows "Measurable.pred M (\<lambda>x. (\<lambda>n. f n x) \<longlonglongrightarrow> c)"
proof -
obtain A :: "nat \<Rightarrow> 'b set" where A:
"\<And>i. open (A i)"
"\<And>i. c \<in> A i"
"\<And>S. open S \<Longrightarrow> c \<in> S \<Longrightarrow> eventually (\<lambda>i. A i \<subseteq> S) sequentially"
by (rule countable_basis_at_decseq) blast
have [measurable]: "\<And>N i. (f N)-`(A i) \<inter> space M \<in> sets M" using A(1) by auto
then have mes: "(\<Inter>i. \<Union>n. \<Inter>N\<in>{n..}. (f N)-`(A i) \<inter> space M) \<in> sets M" by blast
have "(u \<longlonglongrightarrow> c) \<longleftrightarrow> (\<forall>i. eventually (\<lambda>n. u n \<in> A i) sequentially)" for u::"nat \<Rightarrow> 'b"
proof
assume "u \<longlonglongrightarrow> c"
then have "eventually (\<lambda>n. u n \<in> A i) sequentially" for i using A(1)[of i] A(2)[of i]
by (simp add: topological_tendstoD)
then show "(\<forall>i. eventually (\<lambda>n. u n \<in> A i) sequentially)" by auto
next
assume H: "(\<forall>i. eventually (\<lambda>n. u n \<in> A i) sequentially)"
show "(u \<longlonglongrightarrow> c)"
proof (rule topological_tendstoI)
fix S assume "open S" "c \<in> S"
with A(3)[OF this] obtain i where "A i \<subseteq> S"
using eventually_False_sequentially eventually_mono by blast
moreover have "eventually (\<lambda>n. u n \<in> A i) sequentially" using H by simp
ultimately show "\<forall>\<^sub>F n in sequentially. u n \<in> S"
by (simp add: eventually_mono subset_eq)
qed
qed
then have "{x. (\<lambda>n. f n x) \<longlonglongrightarrow> c} = (\<Inter>i. \<Union>n. \<Inter>N\<in>{n..}. (f N)-`(A i))"
by (auto simp add: atLeast_def eventually_at_top_linorder)
then have "{x \<in> space M. (\<lambda>n. f n x) \<longlonglongrightarrow> c} = (\<Inter>i. \<Union>n. \<Inter>N\<in>{n..}. (f N)-`(A i) \<inter> space M)"
by auto
then have "{x \<in> space M. (\<lambda>n. f n x) \<longlonglongrightarrow> c} \<in> sets M" using mes by simp
then show ?thesis by auto
qed
lemma measurable_limit2 [measurable]:
fixes u::"nat \<Rightarrow> 'a \<Rightarrow> real"
assumes [measurable]: "\<And>n. u n \<in> borel_measurable M" "v \<in> borel_measurable M"
shows "Measurable.pred M (\<lambda>x. (\<lambda>n. u n x) \<longlonglongrightarrow> v x)"
proof -
define w where "w = (\<lambda>n x. u n x - v x)"
have [measurable]: "w n \<in> borel_measurable M" for n unfolding w_def by auto
have "((\<lambda>n. u n x) \<longlonglongrightarrow> v x) \<longleftrightarrow> ((\<lambda>n. w n x) \<longlonglongrightarrow> 0)" for x
unfolding w_def using Lim_null by auto
then show ?thesis using measurable_limit by auto
qed
lemma measurable_P_restriction [measurable (raw)]:
assumes [measurable]: "Measurable.pred M P" "A \<in> sets M"
shows "{x \<in> A. P x} \<in> sets M"
proof -
have "A \<subseteq> space M" using sets.sets_into_space[OF assms(2)].
then have "{x \<in> A. P x} = A \<inter> {x \<in> space M. P x}" by blast
then show ?thesis by auto
qed
lemma measurable_sum_nat [measurable (raw)]:
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> nat"
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> measurable M (count_space UNIV)"
shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> measurable M (count_space UNIV)"
proof cases
assume "finite S"
then show ?thesis using assms by induct auto
qed simp
lemma measurable_abs_powr [measurable]:
fixes p::real
assumes [measurable]: "f \<in> borel_measurable M"
shows "(\<lambda>x. \<bar>f x\<bar> powr p) \<in> borel_measurable M"
by simp
text \<open>The next one is a variation around \<open>measurable_restrict_space\<close>.\<close>
lemma measurable_restrict_space3:
assumes "f \<in> measurable M N" and
"f \<in> A \<rightarrow> B"
shows "f \<in> measurable (restrict_space M A) (restrict_space N B)"
proof -
have "f \<in> measurable (restrict_space M A) N" using assms(1) measurable_restrict_space1 by auto
then show ?thesis by (metis Int_iff funcsetI funcset_mem
measurable_restrict_space2[of f, of "restrict_space M A", of B, of N] assms(2) space_restrict_space)
qed
lemma measurable_restrict_mono:
assumes f: "f \<in> restrict_space M A \<rightarrow>\<^sub>M N" and "B \<subseteq> A"
shows "f \<in> restrict_space M B \<rightarrow>\<^sub>M N"
by (rule measurable_compose[OF measurable_restrict_space3 f])
(insert \<open>B \<subseteq> A\<close>, auto)
text \<open>The next one is a variation around \<open>measurable_piecewise_restrict\<close>.\<close>
lemma measurable_piecewise_restrict2:
assumes [measurable]: "\<And>n. A n \<in> sets M"
and "space M = (\<Union>(n::nat). A n)"
"\<And>n. \<exists>h \<in> measurable M N. (\<forall>x \<in> A n. f x = h x)"
shows "f \<in> measurable M N"
proof (rule measurableI)
fix B assume [measurable]: "B \<in> sets N"
{
fix n::nat
obtain h where [measurable]: "h \<in> measurable M N" and "\<forall>x \<in> A n. f x = h x" using assms(3) by blast
then have *: "f-`B \<inter> A n = h-`B \<inter> A n" by auto
have "h-`B \<inter> A n = h-`B \<inter> space M \<inter> A n" using assms(2) sets.sets_into_space by auto
then have "h-`B \<inter> A n \<in> sets M" by simp
then have "f-`B \<inter> A n \<in> sets M" using * by simp
}
then have "(\<Union>n. f-`B \<inter> A n) \<in> sets M" by measurable
moreover have "f-`B \<inter> space M = (\<Union>n. f-`B \<inter> A n)" using assms(2) by blast
ultimately show "f-`B \<inter> space M \<in> sets M" by simp
next
fix x assume "x \<in> space M"
then obtain n where "x \<in> A n" using assms(2) by blast
obtain h where [measurable]: "h \<in> measurable M N" and "\<forall>x \<in> A n. f x = h x" using assms(3) by blast
then have "f x = h x" using \<open>x \<in> A n\<close> by blast
moreover have "h x \<in> space N" by (metis measurable_space \<open>x \<in> space M\<close> \<open>h \<in> measurable M N\<close>)
ultimately show "f x \<in> space N" by simp
qed
end
diff --git a/src/HOL/Analysis/Cartesian_Space.thy b/src/HOL/Analysis/Cartesian_Space.thy
--- a/src/HOL/Analysis/Cartesian_Space.thy
+++ b/src/HOL/Analysis/Cartesian_Space.thy
@@ -1,1442 +1,1442 @@
(* Title: HOL/Analysis/Cartesian_Space.thy
Author: Amine Chaieb, University of Cambridge
Author: Jose Divasón <jose.divasonm at unirioja.es>
Author: Jesús Aransay <jesus-maria.aransay at unirioja.es>
Author: Johannes Hölzl, VU Amsterdam
Author: Fabian Immler, TUM
*)
section "Linear Algebra on Finite Cartesian Products"
theory Cartesian_Space
imports
Finite_Cartesian_Product Linear_Algebra
begin
subsection\<^marker>\<open>tag unimportant\<close> \<open>Type @{typ \<open>'a ^ 'n\<close>} and fields as vector spaces\<close> (*much of the following
is really basic linear algebra, check for overlap? rename subsection? *)
definition "cart_basis = {axis i 1 | i. i\<in>UNIV}"
lemma finite_cart_basis: "finite (cart_basis)" unfolding cart_basis_def
using finite_Atleast_Atmost_nat by fastforce
lemma card_cart_basis: "card (cart_basis::('a::zero_neq_one^'i) set) = CARD('i)"
unfolding cart_basis_def Setcompr_eq_image
by (rule card_image) (auto simp: inj_on_def axis_eq_axis)
interpretation vec: vector_space "(*s) "
by unfold_locales (vector algebra_simps)+
lemma independent_cart_basis:
"vec.independent (cart_basis)"
proof (rule vec.independent_if_scalars_zero)
show "finite (cart_basis)" using finite_cart_basis .
fix f::"('a, 'b) vec \<Rightarrow> 'a" and x::"('a, 'b) vec"
assume eq_0: "(\<Sum>x\<in>cart_basis. f x *s x) = 0" and x_in: "x \<in> cart_basis"
obtain i where x: "x = axis i 1" using x_in unfolding cart_basis_def by auto
have sum_eq_0: "(\<Sum>x\<in>(cart_basis) - {x}. f x * (x $ i)) = 0"
proof (rule sum.neutral, rule ballI)
fix xa assume xa: "xa \<in> cart_basis - {x}"
obtain a where a: "xa = axis a 1" and a_not_i: "a \<noteq> i"
using xa x unfolding cart_basis_def by auto
have "xa $ i = 0" unfolding a axis_def using a_not_i by auto
thus "f xa * xa $ i = 0" by simp
qed
have "0 = (\<Sum>x\<in>cart_basis. f x *s x) $ i" using eq_0 by simp
also have "... = (\<Sum>x\<in>cart_basis. (f x *s x) $ i)" unfolding sum_component ..
also have "... = (\<Sum>x\<in>cart_basis. f x * (x $ i))" unfolding vector_smult_component ..
also have "... = f x * (x $ i) + (\<Sum>x\<in>(cart_basis) - {x}. f x * (x $ i))"
by (rule sum.remove[OF finite_cart_basis x_in])
also have "... = f x * (x $ i)" unfolding sum_eq_0 by simp
also have "... = f x" unfolding x axis_def by auto
finally show "f x = 0" ..
qed
lemma span_cart_basis:
"vec.span (cart_basis) = UNIV"
proof (auto)
fix x::"('a, 'b) vec"
let ?f="\<lambda>v. x $ (THE i. v = axis i 1)"
show "x \<in> vec.span (cart_basis)"
apply (unfold vec.span_finite[OF finite_cart_basis])
apply (rule image_eqI[of _ _ ?f])
apply (subst vec_eq_iff)
apply clarify
proof -
fix i::'b
let ?w = "axis i (1::'a)"
have the_eq_i: "(THE a. ?w = axis a 1) = i"
by (rule the_equality, auto simp: axis_eq_axis)
have sum_eq_0: "(\<Sum>v\<in>(cart_basis) - {?w}. x $ (THE i. v = axis i 1) * v $ i) = 0"
proof (rule sum.neutral, rule ballI)
fix xa::"('a, 'b) vec"
assume xa: "xa \<in> cart_basis - {?w}"
obtain j where j: "xa = axis j 1" and i_not_j: "i \<noteq> j" using xa unfolding cart_basis_def by auto
have the_eq_j: "(THE i. xa = axis i 1) = j"
proof (rule the_equality)
show "xa = axis j 1" using j .
show "\<And>i. xa = axis i 1 \<Longrightarrow> i = j" by (metis axis_eq_axis j zero_neq_one)
qed
show "x $ (THE i. xa = axis i 1) * xa $ i = 0"
apply (subst (2) j)
unfolding the_eq_j unfolding axis_def using i_not_j by simp
qed
have "(\<Sum>v\<in>cart_basis. x $ (THE i. v = axis i 1) *s v) $ i =
(\<Sum>v\<in>cart_basis. (x $ (THE i. v = axis i 1) *s v) $ i)" unfolding sum_component ..
also have "... = (\<Sum>v\<in>cart_basis. x $ (THE i. v = axis i 1) * v $ i)"
unfolding vector_smult_component ..
also have "... = x $ (THE a. ?w = axis a 1) * ?w $ i + (\<Sum>v\<in>(cart_basis) - {?w}. x $ (THE i. v = axis i 1) * v $ i)"
by (rule sum.remove[OF finite_cart_basis], auto simp add: cart_basis_def)
also have "... = x $ (THE a. ?w = axis a 1) * ?w $ i" unfolding sum_eq_0 by simp
also have "... = x $ i" unfolding the_eq_i unfolding axis_def by auto
finally show "x $ i = (\<Sum>v\<in>cart_basis. x $ (THE i. v = axis i 1) *s v) $ i" by simp
qed simp
qed
(*Some interpretations:*)
interpretation vec: finite_dimensional_vector_space "(*s)" "cart_basis"
by (unfold_locales, auto simp add: finite_cart_basis independent_cart_basis span_cart_basis)
lemma matrix_vector_mul_linear_gen[intro, simp]:
"Vector_Spaces.linear (*s) (*s) ((*v) A)"
by unfold_locales
(vector matrix_vector_mult_def sum.distrib algebra_simps)+
lemma span_vec_eq: "vec.span X = span X"
and dim_vec_eq: "vec.dim X = dim X"
and dependent_vec_eq: "vec.dependent X = dependent X"
and subspace_vec_eq: "vec.subspace X = subspace X"
for X::"(real^'n) set"
unfolding span_raw_def dim_raw_def dependent_raw_def subspace_raw_def
by (auto simp: scalar_mult_eq_scaleR)
lemma linear_componentwise:
fixes f:: "'a::field ^'m \<Rightarrow> 'a ^ 'n"
assumes lf: "Vector_Spaces.linear (*s) (*s) f"
shows "(f x)$j = sum (\<lambda>i. (x$i) * (f (axis i 1)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs")
proof -
interpret lf: Vector_Spaces.linear "(*s)" "(*s)" f
using lf .
let ?M = "(UNIV :: 'm set)"
let ?N = "(UNIV :: 'n set)"
have fM: "finite ?M" by simp
have "?rhs = (sum (\<lambda>i. (x$i) *s (f (axis i 1))) ?M)$j"
unfolding sum_component by simp
then show ?thesis
unfolding lf.sum[symmetric] lf.scale[symmetric]
unfolding basis_expansion by auto
qed
interpretation vec: Vector_Spaces.linear "(*s)" "(*s)" "(*v) A"
using matrix_vector_mul_linear_gen.
interpretation vec: finite_dimensional_vector_space_pair "(*s)" cart_basis "(*s)" cart_basis ..
lemma matrix_works:
assumes lf: "Vector_Spaces.linear (*s) (*s) f"
shows "matrix f *v x = f (x::'a::field ^ '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_of_matrix_vector_mul[simp]: "matrix(\<lambda>x. A *v (x :: 'a::field ^ 'n)) = A"
by (simp add: matrix_eq matrix_works)
lemma matrix_compose_gen:
assumes lf: "Vector_Spaces.linear (*s) (*s) (f::'a::{field}^'n \<Rightarrow> 'a^'m)"
and lg: "Vector_Spaces.linear (*s) (*s) (g::'a^'m \<Rightarrow> 'a^_)"
shows "matrix (g o f) = matrix g ** matrix f"
using lf lg Vector_Spaces.linear_compose[OF lf lg] matrix_works[OF Vector_Spaces.linear_compose[OF lf lg]]
by (simp add: matrix_eq matrix_works matrix_vector_mul_assoc[symmetric] o_def)
lemma matrix_compose:
assumes "linear (f::real^'n \<Rightarrow> real^'m)" "linear (g::real^'m \<Rightarrow> real^_)"
shows "matrix (g o f) = matrix g ** matrix f"
using matrix_compose_gen[of f g] assms
by (simp add: linear_def scalar_mult_eq_scaleR)
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 linear_matrix_vector_mul_eq:
"Vector_Spaces.linear (*s) (*s) f \<longleftrightarrow> linear (f :: real^'n \<Rightarrow> real ^'m)"
by (simp add: scalar_mult_eq_scaleR linear_def)
lemma matrix_vector_mul[simp]:
"Vector_Spaces.linear (*s) (*s) g \<Longrightarrow> (\<lambda>y. matrix g *v y) = g"
"linear f \<Longrightarrow> (\<lambda>x. matrix f *v x) = f"
"bounded_linear f \<Longrightarrow> (\<lambda>x. matrix f *v x) = f"
for f :: "real^'n \<Rightarrow> real ^'m"
by (simp_all add: ext matrix_works linear_matrix_vector_mul_eq linear_linear)
lemma matrix_left_invertible_injective:
fixes A :: "'a::field^'n^'m"
shows "(\<exists>B. B ** A = mat 1) \<longleftrightarrow> inj ((*v) A)"
proof safe
fix B
assume B: "B ** A = mat 1"
show "inj ((*v) A)"
unfolding inj_on_def
by (metis B matrix_vector_mul_assoc matrix_vector_mul_lid)
next
assume "inj ((*v) A)"
from vec.linear_injective_left_inverse[OF matrix_vector_mul_linear_gen this]
obtain g where "Vector_Spaces.linear (*s) (*s) g" and g: "g \<circ> (*v) A = id"
by blast
have "matrix g ** A = mat 1"
by (metis matrix_vector_mul_linear_gen \<open>Vector_Spaces.linear (*s) (*s) g\<close> g matrix_compose_gen
matrix_eq matrix_id_mat_1 matrix_vector_mul(1))
then show "\<exists>B. B ** A = mat 1"
by metis
qed
lemma matrix_left_invertible_ker:
"(\<exists>B. (B::'a::{field} ^'m^'n) ** (A::'a::{field}^'n^'m) = mat 1) \<longleftrightarrow> (\<forall>x. A *v x = 0 \<longrightarrow> x = 0)"
unfolding matrix_left_invertible_injective
using vec.inj_on_iff_eq_0[OF vec.subspace_UNIV, of A]
by (simp add: inj_on_def)
lemma matrix_right_invertible_surjective:
"(\<exists>B. (A::'a::field^'n^'m) ** (B::'a::field^'m^'n) = mat 1) \<longleftrightarrow> surj (\<lambda>x. A *v x)"
proof -
{ fix B :: "'a ^'m^'n"
assume AB: "A ** B = mat 1"
{ fix x :: "'a ^ 'm"
have "A *v (B *v x) = x"
by (simp add: matrix_vector_mul_assoc AB) }
hence "surj ((*v) A)" unfolding surj_def by metis }
moreover
{ assume sf: "surj ((*v) A)"
from vec.linear_surjective_right_inverse[OF _ this]
obtain g:: "'a ^'m \<Rightarrow> 'a ^'n" where g: "Vector_Spaces.linear (*s) (*s) g" "(*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::'a^'m^'n) = mat 1" by blast
}
ultimately show ?thesis unfolding surj_def by blast
qed
lemma matrix_left_invertible_independent_columns:
fixes A :: "'a::{field}^'n^'m"
shows "(\<exists>(B::'a ^'m^'n). B ** A = mat 1) \<longleftrightarrow>
(\<forall>c. sum (\<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: "sum (\<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
by (vector matrix_mult_sum)
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 ):: 'a)"
from H[rule_format, of ?c, unfolded matrix_mult_sum[symmetric], OF x]
have "x = 0" by vector }
}
ultimately show ?thesis unfolding matrix_left_invertible_ker by auto
qed
lemma matrix_right_invertible_independent_rows:
fixes A :: "'a::{field}^'n^'m"
shows "(\<exists>(B::'a^'m^'n). A ** B = mat 1) \<longleftrightarrow>
(\<forall>c. sum (\<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:)
lemma matrix_right_invertible_span_columns:
"(\<exists>(B::'a::field ^'n^'m). (A::'a ^'m^'n) ** B = mat 1) \<longleftrightarrow>
vec.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::'a^'m). sum (\<lambda>i. (x$i) *s column i A) ?U = y)"
unfolding matrix_right_invertible_surjective matrix_mult_sum surj_def
by (simp add: eq_commute)
have rhseq: "?rhs \<longleftrightarrow> (\<forall>x. x \<in> vec.span (columns A))" by blast
{ assume h: ?lhs
{ fix x:: "'a ^'n"
from h[unfolded lhseq, rule_format, of x] obtain y :: "'a ^'m"
where y: "sum (\<lambda>i. (y$i) *s column i A) ?U = x" by blast
have "x \<in> vec.span (columns A)"
unfolding y[symmetric] scalar_mult_eq_scaleR
proof (rule vec.span_sum [OF vec.span_scale])
show "column i A \<in> vec.span (columns A)" for i
using columns_def vec.span_superset by auto
qed
}
then have ?rhs unfolding rhseq by blast }
moreover
{ assume h:?rhs
let ?P = "\<lambda>(y::'a ^'n). \<exists>(x::'a^'m). sum (\<lambda>i. (x$i) *s column i A) ?U = y"
{ fix y
have "y \<in> vec.span (columns A)"
unfolding h by blast
then have "?P y"
proof (induction rule: vec.span_induct_alt)
case base
then show ?case
by (metis (full_types) matrix_mult_sum matrix_vector_mult_0_right)
next
case (step c y1 y2)
from step obtain i where i: "i \<in> ?U" "y1 = column i A"
unfolding columns_def by blast
obtain x:: "'a ^'m" where x: "sum (\<lambda>i. (x$i) *s column i A) ?U = y2"
using step by blast
let ?x = "(\<chi> j. if j = i then c + (x$i) else (x$j))::'a^'m"
show ?case
proof (rule exI[where x= "?x"], vector, auto simp add: i x[symmetric] if_distrib distrib_left if_distribR 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 "sum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
else (x$xa) * ((column xa A$j))) ?U = sum (\<lambda>xa. (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))) ?U"
by (rule sum.cong[OF refl]) (use th in blast)
also have "\<dots> = sum (\<lambda>xa. if xa = i then c * ((column i A)$j) else 0) ?U + sum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"
by (simp add: sum.distrib)
also have "\<dots> = c * ((column i A)$j) + sum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"
unfolding sum.delta[OF fU]
using i(1) by simp
finally show "sum (\<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) + sum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U" .
qed
qed
}
then have ?lhs unfolding lhseq ..
}
ultimately show ?thesis by blast
qed
lemma matrix_left_invertible_span_rows_gen:
"(\<exists>(B::'a^'m^'n). B ** (A::'a::field^'n^'m) = mat 1) \<longleftrightarrow> vec.span (rows A) = UNIV"
unfolding right_invertible_transpose[symmetric]
unfolding columns_transpose[symmetric]
unfolding matrix_right_invertible_span_columns
..
lemma matrix_left_invertible_span_rows:
"(\<exists>(B::real^'m^'n). B ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> span (rows A) = UNIV"
using matrix_left_invertible_span_rows_gen[of A] by (simp add: span_vec_eq)
lemma matrix_left_right_inverse:
fixes A A' :: "'a::{field}^'n^'n"
shows "A ** A' = mat 1 \<longleftrightarrow> A' ** A = mat 1"
proof -
{ fix A A' :: "'a ^'n^'n"
assume AA': "A ** A' = mat 1"
have sA: "surj ((*v) A)"
using AA' matrix_right_invertible_surjective by auto
from vec.linear_surjective_isomorphism[OF matrix_vector_mul_linear_gen sA]
obtain f' :: "'a ^'n \<Rightarrow> 'a ^'n"
where f': "Vector_Spaces.linear (*s) (*s) 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] 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')
hence "matrix f' ** A = A' ** A" by simp
hence "A' ** A = mat 1" by (simp add: th)
}
then show ?thesis by blast
qed
lemma invertible_left_inverse:
fixes A :: "'a::{field}^'n^'n"
shows "invertible A \<longleftrightarrow> (\<exists>(B::'a^'n^'n). B ** A = mat 1)"
by (metis invertible_def matrix_left_right_inverse)
lemma invertible_right_inverse:
fixes A :: "'a::{field}^'n^'n"
shows "invertible A \<longleftrightarrow> (\<exists>(B::'a^'n^'n). A** B = mat 1)"
by (metis invertible_def matrix_left_right_inverse)
lemma invertible_mult:
assumes inv_A: "invertible A"
and inv_B: "invertible B"
shows "invertible (A**B)"
proof -
obtain A' where AA': "A ** A' = mat 1" and A'A: "A' ** A = mat 1"
using inv_A unfolding invertible_def by blast
obtain B' where BB': "B ** B' = mat 1" and B'B: "B' ** B = mat 1"
using inv_B unfolding invertible_def by blast
show ?thesis
proof (unfold invertible_def, rule exI[of _ "B'**A'"], rule conjI)
have "A ** B ** (B' ** A') = A ** (B ** (B' ** A'))"
using matrix_mul_assoc[of A B "(B' ** A')", symmetric] .
also have "... = A ** (B ** B' ** A')" unfolding matrix_mul_assoc[of B "B'" "A'"] ..
also have "... = A ** (mat 1 ** A')" unfolding BB' ..
also have "... = A ** A'" unfolding matrix_mul_lid ..
also have "... = mat 1" unfolding AA' ..
finally show "A ** B ** (B' ** A') = mat (1::'a)" .
have "B' ** A' ** (A ** B) = B' ** (A' ** (A ** B))" using matrix_mul_assoc[of B' A' "(A ** B)", symmetric] .
also have "... = B' ** (A' ** A ** B)" unfolding matrix_mul_assoc[of A' A B] ..
also have "... = B' ** (mat 1 ** B)" unfolding A'A ..
also have "... = B' ** B" unfolding matrix_mul_lid ..
also have "... = mat 1" unfolding B'B ..
finally show "B' ** A' ** (A ** B) = mat 1" .
qed
qed
lemma transpose_invertible:
fixes A :: "real^'n^'n"
assumes "invertible A"
shows "invertible (transpose A)"
by (meson assms invertible_def matrix_left_right_inverse right_invertible_transpose)
lemma vector_matrix_mul_assoc:
fixes v :: "('a::comm_semiring_1)^'n"
shows "(v v* M) v* N = v v* (M ** N)"
proof -
from matrix_vector_mul_assoc
have "transpose N *v (transpose M *v v) = (transpose N ** transpose M) *v v" by fast
thus "(v v* M) v* N = v v* (M ** N)"
by (simp add: matrix_transpose_mul [symmetric])
qed
lemma matrix_scaleR_vector_ac:
fixes A :: "real^('m::finite)^'n"
shows "A *v (k *\<^sub>R v) = k *\<^sub>R A *v v"
by (metis matrix_vector_mult_scaleR transpose_scalar vector_scaleR_matrix_ac vector_transpose_matrix)
lemma scaleR_matrix_vector_assoc:
fixes A :: "real^('m::finite)^'n"
shows "k *\<^sub>R (A *v v) = k *\<^sub>R A *v v"
by (metis matrix_scaleR_vector_ac matrix_vector_mult_scaleR)
(*Finally, some interesting theorems and interpretations that don't appear in any file of the
library.*)
locale linear_first_finite_dimensional_vector_space =
l?: Vector_Spaces.linear scaleB scaleC f +
B?: finite_dimensional_vector_space scaleB BasisB
for scaleB :: "('a::field => 'b::ab_group_add => 'b)" (infixr "*b" 75)
and scaleC :: "('a => 'c::ab_group_add => 'c)" (infixr "*c" 75)
and BasisB :: "('b set)"
and f :: "('b=>'c)"
lemma vec_dim_card: "vec.dim (UNIV::('a::{field}^'n) set) = CARD ('n)"
proof -
let ?f="\<lambda>i::'n. axis i (1::'a)"
have "vec.dim (UNIV::('a::{field}^'n) set) = card (cart_basis::('a^'n) set)"
unfolding vec.dim_UNIV ..
also have "... = card ({i. i\<in> UNIV}::('n) set)"
proof (rule bij_betw_same_card[of ?f, symmetric], unfold bij_betw_def, auto)
show "inj (\<lambda>i::'n. axis i (1::'a))" by (simp add: inj_on_def axis_eq_axis)
fix i::'n
show "axis i 1 \<in> cart_basis" unfolding cart_basis_def by auto
fix x::"'a^'n"
assume "x \<in> cart_basis"
thus "x \<in> range (\<lambda>i. axis i 1)" unfolding cart_basis_def by auto
qed
also have "... = CARD('n)" by auto
finally show ?thesis .
qed
interpretation vector_space_over_itself: vector_space "(*) :: 'a::field \<Rightarrow> 'a \<Rightarrow> 'a"
by unfold_locales (simp_all add: algebra_simps)
lemmas [simp del] = vector_space_over_itself.scale_scale
interpretation vector_space_over_itself: finite_dimensional_vector_space
"(*) :: 'a::field => 'a => 'a" "{1}"
by unfold_locales (auto simp: vector_space_over_itself.span_singleton)
lemma dimension_eq_1[code_unfold]: "vector_space_over_itself.dimension TYPE('a::field)= 1"
unfolding vector_space_over_itself.dimension_def by simp
lemma dim_subset_UNIV_cart_gen:
fixes S :: "('a::field^'n) set"
shows "vec.dim S \<le> CARD('n)"
by (metis vec.dim_eq_full vec.dim_subset_UNIV vec.span_UNIV vec_dim_card)
lemma dim_subset_UNIV_cart:
fixes S :: "(real^'n) set"
shows "dim S \<le> CARD('n)"
using dim_subset_UNIV_cart_gen[of S] by (simp add: dim_vec_eq)
text\<open>Two sometimes fruitful ways of looking at matrix-vector multiplication.\<close>
lemma matrix_mult_dot: "A *v x = (\<chi> i. inner (A$i) x)"
by (simp add: matrix_vector_mult_def inner_vec_def)
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
sum_distrib_right sum_distrib_left)
apply (subst sum.swap)
apply (simp add: ac_simps)
done
lemma matrix_adjoint: assumes lf: "linear (f :: real^'n \<Rightarrow> real ^'m)"
shows "matrix(adjoint f) = transpose(matrix f)"
proof -
have "matrix(adjoint f) = matrix(adjoint ((*v) (matrix f)))"
by (simp add: lf)
also have "\<dots> = transpose(matrix f)"
unfolding adjoint_matrix matrix_of_matrix_vector_mul
apply rule
done
finally show ?thesis .
qed
subsection\<open> Rank of a matrix\<close>
text\<open>Equivalence of row and column rank is taken from George Mackiw's paper, Mathematics Magazine 1995, p. 285.\<close>
lemma matrix_vector_mult_in_columnspace_gen:
fixes A :: "'a::field^'n^'m"
shows "(A *v x) \<in> vec.span(columns A)"
apply (simp add: matrix_vector_column columns_def transpose_def column_def)
apply (intro vec.span_sum vec.span_scale)
apply (force intro: vec.span_base)
done
lemma matrix_vector_mult_in_columnspace:
fixes A :: "real^'n^'m"
shows "(A *v x) \<in> span(columns A)"
using matrix_vector_mult_in_columnspace_gen[of A x] by (simp add: span_vec_eq)
lemma subspace_orthogonal_to_vector: "subspace {y. orthogonal x y}"
by (simp add: subspace_def orthogonal_clauses)
lemma orthogonal_nullspace_rowspace:
fixes A :: "real^'n^'m"
assumes 0: "A *v x = 0" and y: "y \<in> span(rows A)"
shows "orthogonal x y"
using y
proof (induction rule: span_induct)
case base
then show ?case
by (simp add: subspace_orthogonal_to_vector)
next
case (step v)
then obtain i where "v = row i A"
by (auto simp: rows_def)
with 0 show ?case
unfolding orthogonal_def inner_vec_def matrix_vector_mult_def row_def
by (simp add: mult.commute) (metis (no_types) vec_lambda_beta zero_index)
qed
lemma nullspace_inter_rowspace:
fixes A :: "real^'n^'m"
shows "A *v x = 0 \<and> x \<in> span(rows A) \<longleftrightarrow> x = 0"
using orthogonal_nullspace_rowspace orthogonal_self span_zero matrix_vector_mult_0_right
by blast
lemma matrix_vector_mul_injective_on_rowspace:
fixes A :: "real^'n^'m"
shows "\<lbrakk>A *v x = A *v y; x \<in> span(rows A); y \<in> span(rows A)\<rbrakk> \<Longrightarrow> x = y"
using nullspace_inter_rowspace [of A "x-y"]
by (metis diff_eq_diff_eq diff_self matrix_vector_mult_diff_distrib span_diff)
definition\<^marker>\<open>tag important\<close> rank :: "'a::field^'n^'m=>nat"
where row_rank_def_gen: "rank A \<equiv> vec.dim(rows A)"
lemma row_rank_def: "rank A = dim (rows A)" for A::"real^'n^'m"
by (auto simp: row_rank_def_gen dim_vec_eq)
lemma dim_rows_le_dim_columns:
fixes A :: "real^'n^'m"
shows "dim(rows A) \<le> dim(columns A)"
proof -
have "dim (span (rows A)) \<le> dim (span (columns A))"
proof -
obtain B where "independent B" "span(rows A) \<subseteq> span B"
and B: "B \<subseteq> span(rows A)""card B = dim (span(rows A))"
using basis_exists [of "span(rows A)"] by metis
with span_subspace have eq: "span B = span(rows A)"
by auto
then have inj: "inj_on ((*v) A) (span B)"
by (simp add: inj_on_def matrix_vector_mul_injective_on_rowspace)
then have ind: "independent ((*v) A ` B)"
by (rule linear_independent_injective_image [OF Finite_Cartesian_Product.matrix_vector_mul_linear \<open>independent B\<close>])
have "dim (span (rows A)) \<le> card ((*v) A ` B)"
unfolding B(2)[symmetric]
using inj
by (auto simp: card_image inj_on_subset span_superset)
also have "\<dots> \<le> dim (span (columns A))"
using _ ind
by (rule independent_card_le_dim) (auto intro!: matrix_vector_mult_in_columnspace)
finally show ?thesis .
qed
then show ?thesis
by (simp)
qed
lemma column_rank_def:
fixes A :: "real^'n^'m"
shows "rank A = dim(columns A)"
unfolding row_rank_def
by (metis columns_transpose dim_rows_le_dim_columns le_antisym rows_transpose)
lemma rank_transpose:
fixes A :: "real^'n^'m"
shows "rank(transpose A) = rank A"
by (metis column_rank_def row_rank_def rows_transpose)
lemma matrix_vector_mult_basis:
fixes A :: "real^'n^'m"
shows "A *v (axis k 1) = column k A"
by (simp add: cart_eq_inner_axis column_def matrix_mult_dot)
lemma columns_image_basis:
fixes A :: "real^'n^'m"
shows "columns A = (*v) A ` (range (\<lambda>i. axis i 1))"
by (force simp: columns_def matrix_vector_mult_basis [symmetric])
lemma rank_dim_range:
fixes A :: "real^'n^'m"
shows "rank A = dim(range (\<lambda>x. A *v x))"
unfolding column_rank_def
proof (rule span_eq_dim)
have "span (columns A) \<subseteq> span (range ((*v) A))" (is "?l \<subseteq> ?r")
by (simp add: columns_image_basis image_subsetI span_mono)
then show "?l = ?r"
by (metis (no_types, lifting) image_subset_iff matrix_vector_mult_in_columnspace
span_eq span_span)
qed
lemma rank_bound:
fixes A :: "real^'n^'m"
shows "rank A \<le> min CARD('m) (CARD('n))"
by (metis (mono_tags, lifting) dim_subset_UNIV_cart min.bounded_iff
column_rank_def row_rank_def)
lemma full_rank_injective:
fixes A :: "real^'n^'m"
shows "rank A = CARD('n) \<longleftrightarrow> inj ((*v) A)"
by (simp add: matrix_left_invertible_injective [symmetric] matrix_left_invertible_span_rows row_rank_def
dim_eq_full [symmetric] card_cart_basis vec.dimension_def)
lemma full_rank_surjective:
fixes A :: "real^'n^'m"
shows "rank A = CARD('m) \<longleftrightarrow> surj ((*v) A)"
by (simp add: matrix_right_invertible_surjective [symmetric] left_invertible_transpose [symmetric]
matrix_left_invertible_injective full_rank_injective [symmetric] rank_transpose)
lemma rank_I: "rank(mat 1::real^'n^'n) = CARD('n)"
by (simp add: full_rank_injective inj_on_def)
lemma less_rank_noninjective:
fixes A :: "real^'n^'m"
shows "rank A < CARD('n) \<longleftrightarrow> \<not> inj ((*v) A)"
using less_le rank_bound by (auto simp: full_rank_injective [symmetric])
lemma matrix_nonfull_linear_equations_eq:
fixes A :: "real^'n^'m"
shows "(\<exists>x. (x \<noteq> 0) \<and> A *v x = 0) \<longleftrightarrow> rank A \<noteq> CARD('n)"
by (meson matrix_left_invertible_injective full_rank_injective matrix_left_invertible_ker)
lemma rank_eq_0: "rank A = 0 \<longleftrightarrow> A = 0" and rank_0 [simp]: "rank (0::real^'n^'m) = 0"
for A :: "real^'n^'m"
by (auto simp: rank_dim_range matrix_eq)
lemma rank_mul_le_right:
fixes A :: "real^'n^'m" and B :: "real^'p^'n"
shows "rank(A ** B) \<le> rank B"
proof -
have "rank(A ** B) \<le> dim ((*v) A ` range ((*v) B))"
by (auto simp: rank_dim_range image_comp o_def matrix_vector_mul_assoc)
also have "\<dots> \<le> rank B"
by (simp add: rank_dim_range dim_image_le)
finally show ?thesis .
qed
lemma rank_mul_le_left:
fixes A :: "real^'n^'m" and B :: "real^'p^'n"
shows "rank(A ** B) \<le> rank A"
by (metis matrix_transpose_mul rank_mul_le_right rank_transpose)
subsection\<^marker>\<open>tag unimportant\<close> \<open>Lemmas for working on \<open>real^1/2/3/4\<close>\<close>
lemma exhaust_2:
fixes x :: 2
shows "x = 1 \<or> x = 2"
proof (induct x)
case (of_int z)
then have "0 \<le> 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 \<le> 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 exhaust_4:
fixes x :: 4
shows "x = 1 \<or> x = 2 \<or> x = 3 \<or> x = 4"
proof (induct x)
case (of_int z)
then have "0 \<le> z" and "z < 4" by simp_all
then have "z = 0 \<or> z = 1 \<or> z = 2 \<or> z = 3" by arith
then show ?case by auto
qed
lemma forall_4: "(\<forall>i::4. P i) \<longleftrightarrow> P 1 \<and> P 2 \<and> P 3 \<and> P 4"
by (metis exhaust_4)
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 UNIV_4: "UNIV = {1::4, 2::4, 3::4, 4::4}"
using exhaust_4 by auto
lemma sum_1: "sum f (UNIV::1 set) = f 1"
unfolding UNIV_1 by simp
lemma sum_2: "sum f (UNIV::2 set) = f 1 + f 2"
unfolding UNIV_2 by simp
lemma sum_3: "sum f (UNIV::3 set) = f 1 + f 2 + f 3"
unfolding UNIV_3 by (simp add: ac_simps)
lemma sum_4: "sum f (UNIV::4 set) = f 1 + f 2 + f 3 + f 4"
unfolding UNIV_4 by (simp add: ac_simps)
subsection\<^marker>\<open>tag unimportant\<close>\<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 dist_vector_1:
fixes x :: "'a::real_normed_vector^1"
shows "dist x y = dist (x$1) (y$1)"
by (simp add: dist_norm norm_vector_1)
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\<^marker>\<open>tag unimportant\<close>\<open>Routine results connecting the types \<^typ>\<open>real^1\<close> and \<^typ>\<open>real\<close>\<close>
lemma vector_one_nth [simp]:
fixes x :: "'a^1" shows "vec (x $ 1) = x"
by (metis vec_def vector_one)
lemma tendsto_at_within_vector_1:
fixes S :: "'a :: metric_space set"
assumes "(f \<longlongrightarrow> fx) (at x within S)"
shows "((\<lambda>y::'a^1. \<chi> i. f (y $ 1)) \<longlongrightarrow> (vec fx::'a^1)) (at (vec x) within vec ` S)"
proof (rule topological_tendstoI)
fix T :: "('a^1) set"
assume "open T" "vec fx \<in> T"
have "\<forall>\<^sub>F x in at x within S. f x \<in> (\<lambda>x. x $ 1) ` T"
using \<open>open T\<close> \<open>vec fx \<in> T\<close> assms open_image_vec_nth tendsto_def by fastforce
then show "\<forall>\<^sub>F x::'a^1 in at (vec x) within vec ` S. (\<chi> i. f (x $ 1)) \<in> T"
unfolding eventually_at dist_norm [symmetric]
by (rule ex_forward)
(use \<open>open T\<close> in
\<open>fastforce simp: dist_norm dist_vec_def L2_set_def image_iff vector_one open_vec_def\<close>)
qed
lemma has_derivative_vector_1:
assumes der_g: "(g has_derivative (\<lambda>x. x * g' a)) (at a within S)"
shows "((\<lambda>x. vec (g (x $ 1))) has_derivative (*\<^sub>R) (g' a))
(at ((vec a)::real^1) within vec ` S)"
using der_g
apply (auto simp: Deriv.has_derivative_within bounded_linear_scaleR_right norm_vector_1)
apply (drule tendsto_at_within_vector_1, vector)
apply (auto simp: algebra_simps eventually_at tendsto_def)
done
subsection\<^marker>\<open>tag unimportant\<close>\<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 [simp]: "(vector[x]) $1 = x"
unfolding vector_def by simp
lemma vector_2 [simp]: "(vector[x,y]) $1 = x" "(vector[x,y] :: 'a^2)$2 = (y::'a::zero)"
unfolding vector_def by simp_all
lemma vector_3 [simp]:
"(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]))"
by (metis vector_1 vector_one)
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
subsection\<^marker>\<open>tag unimportant\<close> \<open>lambda skolemization on cartesian products\<close>
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
text \<open>The same result in terms of square matrices.\<close>
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 dim_substandard_cart: "vec.dim {x::'a::field^'n. \<forall>i. i \<notin> d \<longrightarrow> x$i = 0} = card d"
(is "vec.dim ?A = _")
proof (rule vec.dim_unique)
let ?B = "((\<lambda>x. axis x 1) ` d)"
have subset_basis: "?B \<subseteq> cart_basis"
by (auto simp: cart_basis_def)
show "?B \<subseteq> ?A"
by (auto simp: axis_def)
show "vec.independent ((\<lambda>x. axis x 1) ` d)"
using subset_basis
by (rule vec.independent_mono[OF vec.independent_Basis])
have "x \<in> vec.span ?B" if "\<forall>i. i \<notin> d \<longrightarrow> x $ i = 0" for x::"'a^'n"
proof -
have "finite ?B"
using subset_basis finite_cart_basis
by (rule finite_subset)
have "x = (\<Sum>i\<in>UNIV. x $ i *s axis i 1)"
by (rule basis_expansion[symmetric])
also have "\<dots> = (\<Sum>i\<in>d. (x $ i) *s axis i 1)"
by (rule sum.mono_neutral_cong_right) (auto simp: that)
also have "\<dots> \<in> vec.span ?B"
by (simp add: vec.span_sum vec.span_clauses)
finally show "x \<in> vec.span ?B" .
qed
then show "?A \<subseteq> vec.span ?B" by auto
qed (simp add: card_image inj_on_def axis_eq_axis)
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
by (auto simp add: fun_eq_iff vector_add_ldistrib diff_conv_add_uminus simp del: add_uminus_conv_diff)
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\<^marker>\<open>tag unimportant\<close> \<open>Explicit formulas for low dimensions\<close>
lemma prod_neutral_const: "prod f {(1::nat)..1} = f 1"
by simp
lemma prod_2: "prod f {(1::nat)..2} = f 1 * f 2"
by (simp add: eval_nat_numeral atLeastAtMostSuc_conv mult.commute)
lemma prod_3: "prod f {(1::nat)..3} = f 1 * f 2 * f 3"
by (simp add: eval_nat_numeral atLeastAtMostSuc_conv mult.commute)
subsection \<open>Orthogonality of a matrix\<close>
definition\<^marker>\<open>tag important\<close> "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)
proposition orthogonal_matrix_mul:
fixes A :: "real ^'n^'n"
assumes "orthogonal_matrix A" "orthogonal_matrix B"
shows "orthogonal_matrix(A ** B)"
using assms
by (simp add: orthogonal_matrix matrix_transpose_mul matrix_left_right_inverse matrix_mul_assoc)
proposition 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: "Vector_Spaces.linear (*s) (*s) f" and fd: "\<And>v w. f v \<bullet> f w = v \<bullet> w"
unfolding orthogonal_transformation_def orthogonal_matrix linear_def scalar_mult_eq_scaleR
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[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 sum.delta[OF fU] mat_def axis_def)
}
then have "orthogonal_matrix ?mf"
unfolding orthogonal_matrix
by vector
with lf have ?rhs
unfolding linear_def scalar_mult_eq_scaleR
by blast
}
moreover
{
assume lf: "Vector_Spaces.linear (*s) (*s) f" and om: "orthogonal_matrix ?mf"
from lf om have ?lhs
unfolding orthogonal_matrix_def norm_eq orthogonal_transformation
apply (simp only: matrix_works[OF lf, symmetric] dot_matrix_vector_mul)
apply (simp add: dot_matrix_product linear_def scalar_mult_eq_scaleR)
done
}
ultimately show ?thesis
by (auto simp: linear_def scalar_mult_eq_scaleR)
qed
subsection \<open>Finding an Orthogonal Matrix\<close>
text \<open>We can find an orthogonal matrix taking any unit vector to any other.\<close>
lemma orthogonal_matrix_transpose [simp]:
"orthogonal_matrix(transpose A) \<longleftrightarrow> orthogonal_matrix A"
by (auto simp: orthogonal_matrix_def)
lemma orthogonal_matrix_orthonormal_columns:
fixes A :: "real^'n^'n"
shows "orthogonal_matrix A \<longleftrightarrow>
(\<forall>i. norm(column i A) = 1) \<and>
(\<forall>i j. i \<noteq> j \<longrightarrow> orthogonal (column i A) (column j A))"
by (auto simp: orthogonal_matrix matrix_mult_transpose_dot_column vec_eq_iff mat_def norm_eq_1 orthogonal_def)
lemma orthogonal_matrix_orthonormal_rows:
fixes A :: "real^'n^'n"
shows "orthogonal_matrix A \<longleftrightarrow>
(\<forall>i. norm(row i A) = 1) \<and>
(\<forall>i j. i \<noteq> j \<longrightarrow> orthogonal (row i A) (row j A))"
using orthogonal_matrix_orthonormal_columns [of "transpose A"] by simp
proposition orthogonal_matrix_exists_basis:
fixes a :: "real^'n"
assumes "norm a = 1"
obtains A where "orthogonal_matrix A" "A *v (axis k 1) = a"
proof -
obtain S where "a \<in> S" "pairwise orthogonal S" and noS: "\<And>x. x \<in> S \<Longrightarrow> norm x = 1"
and "independent S" "card S = CARD('n)" "span S = UNIV"
using vector_in_orthonormal_basis assms by force
then obtain f0 where "bij_betw f0 (UNIV::'n set) S"
by (metis finite_class.finite_UNIV finite_same_card_bij finiteI_independent)
then obtain f where f: "bij_betw f (UNIV::'n set) S" and a: "a = f k"
using bij_swap_iff [of k "inv f0 a" f0]
by (metis UNIV_I \<open>a \<in> S\<close> bij_betw_inv_into_right bij_betw_swap_iff swap_apply(1))
show thesis
proof
have [simp]: "\<And>i. norm (f i) = 1"
using bij_betwE [OF \<open>bij_betw f UNIV S\<close>] by (blast intro: noS)
have [simp]: "\<And>i j. i \<noteq> j \<Longrightarrow> orthogonal (f i) (f j)"
using \<open>pairwise orthogonal S\<close> \<open>bij_betw f UNIV S\<close>
by (auto simp: pairwise_def bij_betw_def inj_on_def)
show "orthogonal_matrix (\<chi> i j. f j $ i)"
by (simp add: orthogonal_matrix_orthonormal_columns column_def)
show "(\<chi> i j. f j $ i) *v axis k 1 = a"
by (simp add: matrix_vector_mult_def axis_def a if_distrib cong: if_cong)
qed
qed
lemma orthogonal_transformation_exists_1:
fixes a b :: "real^'n"
assumes "norm a = 1" "norm b = 1"
obtains f where "orthogonal_transformation f" "f a = b"
proof -
obtain k::'n where True
by simp
obtain A B where AB: "orthogonal_matrix A" "orthogonal_matrix B" and eq: "A *v (axis k 1) = a" "B *v (axis k 1) = b"
using orthogonal_matrix_exists_basis assms by metis
let ?f = "\<lambda>x. (B ** transpose A) *v x"
show thesis
proof
show "orthogonal_transformation ?f"
by (subst orthogonal_transformation_matrix)
(auto simp: AB orthogonal_matrix_mul)
next
show "?f a = b"
using \<open>orthogonal_matrix A\<close> unfolding orthogonal_matrix_def
by (metis eq matrix_mul_rid matrix_vector_mul_assoc)
qed
qed
proposition orthogonal_transformation_exists:
fixes a b :: "real^'n"
assumes "norm a = norm b"
obtains f where "orthogonal_transformation f" "f a = b"
proof (cases "a = 0 \<or> b = 0")
case True
with assms show ?thesis
using that by force
next
case False
then obtain f where f: "orthogonal_transformation f" and eq: "f (a /\<^sub>R norm a) = (b /\<^sub>R norm b)"
by (auto intro: orthogonal_transformation_exists_1 [of "a /\<^sub>R norm a" "b /\<^sub>R norm b"])
show ?thesis
proof
interpret linear f
using f by (simp add: orthogonal_transformation_linear)
have "f a /\<^sub>R norm a = f (a /\<^sub>R norm a)"
by (simp add: scale)
also have "\<dots> = b /\<^sub>R norm a"
by (simp add: eq assms [symmetric])
finally show "f a = b"
using False by auto
qed (use f in auto)
qed
subsection \<open>Scaling and isometry\<close>
proposition scaling_linear:
fixes f :: "'a::real_inner \<Rightarrow> 'a::real_inner"
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
have "norm (f x) = c * norm x" for x
by (metis dist_0_norm f0 fd)
then have "f v \<bullet> f w = c\<^sup>2 * (v \<bullet> w)"
unfolding dot_norm_neg dist_norm[symmetric]
by (simp add: fd power2_eq_square field_simps)
}
then show ?thesis
unfolding linear_iff vector_eq[where 'a="'a"] scalar_mult_eq_scaleR
by (simp add: inner_add field_simps)
qed
lemma isometry_linear:
"f (0::'a::real_inner) = (0::'a) \<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>
proposition orthogonal_transformation_isometry:
"orthogonal_transformation f \<longleftrightarrow> f(0::'a::real_inner) = (0::'a) \<and> (\<forall>x y. dist(f x) (f y) = dist x y)"
unfolding orthogonal_transformation
apply (auto simp: linear_0 isometry_linear)
apply (metis (no_types, hide_lams) dist_norm linear_diff)
by (metis dist_0_norm)
text \<open>Can extend an isometry from unit sphere:\<close>
lemma isometry_sphere_extend:
fixes f:: "'a::real_inner \<Rightarrow> 'a"
assumes f1: "\<And>x. norm x = 1 \<Longrightarrow> norm (f x) = 1"
and fd1: "\<And>x y. \<lbrakk>norm x = 1; norm y = 1\<rbrakk> \<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' u v u' v' :: "'a"
assume H: "x = norm x *\<^sub>R u" "y = norm y *\<^sub>R v"
"x' = norm x *\<^sub>R u'" "y' = norm y *\<^sub>R v'"
and J: "norm u = 1" "norm u' = 1" "norm v = 1" "norm v' = 1" "norm(u' - v') = norm(u - v)"
then have *: "u \<bullet> v = u' \<bullet> v' + v' \<bullet> u' - v \<bullet> u "
by (simp add: norm_eq norm_eq_1 inner_add inner_diff)
have "norm (norm x *\<^sub>R u' - norm y *\<^sub>R v') = norm (norm x *\<^sub>R u - norm y *\<^sub>R v)"
using J by (simp add: norm_eq norm_eq_1 inner_diff * field_simps)
then have "norm(x' - y') = norm(x - y)"
using H by metis
}
note norm_eq = this
let ?g = "\<lambda>x. if x = 0 then 0 else norm x *\<^sub>R f (x /\<^sub>R norm x)"
have thfg: "?g x = f x" if "norm x = 1" for x
using that by auto
have thd: "dist (?g x) (?g y) = dist x y" for x y
proof (cases "x=0 \<or> y=0")
case False
show "dist (?g x) (?g y) = dist x y"
unfolding dist_norm
proof (rule norm_eq)
show "x = norm x *\<^sub>R (x /\<^sub>R norm x)" "y = norm y *\<^sub>R (y /\<^sub>R norm y)"
"norm (f (x /\<^sub>R norm x)) = 1" "norm (f (y /\<^sub>R norm y)) = 1"
using False f1 by auto
qed (use False in \<open>auto simp: field_simps intro: f1 fd1[unfolded dist_norm]\<close>)
qed (auto simp: f1)
show ?thesis
unfolding orthogonal_transformation_isometry
by (rule exI[where x= ?g]) (metis thfg thd)
qed
subsection\<open>Induction on matrix row operations\<close>
lemma induct_matrix_row_operations:
fixes P :: "real^'n^'n \<Rightarrow> bool"
assumes zero_row: "\<And>A i. row i A = 0 \<Longrightarrow> P A"
and diagonal: "\<And>A. (\<And>i j. i \<noteq> j \<Longrightarrow> A$i$j = 0) \<Longrightarrow> P A"
and swap_cols: "\<And>A m n. \<lbrakk>P A; m \<noteq> n\<rbrakk> \<Longrightarrow> P(\<chi> i j. A $ i $ Fun.swap m n id j)"
and row_op: "\<And>A m n c. \<lbrakk>P A; m \<noteq> n\<rbrakk>
\<Longrightarrow> P(\<chi> i. if i = m then row m A + c *\<^sub>R row n A else row i A)"
shows "P A"
proof -
have "P A" if "(\<And>i j. \<lbrakk>j \<in> -K; i \<noteq> j\<rbrakk> \<Longrightarrow> A$i$j = 0)" for A K
proof -
have "finite K"
by simp
then show ?thesis using that
proof (induction arbitrary: A rule: finite_induct)
case empty
with diagonal show ?case
by simp
next
case (insert k K)
note insertK = insert
have "P A" if kk: "A$k$k \<noteq> 0"
and 0: "\<And>i j. \<lbrakk>j \<in> - insert k K; i \<noteq> j\<rbrakk> \<Longrightarrow> A$i$j = 0"
"\<And>i. \<lbrakk>i \<in> -L; i \<noteq> k\<rbrakk> \<Longrightarrow> A$i$k = 0" for A L
proof -
have "finite L"
by simp
then show ?thesis using 0 kk
proof (induction arbitrary: A rule: finite_induct)
case (empty B)
show ?case
proof (rule insertK)
fix i j
assume "i \<in> - K" "j \<noteq> i"
show "B $ j $ i = 0"
using \<open>j \<noteq> i\<close> \<open>i \<in> - K\<close> empty
by (metis ComplD ComplI Compl_eq_Diff_UNIV Diff_empty UNIV_I insert_iff)
qed
next
case (insert l L B)
show ?case
proof (cases "k = l")
case True
with insert show ?thesis
by auto
next
case False
let ?C = "\<chi> i. if i = l then row l B - (B $ l $ k / B $ k $ k) *\<^sub>R row k B else row i B"
have 1: "\<lbrakk>j \<in> - insert k K; i \<noteq> j\<rbrakk> \<Longrightarrow> ?C $ i $ j = 0" for j i
by (auto simp: insert.prems(1) row_def)
have 2: "?C $ i $ k = 0"
if "i \<in> - L" "i \<noteq> k" for i
proof (cases "i=l")
case True
with that insert.prems show ?thesis
by (simp add: row_def)
next
case False
with that show ?thesis
by (simp add: insert.prems(2) row_def)
qed
have 3: "?C $ k $ k \<noteq> 0"
by (auto simp: insert.prems row_def \<open>k \<noteq> l\<close>)
have PC: "P ?C"
using insert.IH [OF 1 2 3] by auto
have eqB: "(\<chi> i. if i = l then row l ?C + (B $ l $ k / B $ k $ k) *\<^sub>R row k ?C else row i ?C) = B"
using \<open>k \<noteq> l\<close> by (simp add: vec_eq_iff row_def)
show ?thesis
using row_op [OF PC, of l k, where c = "B$l$k / B$k$k"] eqB \<open>k \<noteq> l\<close>
by (simp add: cong: if_cong)
qed
qed
qed
then have nonzero_hyp: "P A"
if kk: "A$k$k \<noteq> 0" and zeroes: "\<And>i j. j \<in> - insert k K \<and> i\<noteq>j \<Longrightarrow> A$i$j = 0" for A
by (auto simp: intro!: kk zeroes)
show ?case
proof (cases "row k A = 0")
case True
with zero_row show ?thesis by auto
next
case False
then obtain l where l: "A$k$l \<noteq> 0"
by (auto simp: row_def zero_vec_def vec_eq_iff)
show ?thesis
proof (cases "k = l")
case True
with l nonzero_hyp insert.prems show ?thesis
by blast
next
case False
have *: "A $ i $ Fun.swap k l id j = 0" if "j \<noteq> k" "j \<notin> K" "i \<noteq> j" for i j
using False l insert.prems that
by (auto simp: swap_def insert split: if_split_asm)
have "P (\<chi> i j. (\<chi> i j. A $ i $ Fun.swap k l id j) $ i $ Fun.swap k l id j)"
by (rule swap_cols [OF nonzero_hyp False]) (auto simp: l *)
moreover
have "(\<chi> i j. (\<chi> i j. A $ i $ Fun.swap k l id j) $ i $ Fun.swap k l id j) = A"
by (vector Fun.swap_def)
ultimately show ?thesis
by simp
qed
qed
qed
qed
then show ?thesis
by blast
qed
lemma induct_matrix_elementary:
fixes P :: "real^'n^'n \<Rightarrow> bool"
assumes mult: "\<And>A B. \<lbrakk>P A; P B\<rbrakk> \<Longrightarrow> P(A ** B)"
and zero_row: "\<And>A i. row i A = 0 \<Longrightarrow> P A"
and diagonal: "\<And>A. (\<And>i j. i \<noteq> j \<Longrightarrow> A$i$j = 0) \<Longrightarrow> P A"
and swap1: "\<And>m n. m \<noteq> n \<Longrightarrow> P(\<chi> i j. mat 1 $ i $ Fun.swap m n id j)"
and idplus: "\<And>m n c. m \<noteq> n \<Longrightarrow> P(\<chi> i j. if i = m \<and> j = n then c else of_bool (i = j))"
shows "P A"
proof -
have swap: "P (\<chi> i j. A $ i $ Fun.swap m n id j)" (is "P ?C")
if "P A" "m \<noteq> n" for A m n
proof -
have "A ** (\<chi> i j. mat 1 $ i $ Fun.swap m n id j) = ?C"
by (simp add: matrix_matrix_mult_def mat_def vec_eq_iff if_distrib sum.delta_remove)
then show ?thesis
using mult swap1 that by metis
qed
have row: "P (\<chi> i. if i = m then row m A + c *\<^sub>R row n A else row i A)" (is "P ?C")
if "P A" "m \<noteq> n" for A m n c
proof -
let ?B = "\<chi> i j. if i = m \<and> j = n then c else of_bool (i = j)"
have "?B ** A = ?C"
using \<open>m \<noteq> n\<close> unfolding matrix_matrix_mult_def row_def of_bool_def
by (auto simp: vec_eq_iff if_distrib [of "\<lambda>x. x * y" for y] sum.remove cong: if_cong)
then show ?thesis
by (rule subst) (auto simp: that mult idplus)
qed
show ?thesis
by (rule induct_matrix_row_operations [OF zero_row diagonal swap row])
qed
lemma induct_matrix_elementary_alt:
fixes P :: "real^'n^'n \<Rightarrow> bool"
assumes mult: "\<And>A B. \<lbrakk>P A; P B\<rbrakk> \<Longrightarrow> P(A ** B)"
and zero_row: "\<And>A i. row i A = 0 \<Longrightarrow> P A"
and diagonal: "\<And>A. (\<And>i j. i \<noteq> j \<Longrightarrow> A$i$j = 0) \<Longrightarrow> P A"
and swap1: "\<And>m n. m \<noteq> n \<Longrightarrow> P(\<chi> i j. mat 1 $ i $ Fun.swap m n id j)"
and idplus: "\<And>m n. m \<noteq> n \<Longrightarrow> P(\<chi> i j. of_bool (i = m \<and> j = n \<or> i = j))"
shows "P A"
proof -
have *: "P (\<chi> i j. if i = m \<and> j = n then c else of_bool (i = j))"
if "m \<noteq> n" for m n c
proof (cases "c = 0")
case True
with diagonal show ?thesis by auto
next
case False
then have eq: "(\<chi> i j. if i = m \<and> j = n then c else of_bool (i = j)) =
(\<chi> i j. if i = j then (if j = n then inverse c else 1) else 0) **
(\<chi> i j. of_bool (i = m \<and> j = n \<or> i = j)) **
(\<chi> i j. if i = j then if j = n then c else 1 else 0)"
using \<open>m \<noteq> n\<close>
apply (simp add: matrix_matrix_mult_def vec_eq_iff of_bool_def if_distrib [of "\<lambda>x. y * x" for y] cong: if_cong)
apply (simp add: if_if_eq_conj sum.neutral conj_commute cong: conj_cong)
done
show ?thesis
apply (subst eq)
apply (intro mult idplus that)
apply (auto intro: diagonal)
done
qed
show ?thesis
by (rule induct_matrix_elementary) (auto intro: assms *)
qed
lemma matrix_vector_mult_matrix_matrix_mult_compose:
"(*v) (A ** B) = (*v) A \<circ> (*v) B"
by (auto simp: matrix_vector_mul_assoc)
lemma induct_linear_elementary:
fixes f :: "real^'n \<Rightarrow> real^'n"
assumes "linear f"
and comp: "\<And>f g. \<lbrakk>linear f; linear g; P f; P g\<rbrakk> \<Longrightarrow> P(f \<circ> g)"
and zeroes: "\<And>f i. \<lbrakk>linear f; \<And>x. (f x) $ i = 0\<rbrakk> \<Longrightarrow> P f"
and const: "\<And>c. P(\<lambda>x. \<chi> i. c i * x$i)"
and swap: "\<And>m n::'n. m \<noteq> n \<Longrightarrow> P(\<lambda>x. \<chi> i. x $ Fun.swap m n id i)"
and idplus: "\<And>m n::'n. m \<noteq> n \<Longrightarrow> P(\<lambda>x. \<chi> i. if i = m then x$m + x$n else x$i)"
shows "P f"
proof -
have "P ((*v) A)" for A
proof (rule induct_matrix_elementary_alt)
fix A B
assume "P ((*v) A)" and "P ((*v) B)"
then show "P ((*v) (A ** B))"
by (auto simp add: matrix_vector_mult_matrix_matrix_mult_compose intro!: comp)
next
fix A :: "real^'n^'n" and i
assume "row i A = 0"
show "P ((*v) A)"
using matrix_vector_mul_linear
by (rule zeroes[where i=i])
(metis \<open>row i A = 0\<close> inner_zero_left matrix_vector_mul_component row_def vec_lambda_eta)
next
fix A :: "real^'n^'n"
assume 0: "\<And>i j. i \<noteq> j \<Longrightarrow> A $ i $ j = 0"
have "A $ i $ i * x $ i = (\<Sum>j\<in>UNIV. A $ i $ j * x $ j)" for x and i :: "'n"
by (simp add: 0 comm_monoid_add_class.sum.remove [where x=i])
then have "(\<lambda>x. \<chi> i. A $ i $ i * x $ i) = ((*v) A)"
by (auto simp: 0 matrix_vector_mult_def)
then show "P ((*v) A)"
using const [of "\<lambda>i. A $ i $ i"] by simp
next
fix m n :: "'n"
assume "m \<noteq> n"
have eq: "(\<Sum>j\<in>UNIV. if i = Fun.swap m n id j then x $ j else 0) =
(\<Sum>j\<in>UNIV. if j = Fun.swap m n id i then x $ j else 0)"
for i and x :: "real^'n"
unfolding swap_def by (rule sum.cong) auto
have "(\<lambda>x::real^'n. \<chi> i. x $ Fun.swap m n id i) = ((*v) (\<chi> i j. if i = Fun.swap m n id j then 1 else 0))"
by (auto simp: mat_def matrix_vector_mult_def eq if_distrib [of "\<lambda>x. x * y" for y] cong: if_cong)
with swap [OF \<open>m \<noteq> n\<close>] show "P ((*v) (\<chi> i j. mat 1 $ i $ Fun.swap m n id j))"
by (simp add: mat_def matrix_vector_mult_def)
next
fix m n :: "'n"
assume "m \<noteq> n"
then have "x $ m + x $ n = (\<Sum>j\<in>UNIV. of_bool (j = n \<or> m = j) * x $ j)" for x :: "real^'n"
by (auto simp: of_bool_def if_distrib [of "\<lambda>x. x * y" for y] sum.remove cong: if_cong)
then have "(\<lambda>x::real^'n. \<chi> i. if i = m then x $ m + x $ n else x $ i) =
((*v) (\<chi> i j. of_bool (i = m \<and> j = n \<or> i = j)))"
unfolding matrix_vector_mult_def of_bool_def
by (auto simp: vec_eq_iff if_distrib [of "\<lambda>x. x * y" for y] cong: if_cong)
then show "P ((*v) (\<chi> i j. of_bool (i = m \<and> j = n \<or> i = j)))"
using idplus [OF \<open>m \<noteq> n\<close>] by simp
qed
then show ?thesis
- by (metis \<open>linear f\<close> matrix_vector_mul)
+ by (metis \<open>linear f\<close> matrix_vector_mul(2))
qed
end
\ No newline at end of file
diff --git a/src/HOL/Analysis/Change_Of_Vars.thy b/src/HOL/Analysis/Change_Of_Vars.thy
--- a/src/HOL/Analysis/Change_Of_Vars.thy
+++ b/src/HOL/Analysis/Change_Of_Vars.thy
@@ -1,3478 +1,3478 @@
(* Title: HOL/Analysis/Change_Of_Vars.thy
Authors: LC Paulson, based on material from HOL Light
*)
section\<open>Change of Variables Theorems\<close>
theory Change_Of_Vars
imports Vitali_Covering_Theorem Determinants
begin
subsection \<open>Measurable Shear and Stretch\<close>
proposition
fixes a :: "real^'n"
assumes "m \<noteq> n" and ab_ne: "cbox a b \<noteq> {}" and an: "0 \<le> a$n"
shows measurable_shear_interval: "(\<lambda>x. \<chi> i. if i = m then x$m + x$n else x$i) ` (cbox a b) \<in> lmeasurable"
(is "?f ` _ \<in> _")
and measure_shear_interval: "measure lebesgue ((\<lambda>x. \<chi> i. if i = m then x$m + x$n else x$i) ` cbox a b)
= measure lebesgue (cbox a b)" (is "?Q")
proof -
have lin: "linear ?f"
by (rule linearI) (auto simp: plus_vec_def scaleR_vec_def algebra_simps)
show fab: "?f ` cbox a b \<in> lmeasurable"
by (simp add: lin measurable_linear_image_interval)
let ?c = "\<chi> i. if i = m then b$m + b$n else b$i"
let ?mn = "axis m 1 - axis n (1::real)"
have eq1: "measure lebesgue (cbox a ?c)
= measure lebesgue (?f ` cbox a b)
+ measure lebesgue (cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a$m})
+ measure lebesgue (cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m})"
proof (rule measure_Un3_negligible)
show "cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a$m} \<in> lmeasurable" "cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m} \<in> lmeasurable"
by (auto simp: convex_Int convex_halfspace_le convex_halfspace_ge bounded_Int measurable_convex)
have "negligible {x. ?mn \<bullet> x = a$m}"
by (metis \<open>m \<noteq> n\<close> axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane)
moreover have "?f ` cbox a b \<inter> (cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m}) \<subseteq> {x. ?mn \<bullet> x = a$m}"
using \<open>m \<noteq> n\<close> antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis')
ultimately show "negligible ((?f ` cbox a b) \<inter> (cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m}))"
by (rule negligible_subset)
have "negligible {x. ?mn \<bullet> x = b$m}"
by (metis \<open>m \<noteq> n\<close> axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane)
moreover have "(?f ` cbox a b) \<inter> (cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m}) \<subseteq> {x. ?mn \<bullet> x = b$m}"
using \<open>m \<noteq> n\<close> antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis')
ultimately show "negligible (?f ` cbox a b \<inter> (cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m}))"
by (rule negligible_subset)
have "negligible {x. ?mn \<bullet> x = b$m}"
by (metis \<open>m \<noteq> n\<close> axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane)
moreover have "(cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m} \<inter> (cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m})) \<subseteq> {x. ?mn \<bullet> x = b$m}"
using \<open>m \<noteq> n\<close> ab_ne
apply (auto simp: algebra_simps mem_box_cart inner_axis')
apply (drule_tac x=m in spec)+
apply simp
done
ultimately show "negligible (cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m} \<inter> (cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m}))"
by (rule negligible_subset)
show "?f ` cbox a b \<union> cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m} \<union> cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m} = cbox a ?c" (is "?lhs = _")
proof
show "?lhs \<subseteq> cbox a ?c"
by (auto simp: mem_box_cart add_mono) (meson add_increasing2 an order_trans)
show "cbox a ?c \<subseteq> ?lhs"
apply (auto simp: algebra_simps image_iff inner_axis' lambda_add_Galois [OF \<open>m \<noteq> n\<close>])
apply (auto simp: mem_box_cart split: if_split_asm)
done
qed
qed (fact fab)
let ?d = "\<chi> i. if i = m then a $ m - b $ m else 0"
have eq2: "measure lebesgue (cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m}) + measure lebesgue (cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m})
= measure lebesgue (cbox a (\<chi> i. if i = m then a $ m + b $ n else b $ i))"
proof (rule measure_translate_add[of "cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a$m}" "cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m}"
"(\<chi> i. if i = m then a$m - b$m else 0)" "cbox a (\<chi> i. if i = m then a$m + b$n else b$i)"])
show "(cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a$m}) \<in> lmeasurable"
"cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m} \<in> lmeasurable"
by (auto simp: convex_Int convex_halfspace_le convex_halfspace_ge bounded_Int measurable_convex)
have "\<And>x. \<lbrakk>x $ n + a $ m \<le> x $ m\<rbrakk>
\<Longrightarrow> x \<in> (+) (\<chi> i. if i = m then a $ m - b $ m else 0) ` {x. x $ n + b $ m \<le> x $ m}"
using \<open>m \<noteq> n\<close>
by (rule_tac x="x - (\<chi> i. if i = m then a$m - b$m else 0)" in image_eqI)
(simp_all add: mem_box_cart)
then have imeq: "(+) ?d ` {x. b $ m \<le> ?mn \<bullet> x} = {x. a $ m \<le> ?mn \<bullet> x}"
using \<open>m \<noteq> n\<close> by (auto simp: mem_box_cart inner_axis' algebra_simps)
have "\<And>x. \<lbrakk>0 \<le> a $ n; x $ n + a $ m \<le> x $ m;
\<forall>i. i \<noteq> m \<longrightarrow> a $ i \<le> x $ i \<and> x $ i \<le> b $ i\<rbrakk>
\<Longrightarrow> a $ m \<le> x $ m"
using \<open>m \<noteq> n\<close> by force
then have "(+) ?d ` (cbox a ?c \<inter> {x. b $ m \<le> ?mn \<bullet> x})
= cbox a (\<chi> i. if i = m then a $ m + b $ n else b $ i) \<inter> {x. a $ m \<le> ?mn \<bullet> x}"
using an ab_ne
apply (simp add: cbox_translation [symmetric] translation_Int interval_ne_empty_cart imeq)
apply (auto simp: mem_box_cart inner_axis' algebra_simps if_distrib all_if_distrib)
by (metis (full_types) add_mono mult_2_right)
then show "cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m} \<union>
(+) ?d ` (cbox a ?c \<inter> {x. b $ m \<le> ?mn \<bullet> x}) =
cbox a (\<chi> i. if i = m then a $ m + b $ n else b $ i)" (is "?lhs = ?rhs")
using an \<open>m \<noteq> n\<close>
apply (auto simp: mem_box_cart inner_axis' algebra_simps if_distrib all_if_distrib, force)
apply (drule_tac x=n in spec)+
by (meson ab_ne add_mono_thms_linordered_semiring(3) dual_order.trans interval_ne_empty_cart(1))
have "negligible{x. ?mn \<bullet> x = a$m}"
by (metis \<open>m \<noteq> n\<close> axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane)
moreover have "(cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m} \<inter>
(+) ?d ` (cbox a ?c \<inter> {x. b $ m \<le> ?mn \<bullet> x})) \<subseteq> {x. ?mn \<bullet> x = a$m}"
using \<open>m \<noteq> n\<close> antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis')
ultimately show "negligible (cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m} \<inter>
(+) ?d ` (cbox a ?c \<inter> {x. b $ m \<le> ?mn \<bullet> x}))"
by (rule negligible_subset)
qed
have ac_ne: "cbox a ?c \<noteq> {}"
using ab_ne an
by (clarsimp simp: interval_eq_empty_cart) (meson add_less_same_cancel1 le_less_linear less_le_trans)
have ax_ne: "cbox a (\<chi> i. if i = m then a $ m + b $ n else b $ i) \<noteq> {}"
using ab_ne an
by (clarsimp simp: interval_eq_empty_cart) (meson add_less_same_cancel1 le_less_linear less_le_trans)
have eq3: "measure lebesgue (cbox a ?c) = measure lebesgue (cbox a (\<chi> i. if i = m then a$m + b$n else b$i)) + measure lebesgue (cbox a b)"
by (simp add: content_cbox_if_cart ab_ne ac_ne ax_ne algebra_simps prod.delta_remove
if_distrib [of "\<lambda>u. u - z" for z] prod.remove)
show ?Q
using eq1 eq2 eq3
by (simp add: algebra_simps)
qed
proposition
fixes S :: "(real^'n) set"
assumes "S \<in> lmeasurable"
shows measurable_stretch: "((\<lambda>x. \<chi> k. m k * x$k) ` S) \<in> lmeasurable" (is "?f ` S \<in> _")
and measure_stretch: "measure lebesgue ((\<lambda>x. \<chi> k. m k * x$k) ` S) = \<bar>prod m UNIV\<bar> * measure lebesgue S"
(is "?MEQ")
proof -
have "(?f ` S) \<in> lmeasurable \<and> ?MEQ"
proof (cases "\<forall>k. m k \<noteq> 0")
case True
have m0: "0 < \<bar>prod m UNIV\<bar>"
using True by simp
have "(indicat_real (?f ` S) has_integral \<bar>prod m UNIV\<bar> * measure lebesgue S) UNIV"
proof (clarsimp simp add: has_integral_alt [where i=UNIV])
fix e :: "real"
assume "e > 0"
have "(indicat_real S has_integral (measure lebesgue S)) UNIV"
using assms lmeasurable_iff_has_integral by blast
then obtain B where "B>0"
and B: "\<And>a b. ball 0 B \<subseteq> cbox a b \<Longrightarrow>
\<exists>z. (indicat_real S has_integral z) (cbox a b) \<and>
\<bar>z - measure lebesgue S\<bar> < e / \<bar>prod m UNIV\<bar>"
by (simp add: has_integral_alt [where i=UNIV]) (metis (full_types) divide_pos_pos m0 m0 \<open>e > 0\<close>)
show "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
(\<exists>z. (indicat_real (?f ` S) has_integral z) (cbox a b) \<and>
\<bar>z - \<bar>prod m UNIV\<bar> * measure lebesgue S\<bar> < e)"
proof (intro exI conjI allI)
let ?C = "Max (range (\<lambda>k. \<bar>m k\<bar>)) * B"
show "?C > 0"
using True \<open>B > 0\<close> by (simp add: Max_gr_iff)
show "ball 0 ?C \<subseteq> cbox u v \<longrightarrow>
(\<exists>z. (indicat_real (?f ` S) has_integral z) (cbox u v) \<and>
\<bar>z - \<bar>prod m UNIV\<bar> * measure lebesgue S\<bar> < e)" for u v
proof
assume uv: "ball 0 ?C \<subseteq> cbox u v"
with \<open>?C > 0\<close> have cbox_ne: "cbox u v \<noteq> {}"
using centre_in_ball by blast
let ?\<alpha> = "\<lambda>k. u$k / m k"
let ?\<beta> = "\<lambda>k. v$k / m k"
have invm0: "\<And>k. inverse (m k) \<noteq> 0"
using True by auto
have "ball 0 B \<subseteq> (\<lambda>x. \<chi> k. x $ k / m k) ` ball 0 ?C"
proof clarsimp
fix x :: "real^'n"
assume x: "norm x < B"
have [simp]: "\<bar>Max (range (\<lambda>k. \<bar>m k\<bar>))\<bar> = Max (range (\<lambda>k. \<bar>m k\<bar>))"
by (meson Max_ge abs_ge_zero abs_of_nonneg finite finite_imageI order_trans rangeI)
have "norm (\<chi> k. m k * x $ k) \<le> norm (Max (range (\<lambda>k. \<bar>m k\<bar>)) *\<^sub>R x)"
by (rule norm_le_componentwise_cart) (auto simp: abs_mult intro: mult_right_mono)
also have "\<dots> < ?C"
using x by simp (metis \<open>B > 0\<close> \<open>?C > 0\<close> mult.commute real_mult_less_iff1 zero_less_mult_pos)
finally have "norm (\<chi> k. m k * x $ k) < ?C" .
then show "x \<in> (\<lambda>x. \<chi> k. x $ k / m k) ` ball 0 ?C"
using stretch_Galois [of "inverse \<circ> m"] True by (auto simp: image_iff field_simps)
qed
then have Bsub: "ball 0 B \<subseteq> cbox (\<chi> k. min (?\<alpha> k) (?\<beta> k)) (\<chi> k. max (?\<alpha> k) (?\<beta> k))"
using cbox_ne uv image_stretch_interval_cart [of "inverse \<circ> m" u v, symmetric]
by (force simp: field_simps)
obtain z where zint: "(indicat_real S has_integral z) (cbox (\<chi> k. min (?\<alpha> k) (?\<beta> k)) (\<chi> k. max (?\<alpha> k) (?\<beta> k)))"
and zless: "\<bar>z - measure lebesgue S\<bar> < e / \<bar>prod m UNIV\<bar>"
using B [OF Bsub] by blast
have ind: "indicat_real (?f ` S) = (\<lambda>x. indicator S (\<chi> k. x$k / m k))"
using True stretch_Galois [of m] by (force simp: indicator_def)
show "\<exists>z. (indicat_real (?f ` S) has_integral z) (cbox u v) \<and>
\<bar>z - \<bar>prod m UNIV\<bar> * measure lebesgue S\<bar> < e"
proof (simp add: ind, intro conjI exI)
have "((\<lambda>x. indicat_real S (\<chi> k. x $ k/ m k)) has_integral z *\<^sub>R \<bar>prod m UNIV\<bar>)
((\<lambda>x. \<chi> k. x $ k * m k) ` cbox (\<chi> k. min (?\<alpha> k) (?\<beta> k)) (\<chi> k. max (?\<alpha> k) (?\<beta> k)))"
using True has_integral_stretch_cart [OF zint, of "inverse \<circ> m"]
by (simp add: field_simps prod_dividef)
moreover have "((\<lambda>x. \<chi> k. x $ k * m k) ` cbox (\<chi> k. min (?\<alpha> k) (?\<beta> k)) (\<chi> k. max (?\<alpha> k) (?\<beta> k))) = cbox u v"
using True image_stretch_interval_cart [of "inverse \<circ> m" u v, symmetric]
image_stretch_interval_cart [of "\<lambda>k. 1" u v, symmetric] \<open>cbox u v \<noteq> {}\<close>
by (simp add: field_simps image_comp o_def)
ultimately show "((\<lambda>x. indicat_real S (\<chi> k. x $ k/ m k)) has_integral z *\<^sub>R \<bar>prod m UNIV\<bar>) (cbox u v)"
by simp
have "\<bar>z *\<^sub>R \<bar>prod m UNIV\<bar> - \<bar>prod m UNIV\<bar> * measure lebesgue S\<bar>
= \<bar>prod m UNIV\<bar> * \<bar>z - measure lebesgue S\<bar>"
by (metis (no_types, hide_lams) abs_abs abs_scaleR mult.commute real_scaleR_def right_diff_distrib')
also have "\<dots> < e"
using zless True by (simp add: field_simps)
finally show "\<bar>z *\<^sub>R \<bar>prod m UNIV\<bar> - \<bar>prod m UNIV\<bar> * measure lebesgue S\<bar> < e" .
qed
qed
qed
qed
then show ?thesis
by (auto simp: has_integral_integrable integral_unique lmeasure_integral_UNIV measurable_integrable)
next
case False
then obtain k where "m k = 0" and prm: "prod m UNIV = 0"
by auto
have nfS: "negligible (?f ` S)"
by (rule negligible_subset [OF negligible_standard_hyperplane_cart]) (use \<open>m k = 0\<close> in auto)
then have "(?f ` S) \<in> lmeasurable"
by (simp add: negligible_iff_measure)
with nfS show ?thesis
by (simp add: prm negligible_iff_measure0)
qed
then show "(?f ` S) \<in> lmeasurable" ?MEQ
by metis+
qed
proposition
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
assumes "linear f" "S \<in> lmeasurable"
shows measurable_linear_image: "(f ` S) \<in> lmeasurable"
and measure_linear_image: "measure lebesgue (f ` S) = \<bar>det (matrix f)\<bar> * measure lebesgue S" (is "?Q f S")
proof -
have "\<forall>S \<in> lmeasurable. (f ` S) \<in> lmeasurable \<and> ?Q f S"
proof (rule induct_linear_elementary [OF \<open>linear f\<close>]; intro ballI)
fix f g and S :: "(real,'n) vec set"
assume "linear f" and "linear g"
and f [rule_format]: "\<forall>S \<in> lmeasurable. f ` S \<in> lmeasurable \<and> ?Q f S"
and g [rule_format]: "\<forall>S \<in> lmeasurable. g ` S \<in> lmeasurable \<and> ?Q g S"
and S: "S \<in> lmeasurable"
then have gS: "g ` S \<in> lmeasurable"
by blast
show "(f \<circ> g) ` S \<in> lmeasurable \<and> ?Q (f \<circ> g) S"
using f [OF gS] g [OF S] matrix_compose [OF \<open>linear g\<close> \<open>linear f\<close>]
by (simp add: o_def image_comp abs_mult det_mul)
next
fix f :: "real^'n::_ \<Rightarrow> real^'n::_" and i and S :: "(real^'n::_) set"
assume "linear f" and 0: "\<And>x. f x $ i = 0" and "S \<in> lmeasurable"
then have "\<not> inj f"
by (metis (full_types) linear_injective_imp_surjective one_neq_zero surjE vec_component)
have detf: "det (matrix f) = 0"
using \<open>\<not> inj f\<close> det_nz_iff_inj[OF \<open>linear f\<close>] by blast
show "f ` S \<in> lmeasurable \<and> ?Q f S"
proof
show "f ` S \<in> lmeasurable"
using lmeasurable_iff_indicator_has_integral \<open>linear f\<close> \<open>\<not> inj f\<close> negligible_UNIV negligible_linear_singular_image by blast
have "measure lebesgue (f ` S) = 0"
by (meson \<open>\<not> inj f\<close> \<open>linear f\<close> negligible_imp_measure0 negligible_linear_singular_image)
also have "\<dots> = \<bar>det (matrix f)\<bar> * measure lebesgue S"
by (simp add: detf)
finally show "?Q f S" .
qed
next
fix c and S :: "(real^'n::_) set"
assume "S \<in> lmeasurable"
show "(\<lambda>a. \<chi> i. c i * a $ i) ` S \<in> lmeasurable \<and> ?Q (\<lambda>a. \<chi> i. c i * a $ i) S"
proof
show "(\<lambda>a. \<chi> i. c i * a $ i) ` S \<in> lmeasurable"
by (simp add: \<open>S \<in> lmeasurable\<close> measurable_stretch)
show "?Q (\<lambda>a. \<chi> i. c i * a $ i) S"
by (simp add: measure_stretch [OF \<open>S \<in> lmeasurable\<close>, of c] axis_def matrix_def det_diagonal)
qed
next
fix m :: "'n" and n :: "'n" and S :: "(real, 'n) vec set"
assume "m \<noteq> n" and "S \<in> lmeasurable"
let ?h = "\<lambda>v::(real, 'n) vec. \<chi> i. v $ Fun.swap m n id i"
have lin: "linear ?h"
by (rule linearI) (simp_all add: plus_vec_def scaleR_vec_def)
have meq: "measure lebesgue ((\<lambda>v::(real, 'n) vec. \<chi> i. v $ Fun.swap m n id i) ` cbox a b)
= measure lebesgue (cbox a b)" for a b
proof (cases "cbox a b = {}")
case True then show ?thesis
by simp
next
case False
then have him: "?h ` (cbox a b) \<noteq> {}"
by blast
have eq: "?h ` (cbox a b) = cbox (?h a) (?h b)"
by (auto simp: image_iff lambda_swap_Galois mem_box_cart) (metis swap_id_eq)+
show ?thesis
using him prod.permute [OF permutes_swap_id, where S=UNIV and g="\<lambda>i. (b - a)$i", symmetric]
by (simp add: eq content_cbox_cart False)
qed
have "(\<chi> i j. if Fun.swap m n id i = j then 1 else 0) = (\<chi> i j. if j = Fun.swap m n id i then 1 else (0::real))"
by (auto intro!: Cart_lambda_cong)
then have "matrix ?h = transpose(\<chi> i j. mat 1 $ i $ Fun.swap m n id j)"
by (auto simp: matrix_eq transpose_def axis_def mat_def matrix_def)
then have 1: "\<bar>det (matrix ?h)\<bar> = 1"
by (simp add: det_permute_columns permutes_swap_id sign_swap_id abs_mult)
show "?h ` S \<in> lmeasurable \<and> ?Q ?h S"
proof
show "?h ` S \<in> lmeasurable" "?Q ?h S"
using measure_linear_sufficient [OF lin \<open>S \<in> lmeasurable\<close>] meq 1 by force+
qed
next
fix m n :: "'n" and S :: "(real, 'n) vec set"
assume "m \<noteq> n" and "S \<in> lmeasurable"
let ?h = "\<lambda>v::(real, 'n) vec. \<chi> i. if i = m then v $ m + v $ n else v $ i"
have lin: "linear ?h"
by (rule linearI) (auto simp: algebra_simps plus_vec_def scaleR_vec_def vec_eq_iff)
consider "m < n" | " n < m"
using \<open>m \<noteq> n\<close> less_linear by blast
then have 1: "det(matrix ?h) = 1"
proof cases
assume "m < n"
have *: "matrix ?h $ i $ j = (0::real)" if "j < i" for i j :: 'n
proof -
have "axis j 1 = (\<chi> n. if n = j then 1 else (0::real))"
using axis_def by blast
then have "(\<chi> p q. if p = m then axis q 1 $ m + axis q 1 $ n else axis q 1 $ p) $ i $ j = (0::real)"
using \<open>j < i\<close> axis_def \<open>m < n\<close> by auto
with \<open>m < n\<close> show ?thesis
by (auto simp: matrix_def axis_def cong: if_cong)
qed
show ?thesis
using \<open>m \<noteq> n\<close> by (subst det_upperdiagonal [OF *]) (auto simp: matrix_def axis_def cong: if_cong)
next
assume "n < m"
have *: "matrix ?h $ i $ j = (0::real)" if "j > i" for i j :: 'n
proof -
have "axis j 1 = (\<chi> n. if n = j then 1 else (0::real))"
using axis_def by blast
then have "(\<chi> p q. if p = m then axis q 1 $ m + axis q 1 $ n else axis q 1 $ p) $ i $ j = (0::real)"
using \<open>j > i\<close> axis_def \<open>m > n\<close> by auto
with \<open>m > n\<close> show ?thesis
by (auto simp: matrix_def axis_def cong: if_cong)
qed
show ?thesis
using \<open>m \<noteq> n\<close>
by (subst det_lowerdiagonal [OF *]) (auto simp: matrix_def axis_def cong: if_cong)
qed
have meq: "measure lebesgue (?h ` (cbox a b)) = measure lebesgue (cbox a b)" for a b
proof (cases "cbox a b = {}")
case True then show ?thesis by simp
next
case False
then have ne: "(+) (\<chi> i. if i = n then - a $ n else 0) ` cbox a b \<noteq> {}"
by auto
let ?v = "\<chi> i. if i = n then - a $ n else 0"
have "?h ` cbox a b
= (+) (\<chi> i. if i = m \<or> i = n then a $ n else 0) ` ?h ` (+) ?v ` (cbox a b)"
using \<open>m \<noteq> n\<close> unfolding image_comp o_def by (force simp: vec_eq_iff)
then have "measure lebesgue (?h ` (cbox a b))
= measure lebesgue ((\<lambda>v. \<chi> i. if i = m then v $ m + v $ n else v $ i) `
(+) ?v ` cbox a b)"
by (rule ssubst) (rule measure_translation)
also have "\<dots> = measure lebesgue ((\<lambda>v. \<chi> i. if i = m then v $ m + v $ n else v $ i) ` cbox (?v +a) (?v + b))"
by (metis (no_types, lifting) cbox_translation)
also have "\<dots> = measure lebesgue ((+) (\<chi> i. if i = n then - a $ n else 0) ` cbox a b)"
apply (subst measure_shear_interval)
using \<open>m \<noteq> n\<close> ne apply auto
apply (simp add: cbox_translation)
by (metis cbox_borel cbox_translation measure_completion sets_lborel)
also have "\<dots> = measure lebesgue (cbox a b)"
by (rule measure_translation)
finally show ?thesis .
qed
show "?h ` S \<in> lmeasurable \<and> ?Q ?h S"
using measure_linear_sufficient [OF lin \<open>S \<in> lmeasurable\<close>] meq 1 by force
qed
with assms show "(f ` S) \<in> lmeasurable" "?Q f S"
by metis+
qed
lemma
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
assumes f: "orthogonal_transformation f" and S: "S \<in> lmeasurable"
shows measurable_orthogonal_image: "f ` S \<in> lmeasurable"
and measure_orthogonal_image: "measure lebesgue (f ` S) = measure lebesgue S"
proof -
have "linear f"
by (simp add: f orthogonal_transformation_linear)
then show "f ` S \<in> lmeasurable"
by (metis S measurable_linear_image)
show "measure lebesgue (f ` S) = measure lebesgue S"
by (simp add: measure_linear_image \<open>linear f\<close> S f)
qed
proposition measure_semicontinuous_with_hausdist_explicit:
assumes "bounded S" and neg: "negligible(frontier S)" and "e > 0"
obtains d where "d > 0"
"\<And>T. \<lbrakk>T \<in> lmeasurable; \<And>y. y \<in> T \<Longrightarrow> \<exists>x. x \<in> S \<and> dist x y < d\<rbrakk>
\<Longrightarrow> measure lebesgue T < measure lebesgue S + e"
proof (cases "S = {}")
case True
with that \<open>e > 0\<close> show ?thesis by force
next
case False
then have frS: "frontier S \<noteq> {}"
using \<open>bounded S\<close> frontier_eq_empty not_bounded_UNIV by blast
have "S \<in> lmeasurable"
by (simp add: \<open>bounded S\<close> measurable_Jordan neg)
have null: "(frontier S) \<in> null_sets lebesgue"
by (metis neg negligible_iff_null_sets)
have "frontier S \<in> lmeasurable" and mS0: "measure lebesgue (frontier S) = 0"
using neg negligible_imp_measurable negligible_iff_measure by blast+
with \<open>e > 0\<close> sets_lebesgue_outer_open
obtain U where "open U"
and U: "frontier S \<subseteq> U" "U - frontier S \<in> lmeasurable" "emeasure lebesgue (U - frontier S) < e"
by (metis fmeasurableD)
with null have "U \<in> lmeasurable"
by (metis borel_open measurable_Diff_null_set sets_completionI_sets sets_lborel)
have "measure lebesgue (U - frontier S) = measure lebesgue U"
using mS0 by (simp add: \<open>U \<in> lmeasurable\<close> fmeasurableD measure_Diff_null_set null)
with U have mU: "measure lebesgue U < e"
by (simp add: emeasure_eq_measure2 ennreal_less_iff)
show ?thesis
proof
have "U \<noteq> UNIV"
using \<open>U \<in> lmeasurable\<close> by auto
then have "- U \<noteq> {}"
by blast
with \<open>open U\<close> \<open>frontier S \<subseteq> U\<close> show "setdist (frontier S) (- U) > 0"
by (auto simp: \<open>bounded S\<close> open_closed compact_frontier_bounded setdist_gt_0_compact_closed frS)
fix T
assume "T \<in> lmeasurable"
and T: "\<And>t. t \<in> T \<Longrightarrow> \<exists>y. y \<in> S \<and> dist y t < setdist (frontier S) (- U)"
then have "measure lebesgue T - measure lebesgue S \<le> measure lebesgue (T - S)"
by (simp add: \<open>S \<in> lmeasurable\<close> measure_diff_le_measure_setdiff)
also have "\<dots> \<le> measure lebesgue U"
proof -
have "T - S \<subseteq> U"
proof clarify
fix x
assume "x \<in> T" and "x \<notin> S"
then obtain y where "y \<in> S" and y: "dist y x < setdist (frontier S) (- U)"
using T by blast
have "closed_segment x y \<inter> frontier S \<noteq> {}"
using connected_Int_frontier \<open>x \<notin> S\<close> \<open>y \<in> S\<close> by blast
then obtain z where z: "z \<in> closed_segment x y" "z \<in> frontier S"
by auto
with y have "dist z x < setdist(frontier S) (- U)"
by (auto simp: dist_commute dest!: dist_in_closed_segment)
with z have False if "x \<in> -U"
using setdist_le_dist [OF \<open>z \<in> frontier S\<close> that] by auto
then show "x \<in> U"
by blast
qed
then show ?thesis
by (simp add: \<open>S \<in> lmeasurable\<close> \<open>T \<in> lmeasurable\<close> \<open>U \<in> lmeasurable\<close> fmeasurableD measure_mono_fmeasurable sets.Diff)
qed
finally have "measure lebesgue T - measure lebesgue S \<le> measure lebesgue U" .
with mU show "measure lebesgue T < measure lebesgue S + e"
by linarith
qed
qed
proposition
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
assumes S: "S \<in> lmeasurable"
and deriv: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
and int: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on S"
and bounded: "\<And>x. x \<in> S \<Longrightarrow> \<bar>det (matrix (f' x))\<bar> \<le> B"
shows measurable_bounded_differentiable_image:
"f ` S \<in> lmeasurable"
and measure_bounded_differentiable_image:
"measure lebesgue (f ` S) \<le> B * measure lebesgue S" (is "?M")
proof -
have "f ` S \<in> lmeasurable \<and> measure lebesgue (f ` S) \<le> B * measure lebesgue S"
proof (cases "B < 0")
case True
then have "S = {}"
by (meson abs_ge_zero bounded empty_iff equalityI less_le_trans linorder_not_less subsetI)
then show ?thesis
by auto
next
case False
then have "B \<ge> 0"
by arith
let ?\<mu> = "measure lebesgue"
have f_diff: "f differentiable_on S"
using deriv by (auto simp: differentiable_on_def differentiable_def)
have eps: "f ` S \<in> lmeasurable" "?\<mu> (f ` S) \<le> (B+e) * ?\<mu> S" (is "?ME")
if "e > 0" for e
proof -
have eps_d: "f ` S \<in> lmeasurable" "?\<mu> (f ` S) \<le> (B+e) * (?\<mu> S + d)" (is "?MD")
if "d > 0" for d
proof -
obtain T where T: "open T" "S \<subseteq> T" and TS: "(T-S) \<in> lmeasurable" and "emeasure lebesgue (T-S) < ennreal d"
using S \<open>d > 0\<close> sets_lebesgue_outer_open by blast
then have "?\<mu> (T-S) < d"
by (metis emeasure_eq_measure2 ennreal_leI not_less)
with S T TS have "T \<in> lmeasurable" and Tless: "?\<mu> T < ?\<mu> S + d"
by (auto simp: measurable_measure_Diff dest!: fmeasurable_Diff_D)
have "\<exists>r. 0 < r \<and> r < d \<and> ball x r \<subseteq> T \<and> f ` (S \<inter> ball x r) \<in> lmeasurable \<and>
?\<mu> (f ` (S \<inter> ball x r)) \<le> (B + e) * ?\<mu> (ball x r)"
if "x \<in> S" "d > 0" for x d
proof -
have lin: "linear (f' x)"
and lim0: "((\<lambda>y. (f y - (f x + f' x (y - x))) /\<^sub>R norm(y - x)) \<longlongrightarrow> 0) (at x within S)"
using deriv \<open>x \<in> S\<close> by (auto simp: has_derivative_within bounded_linear.linear field_simps)
have bo: "bounded (f' x ` ball 0 1)"
by (simp add: bounded_linear_image linear_linear lin)
have neg: "negligible (frontier (f' x ` ball 0 1))"
using deriv has_derivative_linear \<open>x \<in> S\<close>
by (auto intro!: negligible_convex_frontier [OF convex_linear_image])
let ?unit_vol = "content (ball (0 :: real ^ 'n :: {finite, wellorder}) 1)"
have 0: "0 < e * ?unit_vol"
using \<open>e > 0\<close> by (simp add: content_ball_pos)
obtain k where "k > 0" and k:
"\<And>U. \<lbrakk>U \<in> lmeasurable; \<And>y. y \<in> U \<Longrightarrow> \<exists>z. z \<in> f' x ` ball 0 1 \<and> dist z y < k\<rbrakk>
\<Longrightarrow> ?\<mu> U < ?\<mu> (f' x ` ball 0 1) + e * ?unit_vol"
using measure_semicontinuous_with_hausdist_explicit [OF bo neg 0] by blast
obtain l where "l > 0" and l: "ball x l \<subseteq> T"
using \<open>x \<in> S\<close> \<open>open T\<close> \<open>S \<subseteq> T\<close> openE by blast
obtain \<zeta> where "0 < \<zeta>"
and \<zeta>: "\<And>y. \<lbrakk>y \<in> S; y \<noteq> x; dist y x < \<zeta>\<rbrakk>
\<Longrightarrow> norm (f y - (f x + f' x (y - x))) / norm (y - x) < k"
using lim0 \<open>k > 0\<close> by (simp add: Lim_within) (auto simp add: field_simps)
define r where "r \<equiv> min (min l (\<zeta>/2)) (min 1 (d/2))"
show ?thesis
proof (intro exI conjI)
show "r > 0" "r < d"
using \<open>l > 0\<close> \<open>\<zeta> > 0\<close> \<open>d > 0\<close> by (auto simp: r_def)
have "r \<le> l"
by (auto simp: r_def)
with l show "ball x r \<subseteq> T"
by auto
have ex_lessK: "\<exists>x' \<in> ball 0 1. dist (f' x x') ((f y - f x) /\<^sub>R r) < k"
if "y \<in> S" and "dist x y < r" for y
proof (cases "y = x")
case True
with lin linear_0 \<open>k > 0\<close> that show ?thesis
by (rule_tac x=0 in bexI) (auto simp: linear_0)
next
case False
then show ?thesis
proof (rule_tac x="(y - x) /\<^sub>R r" in bexI)
have "f' x ((y - x) /\<^sub>R r) = f' x (y - x) /\<^sub>R r"
by (simp add: lin linear_scale)
then have "dist (f' x ((y - x) /\<^sub>R r)) ((f y - f x) /\<^sub>R r) = norm (f' x (y - x) /\<^sub>R r - (f y - f x) /\<^sub>R r)"
by (simp add: dist_norm)
also have "\<dots> = norm (f' x (y - x) - (f y - f x)) / r"
using \<open>r > 0\<close> by (simp add: divide_simps scale_right_diff_distrib [symmetric])
also have "\<dots> \<le> norm (f y - (f x + f' x (y - x))) / norm (y - x)"
- using that \<open>r > 0\<close> False by (simp add: algebra_simps field_split_simps dist_norm norm_minus_commute mult_right_mono)
+ using that \<open>r > 0\<close> False by (simp add: field_split_simps dist_norm norm_minus_commute mult_right_mono)
also have "\<dots> < k"
using that \<open>0 < \<zeta>\<close> by (simp add: dist_commute r_def \<zeta> [OF \<open>y \<in> S\<close> False])
finally show "dist (f' x ((y - x) /\<^sub>R r)) ((f y - f x) /\<^sub>R r) < k" .
show "(y - x) /\<^sub>R r \<in> ball 0 1"
using that \<open>r > 0\<close> by (simp add: dist_norm divide_simps norm_minus_commute)
qed
qed
let ?rfs = "(\<lambda>x. x /\<^sub>R r) ` (+) (- f x) ` f ` (S \<inter> ball x r)"
have rfs_mble: "?rfs \<in> lmeasurable"
proof (rule bounded_set_imp_lmeasurable)
have "f differentiable_on S \<inter> ball x r"
using f_diff by (auto simp: fmeasurableD differentiable_on_subset)
with S show "?rfs \<in> sets lebesgue"
by (auto simp: sets.Int intro!: lebesgue_sets_translation differentiable_image_in_sets_lebesgue)
let ?B = "(\<lambda>(x, y). x + y) ` (f' x ` ball 0 1 \<times> ball 0 k)"
have "bounded ?B"
by (simp add: bounded_plus [OF bo])
moreover have "?rfs \<subseteq> ?B"
apply (auto simp: dist_norm image_iff dest!: ex_lessK)
by (metis (no_types, hide_lams) add.commute diff_add_cancel dist_0_norm dist_commute dist_norm mem_ball)
ultimately show "bounded (?rfs)"
by (rule bounded_subset)
qed
then have "(\<lambda>x. r *\<^sub>R x) ` ?rfs \<in> lmeasurable"
by (simp add: measurable_linear_image)
with \<open>r > 0\<close> have "(+) (- f x) ` f ` (S \<inter> ball x r) \<in> lmeasurable"
by (simp add: image_comp o_def)
then have "(+) (f x) ` (+) (- f x) ` f ` (S \<inter> ball x r) \<in> lmeasurable"
using measurable_translation by blast
then show fsb: "f ` (S \<inter> ball x r) \<in> lmeasurable"
by (simp add: image_comp o_def)
have "?\<mu> (f ` (S \<inter> ball x r)) = ?\<mu> (?rfs) * r ^ CARD('n)"
using \<open>r > 0\<close> fsb
by (simp add: measure_linear_image measure_translation_subtract measurable_translation_subtract field_simps cong: image_cong_simp)
also have "\<dots> \<le> (\<bar>det (matrix (f' x))\<bar> * ?unit_vol + e * ?unit_vol) * r ^ CARD('n)"
proof -
have "?\<mu> (?rfs) < ?\<mu> (f' x ` ball 0 1) + e * ?unit_vol"
using rfs_mble by (force intro: k dest!: ex_lessK)
then have "?\<mu> (?rfs) < \<bar>det (matrix (f' x))\<bar> * ?unit_vol + e * ?unit_vol"
by (simp add: lin measure_linear_image [of "f' x"])
with \<open>r > 0\<close> show ?thesis
by auto
qed
also have "\<dots> \<le> (B + e) * ?\<mu> (ball x r)"
using bounded [OF \<open>x \<in> S\<close>] \<open>r > 0\<close>
by (simp add: algebra_simps content_ball_conv_unit_ball[of r] content_ball_pos)
finally show "?\<mu> (f ` (S \<inter> ball x r)) \<le> (B + e) * ?\<mu> (ball x r)" .
qed
qed
then obtain r where
r0d: "\<And>x d. \<lbrakk>x \<in> S; d > 0\<rbrakk> \<Longrightarrow> 0 < r x d \<and> r x d < d"
and rT: "\<And>x d. \<lbrakk>x \<in> S; d > 0\<rbrakk> \<Longrightarrow> ball x (r x d) \<subseteq> T"
and r: "\<And>x d. \<lbrakk>x \<in> S; d > 0\<rbrakk> \<Longrightarrow>
(f ` (S \<inter> ball x (r x d))) \<in> lmeasurable \<and>
?\<mu> (f ` (S \<inter> ball x (r x d))) \<le> (B + e) * ?\<mu> (ball x (r x d))"
by metis
obtain C where "countable C" and Csub: "C \<subseteq> {(x,r x t) |x t. x \<in> S \<and> 0 < t}"
and pwC: "pairwise (\<lambda>i j. disjnt (ball (fst i) (snd i)) (ball (fst j) (snd j))) C"
and negC: "negligible(S - (\<Union>i \<in> C. ball (fst i) (snd i)))"
apply (rule Vitali_covering_theorem_balls [of S "{(x,r x t) |x t. x \<in> S \<and> 0 < t}" fst snd])
apply auto
by (metis dist_eq_0_iff r0d)
let ?UB = "(\<Union>(x,s) \<in> C. ball x s)"
have eq: "f ` (S \<inter> ?UB) = (\<Union>(x,s) \<in> C. f ` (S \<inter> ball x s))"
by auto
have mle: "?\<mu> (\<Union>(x,s) \<in> K. f ` (S \<inter> ball x s)) \<le> (B + e) * (?\<mu> S + d)" (is "?l \<le> ?r")
if "K \<subseteq> C" and "finite K" for K
proof -
have gt0: "b > 0" if "(a, b) \<in> K" for a b
using Csub that \<open>K \<subseteq> C\<close> r0d by auto
have inj: "inj_on (\<lambda>(x, y). ball x y) K"
by (force simp: inj_on_def ball_eq_ball_iff dest: gt0)
have disjnt: "disjoint ((\<lambda>(x, y). ball x y) ` K)"
using pwC that
apply (clarsimp simp: pairwise_def case_prod_unfold ball_eq_ball_iff)
by (metis subsetD fst_conv snd_conv)
have "?l \<le> (\<Sum>i\<in>K. ?\<mu> (case i of (x, s) \<Rightarrow> f ` (S \<inter> ball x s)))"
proof (rule measure_UNION_le [OF \<open>finite K\<close>], clarify)
fix x r
assume "(x,r) \<in> K"
then have "x \<in> S"
using Csub \<open>K \<subseteq> C\<close> by auto
show "f ` (S \<inter> ball x r) \<in> sets lebesgue"
by (meson Int_lower1 S differentiable_on_subset f_diff fmeasurableD lmeasurable_ball order_refl sets.Int differentiable_image_in_sets_lebesgue)
qed
also have "\<dots> \<le> (\<Sum>(x,s) \<in> K. (B + e) * ?\<mu> (ball x s))"
apply (rule sum_mono)
using Csub r \<open>K \<subseteq> C\<close> by auto
also have "\<dots> = (B + e) * (\<Sum>(x,s) \<in> K. ?\<mu> (ball x s))"
by (simp add: prod.case_distrib sum_distrib_left)
also have "\<dots> = (B + e) * sum ?\<mu> ((\<lambda>(x, y). ball x y) ` K)"
using \<open>B \<ge> 0\<close> \<open>e > 0\<close> by (simp add: inj sum.reindex prod.case_distrib)
also have "\<dots> = (B + e) * ?\<mu> (\<Union>(x,s) \<in> K. ball x s)"
using \<open>B \<ge> 0\<close> \<open>e > 0\<close> that
by (subst measure_Union') (auto simp: disjnt measure_Union')
also have "\<dots> \<le> (B + e) * ?\<mu> T"
using \<open>B \<ge> 0\<close> \<open>e > 0\<close> that apply simp
apply (rule measure_mono_fmeasurable [OF _ _ \<open>T \<in> lmeasurable\<close>])
using Csub rT by force+
also have "\<dots> \<le> (B + e) * (?\<mu> S + d)"
using \<open>B \<ge> 0\<close> \<open>e > 0\<close> Tless by simp
finally show ?thesis .
qed
have fSUB_mble: "(f ` (S \<inter> ?UB)) \<in> lmeasurable"
unfolding eq using Csub r False \<open>e > 0\<close> that
by (auto simp: intro!: fmeasurable_UN_bound [OF \<open>countable C\<close> _ mle])
have fSUB_meas: "?\<mu> (f ` (S \<inter> ?UB)) \<le> (B + e) * (?\<mu> S + d)" (is "?MUB")
unfolding eq using Csub r False \<open>e > 0\<close> that
by (auto simp: intro!: measure_UN_bound [OF \<open>countable C\<close> _ mle])
have neg: "negligible ((f ` (S \<inter> ?UB) - f ` S) \<union> (f ` S - f ` (S \<inter> ?UB)))"
proof (rule negligible_subset [OF negligible_differentiable_image_negligible [OF order_refl negC, where f=f]])
show "f differentiable_on S - (\<Union>i\<in>C. ball (fst i) (snd i))"
by (meson DiffE differentiable_on_subset subsetI f_diff)
qed force
show "f ` S \<in> lmeasurable"
by (rule lmeasurable_negligible_symdiff [OF fSUB_mble neg])
show ?MD
using fSUB_meas measure_negligible_symdiff [OF fSUB_mble neg] by simp
qed
show "f ` S \<in> lmeasurable"
using eps_d [of 1] by simp
show ?ME
proof (rule field_le_epsilon)
fix \<delta> :: real
assume "0 < \<delta>"
then show "?\<mu> (f ` S) \<le> (B + e) * ?\<mu> S + \<delta>"
using eps_d [of "\<delta> / (B+e)"] \<open>e > 0\<close> \<open>B \<ge> 0\<close> by (auto simp: divide_simps mult_ac)
qed
qed
show ?thesis
proof (cases "?\<mu> S = 0")
case True
with eps have "?\<mu> (f ` S) = 0"
by (metis mult_zero_right not_le zero_less_measure_iff)
then show ?thesis
using eps [of 1] by (simp add: True)
next
case False
have "?\<mu> (f ` S) \<le> B * ?\<mu> S"
proof (rule field_le_epsilon)
fix e :: real
assume "e > 0"
then show "?\<mu> (f ` S) \<le> B * ?\<mu> S + e"
using eps [of "e / ?\<mu> S"] False by (auto simp: algebra_simps zero_less_measure_iff)
qed
with eps [of 1] show ?thesis by auto
qed
qed
then show "f ` S \<in> lmeasurable" ?M by blast+
qed
lemma m_diff_image_weak:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
assumes S: "S \<in> lmeasurable"
and deriv: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
and int: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on S"
shows "f ` S \<in> lmeasurable \<and> measure lebesgue (f ` S) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
proof -
let ?\<mu> = "measure lebesgue"
have aint_S: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) absolutely_integrable_on S"
using int unfolding absolutely_integrable_on_def by auto
define m where "m \<equiv> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
have *: "f ` S \<in> lmeasurable" "?\<mu> (f ` S) \<le> m + e * ?\<mu> S"
if "e > 0" for e
proof -
define T where "T \<equiv> \<lambda>n. {x \<in> S. n * e \<le> \<bar>det (matrix (f' x))\<bar> \<and>
\<bar>det (matrix (f' x))\<bar> < (Suc n) * e}"
have meas_t: "T n \<in> lmeasurable" for n
proof -
have *: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) \<in> borel_measurable (lebesgue_on S)"
using aint_S by (simp add: S borel_measurable_restrict_space_iff fmeasurableD set_integrable_def)
have [intro]: "x \<in> sets (lebesgue_on S) \<Longrightarrow> x \<in> sets lebesgue" for x
using S sets_restrict_space_subset by blast
have "{x \<in> S. real n * e \<le> \<bar>det (matrix (f' x))\<bar>} \<in> sets lebesgue"
using * by (auto simp: borel_measurable_iff_halfspace_ge space_restrict_space)
then have 1: "{x \<in> S. real n * e \<le> \<bar>det (matrix (f' x))\<bar>} \<in> lmeasurable"
using S by (simp add: fmeasurableI2)
have "{x \<in> S. \<bar>det (matrix (f' x))\<bar> < (1 + real n) * e} \<in> sets lebesgue"
using * by (auto simp: borel_measurable_iff_halfspace_less space_restrict_space)
then have 2: "{x \<in> S. \<bar>det (matrix (f' x))\<bar> < (1 + real n) * e} \<in> lmeasurable"
using S by (simp add: fmeasurableI2)
show ?thesis
using fmeasurable.Int [OF 1 2] by (simp add: T_def Int_def cong: conj_cong)
qed
have aint_T: "\<And>k. (\<lambda>x. \<bar>det (matrix (f' x))\<bar>) absolutely_integrable_on T k"
using set_integrable_subset [OF aint_S] meas_t T_def by blast
have Seq: "S = (\<Union>n. T n)"
apply (auto simp: T_def)
apply (rule_tac x="nat(floor(abs(det(matrix(f' x))) / e))" in exI)
using that apply auto
using of_int_floor_le pos_le_divide_eq apply blast
by (metis add.commute pos_divide_less_eq real_of_int_floor_add_one_gt)
have meas_ft: "f ` T n \<in> lmeasurable" for n
proof (rule measurable_bounded_differentiable_image)
show "T n \<in> lmeasurable"
by (simp add: meas_t)
next
fix x :: "(real,'n) vec"
assume "x \<in> T n"
show "(f has_derivative f' x) (at x within T n)"
by (metis (no_types, lifting) \<open>x \<in> T n\<close> deriv has_derivative_within_subset mem_Collect_eq subsetI T_def)
show "\<bar>det (matrix (f' x))\<bar> \<le> (Suc n) * e"
using \<open>x \<in> T n\<close> T_def by auto
next
show "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on T n"
using aint_T absolutely_integrable_on_def by blast
qed
have disT: "disjoint (range T)"
unfolding disjoint_def
proof clarsimp
show "T m \<inter> T n = {}" if "T m \<noteq> T n" for m n
using that
proof (induction m n rule: linorder_less_wlog)
case (less m n)
with \<open>e > 0\<close> show ?case
unfolding T_def
proof (clarsimp simp add: Collect_conj_eq [symmetric])
fix x
assume "e > 0" "m < n" "n * e \<le> \<bar>det (matrix (f' x))\<bar>" "\<bar>det (matrix (f' x))\<bar> < (1 + real m) * e"
then have "n < 1 + real m"
by (metis (no_types, hide_lams) less_le_trans mult.commute not_le real_mult_le_cancel_iff2)
then show "False"
using less.hyps by linarith
qed
qed auto
qed
have injT: "inj_on T ({n. T n \<noteq> {}})"
unfolding inj_on_def
proof clarsimp
show "m = n" if "T m = T n" "T n \<noteq> {}" for m n
using that
proof (induction m n rule: linorder_less_wlog)
case (less m n)
have False if "T n \<subseteq> T m" "x \<in> T n" for x
using \<open>e > 0\<close> \<open>m < n\<close> that
apply (auto simp: T_def mult.commute intro: less_le_trans dest!: subsetD)
by (metis add.commute less_le_trans nat_less_real_le not_le real_mult_le_cancel_iff2)
then show ?case
using less.prems by blast
qed auto
qed
have sum_eq_Tim: "(\<Sum>k\<le>n. f (T k)) = sum f (T ` {..n})" if "f {} = 0" for f :: "_ \<Rightarrow> real" and n
proof (subst sum.reindex_nontrivial)
fix i j assume "i \<in> {..n}" "j \<in> {..n}" "i \<noteq> j" "T i = T j"
with that injT [unfolded inj_on_def] show "f (T i) = 0"
by simp metis
qed (use atMost_atLeast0 in auto)
let ?B = "m + e * ?\<mu> S"
have "(\<Sum>k\<le>n. ?\<mu> (f ` T k)) \<le> ?B" for n
proof -
have "(\<Sum>k\<le>n. ?\<mu> (f ` T k)) \<le> (\<Sum>k\<le>n. ((k+1) * e) * ?\<mu>(T k))"
proof (rule sum_mono [OF measure_bounded_differentiable_image])
show "(f has_derivative f' x) (at x within T k)" if "x \<in> T k" for k x
using that unfolding T_def by (blast intro: deriv has_derivative_within_subset)
show "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on T k" for k
using absolutely_integrable_on_def aint_T by blast
show "\<bar>det (matrix (f' x))\<bar> \<le> real (k + 1) * e" if "x \<in> T k" for k x
using T_def that by auto
qed (use meas_t in auto)
also have "\<dots> \<le> (\<Sum>k\<le>n. (k * e) * ?\<mu>(T k)) + (\<Sum>k\<le>n. e * ?\<mu>(T k))"
by (simp add: algebra_simps sum.distrib)
also have "\<dots> \<le> ?B"
proof (rule add_mono)
have "(\<Sum>k\<le>n. real k * e * ?\<mu> (T k)) = (\<Sum>k\<le>n. integral (T k) (\<lambda>x. k * e))"
by (simp add: lmeasure_integral [OF meas_t]
flip: integral_mult_right integral_mult_left)
also have "\<dots> \<le> (\<Sum>k\<le>n. integral (T k) (\<lambda>x. (abs (det (matrix (f' x))))))"
proof (rule sum_mono)
fix k
assume "k \<in> {..n}"
show "integral (T k) (\<lambda>x. k * e) \<le> integral (T k) (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
proof (rule integral_le [OF integrable_on_const [OF meas_t]])
show "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on T k"
using absolutely_integrable_on_def aint_T by blast
next
fix x assume "x \<in> T k"
show "k * e \<le> \<bar>det (matrix (f' x))\<bar>"
using \<open>x \<in> T k\<close> T_def by blast
qed
qed
also have "\<dots> = sum (\<lambda>T. integral T (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)) (T ` {..n})"
by (auto intro: sum_eq_Tim)
also have "\<dots> = integral (\<Union>k\<le>n. T k) (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
proof (rule integral_unique [OF has_integral_Union, symmetric])
fix S assume "S \<in> T ` {..n}"
then show "((\<lambda>x. \<bar>det (matrix (f' x))\<bar>) has_integral integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)) S"
using absolutely_integrable_on_def aint_T by blast
next
show "pairwise (\<lambda>S S'. negligible (S \<inter> S')) (T ` {..n})"
using disT unfolding disjnt_iff by (auto simp: pairwise_def intro!: empty_imp_negligible)
qed auto
also have "\<dots> \<le> m"
unfolding m_def
proof (rule integral_subset_le)
have "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) absolutely_integrable_on (\<Union>k\<le>n. T k)"
apply (rule set_integrable_subset [OF aint_S])
apply (intro measurable meas_t fmeasurableD)
apply (force simp: Seq)
done
then show "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on (\<Union>k\<le>n. T k)"
using absolutely_integrable_on_def by blast
qed (use Seq int in auto)
finally show "(\<Sum>k\<le>n. real k * e * ?\<mu> (T k)) \<le> m" .
next
have "(\<Sum>k\<le>n. ?\<mu> (T k)) = sum ?\<mu> (T ` {..n})"
by (auto intro: sum_eq_Tim)
also have "\<dots> = ?\<mu> (\<Union>k\<le>n. T k)"
using S disT by (auto simp: pairwise_def meas_t intro: measure_Union' [symmetric])
also have "\<dots> \<le> ?\<mu> S"
using S by (auto simp: Seq intro: meas_t fmeasurableD measure_mono_fmeasurable)
finally have "(\<Sum>k\<le>n. ?\<mu> (T k)) \<le> ?\<mu> S" .
then show "(\<Sum>k\<le>n. e * ?\<mu> (T k)) \<le> e * ?\<mu> S"
by (metis less_eq_real_def ordered_comm_semiring_class.comm_mult_left_mono sum_distrib_left that)
qed
finally show "(\<Sum>k\<le>n. ?\<mu> (f ` T k)) \<le> ?B" .
qed
moreover have "measure lebesgue (\<Union>k\<le>n. f ` T k) \<le> (\<Sum>k\<le>n. ?\<mu> (f ` T k))" for n
by (simp add: fmeasurableD meas_ft measure_UNION_le)
ultimately have B_ge_m: "?\<mu> (\<Union>k\<le>n. (f ` T k)) \<le> ?B" for n
by (meson order_trans)
have "(\<Union>n. f ` T n) \<in> lmeasurable"
by (rule fmeasurable_countable_Union [OF meas_ft B_ge_m])
moreover have "?\<mu> (\<Union>n. f ` T n) \<le> m + e * ?\<mu> S"
by (rule measure_countable_Union_le [OF meas_ft B_ge_m])
ultimately show "f ` S \<in> lmeasurable" "?\<mu> (f ` S) \<le> m + e * ?\<mu> S"
by (auto simp: Seq image_Union)
qed
show ?thesis
proof
show "f ` S \<in> lmeasurable"
using * linordered_field_no_ub by blast
let ?x = "m - ?\<mu> (f ` S)"
have False if "?\<mu> (f ` S) > integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
proof -
have ml: "m < ?\<mu> (f ` S)"
using m_def that by blast
then have "?\<mu> S \<noteq> 0"
using "*"(2) bgauge_existence_lemma by fastforce
with ml have 0: "0 < - (m - ?\<mu> (f ` S))/2 / ?\<mu> S"
using that zero_less_measure_iff by force
then show ?thesis
using * [OF 0] that by (auto simp: field_split_simps m_def split: if_split_asm)
qed
then show "?\<mu> (f ` S) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
by fastforce
qed
qed
theorem
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
assumes S: "S \<in> sets lebesgue"
and deriv: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
and int: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on S"
shows measurable_differentiable_image: "f ` S \<in> lmeasurable"
and measure_differentiable_image:
"measure lebesgue (f ` S) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)" (is "?M")
proof -
let ?I = "\<lambda>n::nat. cbox (vec (-n)) (vec n) \<inter> S"
let ?\<mu> = "measure lebesgue"
have "x \<in> cbox (vec (- real (nat \<lceil>norm x\<rceil>))) (vec (real (nat \<lceil>norm x\<rceil>)))" for x :: "real^'n::_"
apply (auto simp: mem_box_cart)
apply (metis abs_le_iff component_le_norm_cart minus_le_iff of_nat_ceiling order.trans)
by (meson abs_le_D1 norm_bound_component_le_cart real_nat_ceiling_ge)
then have Seq: "S = (\<Union>n. ?I n)"
by auto
have fIn: "f ` ?I n \<in> lmeasurable"
and mfIn: "?\<mu> (f ` ?I n) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)" (is ?MN) for n
proof -
have In: "?I n \<in> lmeasurable"
by (simp add: S bounded_Int bounded_set_imp_lmeasurable sets.Int)
moreover have "\<And>x. x \<in> ?I n \<Longrightarrow> (f has_derivative f' x) (at x within ?I n)"
by (meson Int_iff deriv has_derivative_within_subset subsetI)
moreover have int_In: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on ?I n"
proof -
have "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) absolutely_integrable_on S"
using int absolutely_integrable_integrable_bound by force
then have "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) absolutely_integrable_on ?I n"
by (metis (no_types) Int_lower1 In fmeasurableD inf_commute set_integrable_subset)
then show ?thesis
using absolutely_integrable_on_def by blast
qed
ultimately have "f ` ?I n \<in> lmeasurable" "?\<mu> (f ` ?I n) \<le> integral (?I n) (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
using m_diff_image_weak by metis+
moreover have "integral (?I n) (\<lambda>x. \<bar>det (matrix (f' x))\<bar>) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
by (simp add: int_In int integral_subset_le)
ultimately show "f ` ?I n \<in> lmeasurable" ?MN
by auto
qed
have "?I k \<subseteq> ?I n" if "k \<le> n" for k n
by (rule Int_mono) (use that in \<open>auto simp: subset_interval_imp_cart\<close>)
then have "(\<Union>k\<le>n. f ` ?I k) = f ` ?I n" for n
by (fastforce simp add:)
with mfIn have "?\<mu> (\<Union>k\<le>n. f ` ?I k) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)" for n
by simp
then have "(\<Union>n. f ` ?I n) \<in> lmeasurable" "?\<mu> (\<Union>n. f ` ?I n) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
by (rule fmeasurable_countable_Union [OF fIn] measure_countable_Union_le [OF fIn])+
then show "f ` S \<in> lmeasurable" ?M
by (metis Seq image_UN)+
qed
lemma borel_measurable_simple_function_limit_increasing:
fixes f :: "'a::euclidean_space \<Rightarrow> real"
shows "(f \<in> borel_measurable lebesgue \<and> (\<forall>x. 0 \<le> f x)) \<longleftrightarrow>
(\<exists>g. (\<forall>n x. 0 \<le> g n x \<and> g n x \<le> f x) \<and> (\<forall>n x. g n x \<le> (g(Suc n) x)) \<and>
(\<forall>n. g n \<in> borel_measurable lebesgue) \<and> (\<forall>n. finite(range (g n))) \<and>
(\<forall>x. (\<lambda>n. g n x) \<longlonglongrightarrow> f x))"
(is "?lhs = ?rhs")
proof
assume f: ?lhs
have leb_f: "{x. a \<le> f x \<and> f x < b} \<in> sets lebesgue" for a b
proof -
have "{x. a \<le> f x \<and> f x < b} = {x. f x < b} - {x. f x < a}"
by auto
also have "\<dots> \<in> sets lebesgue"
using borel_measurable_vimage_halfspace_component_lt [of f UNIV] f by auto
finally show ?thesis .
qed
have "g n x \<le> f x"
if inc_g: "\<And>n x. 0 \<le> g n x \<and> g n x \<le> g (Suc n) x"
and meas_g: "\<And>n. g n \<in> borel_measurable lebesgue"
and fin: "\<And>n. finite(range (g n))" and lim: "\<And>x. (\<lambda>n. g n x) \<longlonglongrightarrow> f x" for g n x
proof -
have "\<exists>r>0. \<forall>N. \<exists>n\<ge>N. dist (g n x) (f x) \<ge> r" if "g n x > f x"
proof -
have g: "g n x \<le> g (N + n) x" for N
by (rule transitive_stepwise_le) (use inc_g in auto)
have "\<exists>na\<ge>N. g n x - f x \<le> dist (g na x) (f x)" for N
apply (rule_tac x="N+n" in exI)
using g [of N] by (auto simp: dist_norm)
with that show ?thesis
using diff_gt_0_iff_gt by blast
qed
with lim show ?thesis
apply (auto simp: lim_sequentially)
by (meson less_le_not_le not_le_imp_less)
qed
moreover
let ?\<Omega> = "\<lambda>n k. indicator {y. k/2^n \<le> f y \<and> f y < (k+1)/2^n}"
let ?g = "\<lambda>n x. (\<Sum>k::real | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k/2^n * ?\<Omega> n k x)"
have "\<exists>g. (\<forall>n x. 0 \<le> g n x \<and> g n x \<le> (g(Suc n) x)) \<and>
(\<forall>n. g n \<in> borel_measurable lebesgue) \<and> (\<forall>n. finite(range (g n))) \<and>(\<forall>x. (\<lambda>n. g n x) \<longlonglongrightarrow> f x)"
proof (intro exI allI conjI)
show "0 \<le> ?g n x" for n x
proof (clarify intro!: ordered_comm_monoid_add_class.sum_nonneg)
fix k::real
assume "k \<in> \<int>" and k: "\<bar>k\<bar> \<le> 2 ^ (2*n)"
show "0 \<le> k/2^n * ?\<Omega> n k x"
using f \<open>k \<in> \<int>\<close> apply (auto simp: indicator_def field_split_simps Ints_def)
apply (drule spec [where x=x])
using zero_le_power [of "2::real" n] mult_nonneg_nonneg [of "f x" "2^n"]
by linarith
qed
show "?g n x \<le> ?g (Suc n) x" for n x
proof -
have "?g n x =
(\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n).
k/2^n * (indicator {y. k/2^n \<le> f y \<and> f y < (k+1/2)/2^n} x +
indicator {y. (k+1/2)/2^n \<le> f y \<and> f y < (k+1)/2^n} x))"
by (rule sum.cong [OF refl]) (simp add: indicator_def field_split_simps)
also have "\<dots> = (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k/2^n * indicator {y. k/2^n \<le> f y \<and> f y < (k+1/2)/2^n} x) +
(\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k/2^n * indicator {y. (k+1/2)/2^n \<le> f y \<and> f y < (k+1)/2^n} x)"
by (simp add: comm_monoid_add_class.sum.distrib algebra_simps)
also have "\<dots> = (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). (2 * k)/2 ^ Suc n * indicator {y. (2 * k)/2 ^ Suc n \<le> f y \<and> f y < (2 * k+1)/2 ^ Suc n} x) +
(\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). (2 * k)/2 ^ Suc n * indicator {y. (2 * k+1)/2 ^ Suc n \<le> f y \<and> f y < ((2 * k+1) + 1)/2 ^ Suc n} x)"
by (force simp: field_simps indicator_def intro: sum.cong)
also have "\<dots> \<le> (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2 * Suc n). k/2 ^ Suc n * (indicator {y. k/2 ^ Suc n \<le> f y \<and> f y < (k+1)/2 ^ Suc n} x))"
(is "?a + _ \<le> ?b")
proof -
have *: "\<lbrakk>sum f I \<le> sum h I; a + sum h I \<le> b\<rbrakk> \<Longrightarrow> a + sum f I \<le> b" for I a b f and h :: "real\<Rightarrow>real"
by linarith
let ?h = "\<lambda>k. (2*k+1)/2 ^ Suc n *
(indicator {y. (2 * k+1)/2 ^ Suc n \<le> f y \<and> f y < ((2*k+1) + 1)/2 ^ Suc n} x)"
show ?thesis
proof (rule *)
show "(\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n).
2 * k/2 ^ Suc n * indicator {y. (2 * k+1)/2 ^ Suc n \<le> f y \<and> f y < (2 * k+1 + 1)/2 ^ Suc n} x)
\<le> sum ?h {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}"
by (rule sum_mono) (simp add: indicator_def field_split_simps)
next
have \<alpha>: "?a = (\<Sum>k \<in> (*)2 ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}.
k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \<le> f y \<and> f y < (k+1)/2 ^ Suc n} x)"
by (auto simp: inj_on_def field_simps comm_monoid_add_class.sum.reindex)
have \<beta>: "sum ?h {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}
= (\<Sum>k \<in> (\<lambda>x. 2*x + 1) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}.
k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \<le> f y \<and> f y < (k+1)/2 ^ Suc n} x)"
by (auto simp: inj_on_def field_simps comm_monoid_add_class.sum.reindex)
have 0: "(*) 2 ` {k \<in> \<int>. P k} \<inter> (\<lambda>x. 2 * x + 1) ` {k \<in> \<int>. P k} = {}" for P :: "real \<Rightarrow> bool"
proof -
have "2 * i \<noteq> 2 * j + 1" for i j :: int by arith
thus ?thesis
unfolding Ints_def by auto (use of_int_eq_iff in fastforce)
qed
have "?a + sum ?h {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}
= (\<Sum>k \<in> (*)2 ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)} \<union> (\<lambda>x. 2*x + 1) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}.
k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \<le> f y \<and> f y < (k+1)/2 ^ Suc n} x)"
unfolding \<alpha> \<beta>
using finite_abs_int_segment [of "2 ^ (2*n)"]
by (subst sum_Un) (auto simp: 0)
also have "\<dots> \<le> ?b"
proof (rule sum_mono2)
show "finite {k::real. k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2 * Suc n)}"
by (rule finite_abs_int_segment)
show "(*) 2 ` {k::real. k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2^(2*n)} \<union> (\<lambda>x. 2*x + 1) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2^(2*n)} \<subseteq> {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2 * Suc n)}"
apply auto
using one_le_power [of "2::real" "2*n"] by linarith
have *: "\<lbrakk>x \<in> (S \<union> T) - U; \<And>x. x \<in> S \<Longrightarrow> x \<in> U; \<And>x. x \<in> T \<Longrightarrow> x \<in> U\<rbrakk> \<Longrightarrow> P x" for S T U P
by blast
have "0 \<le> b" if "b \<in> \<int>" "f x * (2 * 2^n) < b + 1" for b
proof -
have "0 \<le> f x * (2 * 2^n)"
by (simp add: f)
also have "\<dots> < b+1"
by (simp add: that)
finally show "0 \<le> b"
using \<open>b \<in> \<int>\<close> by (auto simp: elim!: Ints_cases)
qed
then show "0 \<le> b/2 ^ Suc n * indicator {y. b/2 ^ Suc n \<le> f y \<and> f y < (b + 1)/2 ^ Suc n} x"
if "b \<in> {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2 * Suc n)} -
((*) 2 ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)} \<union> (\<lambda>x. 2*x + 1) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)})" for b
using that by (simp add: indicator_def divide_simps)
qed
finally show "?a + sum ?h {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)} \<le> ?b" .
qed
qed
finally show ?thesis .
qed
show "?g n \<in> borel_measurable lebesgue" for n
apply (intro borel_measurable_indicator borel_measurable_times borel_measurable_sum)
using leb_f sets_restrict_UNIV by auto
show "finite (range (?g n))" for n
proof -
have "(\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k/2^n * ?\<Omega> n k x)
\<in> (\<lambda>k. k/2^n) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}" for x
proof (cases "\<exists>k. k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n) \<and> k/2^n \<le> f x \<and> f x < (k+1)/2^n")
case True
then show ?thesis
by (blast intro: indicator_sum_eq)
next
case False
then have "(\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k/2^n * ?\<Omega> n k x) = 0"
by auto
then show ?thesis by force
qed
then have "range (?g n) \<subseteq> ((\<lambda>k. (k/2^n)) ` {k. k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n)})"
by auto
moreover have "finite ((\<lambda>k::real. (k/2^n)) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)})"
by (intro finite_imageI finite_abs_int_segment)
ultimately show ?thesis
by (rule finite_subset)
qed
show "(\<lambda>n. ?g n x) \<longlonglongrightarrow> f x" for x
proof (clarsimp simp add: lim_sequentially)
fix e::real
assume "e > 0"
obtain N1 where N1: "2 ^ N1 > abs(f x)"
using real_arch_pow by fastforce
obtain N2 where N2: "(1/2) ^ N2 < e"
using real_arch_pow_inv \<open>e > 0\<close> by fastforce
have "dist (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k/2^n * ?\<Omega> n k x) (f x) < e" if "N1 + N2 \<le> n" for n
proof -
let ?m = "real_of_int \<lfloor>2^n * f x\<rfloor>"
have "\<bar>?m\<bar> \<le> 2^n * 2^N1"
using N1 apply (simp add: f)
by (meson floor_mono le_floor_iff less_le_not_le mult_le_cancel_left_pos zero_less_numeral zero_less_power)
also have "\<dots> \<le> 2 ^ (2*n)"
by (metis that add_leD1 add_le_cancel_left mult.commute mult_2_right one_less_numeral_iff
power_add power_increasing_iff semiring_norm(76))
finally have m_le: "\<bar>?m\<bar> \<le> 2 ^ (2*n)" .
have "?m/2^n \<le> f x" "f x < (?m + 1)/2^n"
by (auto simp: mult.commute pos_divide_le_eq mult_imp_less_div_pos)
then have eq: "dist (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k/2^n * ?\<Omega> n k x) (f x)
= dist (?m/2^n) (f x)"
by (subst indicator_sum_eq [of ?m]) (auto simp: m_le)
have "\<bar>2^n\<bar> * \<bar>?m/2^n - f x\<bar> = \<bar>2^n * (?m/2^n - f x)\<bar>"
by (simp add: abs_mult)
also have "\<dots> < 2 ^ N2 * e"
using N2 by (simp add: divide_simps mult.commute) linarith
also have "\<dots> \<le> \<bar>2^n\<bar> * e"
using that \<open>e > 0\<close> by auto
finally have "dist (?m/2^n) (f x) < e"
by (simp add: dist_norm)
then show ?thesis
using eq by linarith
qed
then show "\<exists>no. \<forall>n\<ge>no. dist (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k * ?\<Omega> n k x/2^n) (f x) < e"
by force
qed
qed
ultimately show ?rhs
by metis
next
assume RHS: ?rhs
with borel_measurable_simple_function_limit [of f UNIV, unfolded lebesgue_on_UNIV_eq]
show ?lhs
by (blast intro: order_trans)
qed
subsection\<open>Borel measurable Jacobian determinant\<close>
lemma lemma_partial_derivatives0:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "linear f" and lim0: "((\<lambda>x. f x /\<^sub>R norm x) \<longlongrightarrow> 0) (at 0 within S)"
and lb: "\<And>v. v \<noteq> 0 \<Longrightarrow> (\<exists>k>0. \<forall>e>0. \<exists>x. x \<in> S - {0} \<and> norm x < e \<and> k * norm x \<le> \<bar>v \<bullet> x\<bar>)"
shows "f x = 0"
proof -
interpret linear f by fact
have "dim {x. f x = 0} \<le> DIM('a)"
by (rule dim_subset_UNIV)
moreover have False if less: "dim {x. f x = 0} < DIM('a)"
proof -
obtain d where "d \<noteq> 0" and d: "\<And>y. f y = 0 \<Longrightarrow> d \<bullet> y = 0"
using orthogonal_to_subspace_exists [OF less] orthogonal_def
by (metis (mono_tags, lifting) mem_Collect_eq span_base)
then obtain k where "k > 0"
and k: "\<And>e. e > 0 \<Longrightarrow> \<exists>y. y \<in> S - {0} \<and> norm y < e \<and> k * norm y \<le> \<bar>d \<bullet> y\<bar>"
using lb by blast
have "\<exists>h. \<forall>n. ((h n \<in> S \<and> h n \<noteq> 0 \<and> k * norm (h n) \<le> \<bar>d \<bullet> h n\<bar>) \<and> norm (h n) < 1 / real (Suc n)) \<and>
norm (h (Suc n)) < norm (h n)"
proof (rule dependent_nat_choice)
show "\<exists>y. (y \<in> S \<and> y \<noteq> 0 \<and> k * norm y \<le> \<bar>d \<bullet> y\<bar>) \<and> norm y < 1 / real (Suc 0)"
by simp (metis DiffE insertCI k not_less not_one_le_zero)
qed (use k [of "min (norm x) (1/(Suc n + 1))" for x n] in auto)
then obtain \<alpha> where \<alpha>: "\<And>n. \<alpha> n \<in> S - {0}" and kd: "\<And>n. k * norm(\<alpha> n) \<le> \<bar>d \<bullet> \<alpha> n\<bar>"
and norm_lt: "\<And>n. norm(\<alpha> n) < 1/(Suc n)"
by force
let ?\<beta> = "\<lambda>n. \<alpha> n /\<^sub>R norm (\<alpha> n)"
have com: "\<And>g. (\<forall>n. g n \<in> sphere (0::'a) 1)
\<Longrightarrow> \<exists>l \<in> sphere 0 1. \<exists>\<rho>::nat\<Rightarrow>nat. strict_mono \<rho> \<and> (g \<circ> \<rho>) \<longlonglongrightarrow> l"
using compact_sphere compact_def by metis
moreover have "\<forall>n. ?\<beta> n \<in> sphere 0 1"
using \<alpha> by auto
ultimately obtain l::'a and \<rho>::"nat\<Rightarrow>nat"
where l: "l \<in> sphere 0 1" and "strict_mono \<rho>" and to_l: "(?\<beta> \<circ> \<rho>) \<longlonglongrightarrow> l"
by meson
moreover have "continuous (at l) (\<lambda>x. (\<bar>d \<bullet> x\<bar> - k))"
by (intro continuous_intros)
ultimately have lim_dl: "((\<lambda>x. (\<bar>d \<bullet> x\<bar> - k)) \<circ> (?\<beta> \<circ> \<rho>)) \<longlonglongrightarrow> (\<bar>d \<bullet> l\<bar> - k)"
by (meson continuous_imp_tendsto)
have "\<forall>\<^sub>F i in sequentially. 0 \<le> ((\<lambda>x. \<bar>d \<bullet> x\<bar> - k) \<circ> ((\<lambda>n. \<alpha> n /\<^sub>R norm (\<alpha> n)) \<circ> \<rho>)) i"
using \<alpha> kd by (auto simp: field_split_simps)
then have "k \<le> \<bar>d \<bullet> l\<bar>"
using tendsto_lowerbound [OF lim_dl, of 0] by auto
moreover have "d \<bullet> l = 0"
proof (rule d)
show "f l = 0"
proof (rule LIMSEQ_unique [of "f \<circ> ?\<beta> \<circ> \<rho>"])
have "isCont f l"
using \<open>linear f\<close> linear_continuous_at linear_conv_bounded_linear by blast
then show "(f \<circ> (\<lambda>n. \<alpha> n /\<^sub>R norm (\<alpha> n)) \<circ> \<rho>) \<longlonglongrightarrow> f l"
unfolding comp_assoc
using to_l continuous_imp_tendsto by blast
have "\<alpha> \<longlonglongrightarrow> 0"
using norm_lt LIMSEQ_norm_0 by metis
with \<open>strict_mono \<rho>\<close> have "(\<alpha> \<circ> \<rho>) \<longlonglongrightarrow> 0"
by (metis LIMSEQ_subseq_LIMSEQ)
with lim0 \<alpha> have "((\<lambda>x. f x /\<^sub>R norm x) \<circ> (\<alpha> \<circ> \<rho>)) \<longlonglongrightarrow> 0"
by (force simp: tendsto_at_iff_sequentially)
then show "(f \<circ> (\<lambda>n. \<alpha> n /\<^sub>R norm (\<alpha> n)) \<circ> \<rho>) \<longlonglongrightarrow> 0"
by (simp add: o_def scale)
qed
qed
ultimately show False
using \<open>k > 0\<close> by auto
qed
ultimately have dim: "dim {x. f x = 0} = DIM('a)"
by force
then show ?thesis
using dim_eq_full
by (metis (mono_tags, lifting) eq_0_on_span eucl.span_Basis linear_axioms linear_eq_stdbasis
mem_Collect_eq module_hom_zero span_base span_raw_def)
qed
lemma lemma_partial_derivatives:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "linear f" and lim: "((\<lambda>x. f (x - a) /\<^sub>R norm (x - a)) \<longlongrightarrow> 0) (at a within S)"
and lb: "\<And>v. v \<noteq> 0 \<Longrightarrow> (\<exists>k>0. \<forall>e>0. \<exists>x \<in> S - {a}. norm(a - x) < e \<and> k * norm(a - x) \<le> \<bar>v \<bullet> (x - a)\<bar>)"
shows "f x = 0"
proof -
have "((\<lambda>x. f x /\<^sub>R norm x) \<longlongrightarrow> 0) (at 0 within (\<lambda>x. x-a) ` S)"
using lim by (simp add: Lim_within dist_norm)
then show ?thesis
proof (rule lemma_partial_derivatives0 [OF \<open>linear f\<close>])
fix v :: "'a"
assume v: "v \<noteq> 0"
show "\<exists>k>0. \<forall>e>0. \<exists>x. x \<in> (\<lambda>x. x - a) ` S - {0} \<and> norm x < e \<and> k * norm x \<le> \<bar>v \<bullet> x\<bar>"
using lb [OF v] by (force simp: norm_minus_commute)
qed
qed
proposition borel_measurable_partial_derivatives:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n"
assumes S: "S \<in> sets lebesgue"
and f: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
shows "(\<lambda>x. (matrix(f' x)$m$n)) \<in> borel_measurable (lebesgue_on S)"
proof -
have contf: "continuous_on S f"
using continuous_on_eq_continuous_within f has_derivative_continuous by blast
have "{x \<in> S. (matrix (f' x)$m$n) \<le> b} \<in> sets lebesgue" for b
proof (rule sets_negligible_symdiff)
let ?T = "{x \<in> S. \<forall>e>0. \<exists>d>0. \<exists>A. A$m$n < b \<and> (\<forall>i j. A$i$j \<in> \<rat>) \<and>
(\<forall>y \<in> S. norm(y - x) < d \<longrightarrow> norm(f y - f x - A *v (y - x)) \<le> e * norm(y - x))}"
let ?U = "S \<inter>
(\<Inter>e \<in> {e \<in> \<rat>. e > 0}.
\<Union>A \<in> {A. A$m$n < b \<and> (\<forall>i j. A$i$j \<in> \<rat>)}.
\<Union>d \<in> {d \<in> \<rat>. 0 < d}.
S \<inter> (\<Inter>y \<in> S. {x \<in> S. norm(y - x) < d \<longrightarrow> norm(f y - f x - A *v (y - x)) \<le> e * norm(y - x)}))"
have "?T = ?U"
proof (intro set_eqI iffI)
fix x
assume xT: "x \<in> ?T"
then show "x \<in> ?U"
proof (clarsimp simp add:)
fix q :: real
assume "q \<in> \<rat>" "q > 0"
then obtain d A where "d > 0" and A: "A $ m $ n < b" "\<And>i j. A $ i $ j \<in> \<rat>"
"\<And>y. \<lbrakk>y\<in>S; norm (y - x) < d\<rbrakk> \<Longrightarrow> norm (f y - f x - A *v (y - x)) \<le> q * norm (y - x)"
using xT by auto
then obtain \<delta> where "d > \<delta>" "\<delta> > 0" "\<delta> \<in> \<rat>"
using Rats_dense_in_real by blast
with A show "\<exists>A. A $ m $ n < b \<and> (\<forall>i j. A $ i $ j \<in> \<rat>) \<and>
(\<exists>s. s \<in> \<rat> \<and> 0 < s \<and> (\<forall>y\<in>S. norm (y - x) < s \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> q * norm (y - x)))"
by force
qed
next
fix x
assume xU: "x \<in> ?U"
then show "x \<in> ?T"
proof clarsimp
fix e :: "real"
assume "e > 0"
then obtain \<epsilon> where \<epsilon>: "e > \<epsilon>" "\<epsilon> > 0" "\<epsilon> \<in> \<rat>"
using Rats_dense_in_real by blast
with xU obtain A r where "x \<in> S" and Ar: "A $ m $ n < b" "\<forall>i j. A $ i $ j \<in> \<rat>" "r \<in> \<rat>" "r > 0"
and "\<forall>y\<in>S. norm (y - x) < r \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> \<epsilon> * norm (y - x)"
by (auto simp: split: if_split_asm)
then have "\<forall>y\<in>S. norm (y - x) < r \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> e * norm (y - x)"
by (meson \<open>e > \<epsilon>\<close> less_eq_real_def mult_right_mono norm_ge_zero order_trans)
then show "\<exists>d>0. \<exists>A. A $ m $ n < b \<and> (\<forall>i j. A $ i $ j \<in> \<rat>) \<and> (\<forall>y\<in>S. norm (y - x) < d \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> e * norm (y - x))"
using \<open>x \<in> S\<close> Ar by blast
qed
qed
moreover have "?U \<in> sets lebesgue"
proof -
have coQ: "countable {e \<in> \<rat>. 0 < e}"
using countable_Collect countable_rat by blast
have ne: "{e \<in> \<rat>. (0::real) < e} \<noteq> {}"
using zero_less_one Rats_1 by blast
have coA: "countable {A. A $ m $ n < b \<and> (\<forall>i j. A $ i $ j \<in> \<rat>)}"
proof (rule countable_subset)
show "countable {A. \<forall>i j. A $ i $ j \<in> \<rat>}"
using countable_vector [OF countable_vector, of "\<lambda>i j. \<rat>"] by (simp add: countable_rat)
qed blast
have *: "\<lbrakk>U \<noteq> {} \<Longrightarrow> closedin (top_of_set S) (S \<inter> \<Inter> U)\<rbrakk>
\<Longrightarrow> closedin (top_of_set S) (S \<inter> \<Inter> U)" for U
by fastforce
have eq: "{x::(real,'m)vec. P x \<and> (Q x \<longrightarrow> R x)} = {x. P x \<and> \<not> Q x} \<union> {x. P x \<and> R x}" for P Q R
by auto
have sets: "S \<inter> (\<Inter>y\<in>S. {x \<in> S. norm (y - x) < d \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> e * norm (y - x)})
\<in> sets lebesgue" for e A d
proof -
have clo: "closedin (top_of_set S)
{x \<in> S. norm (y - x) < d \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> e * norm (y - x)}"
for y
proof -
have cont1: "continuous_on S (\<lambda>x. norm (y - x))"
and cont2: "continuous_on S (\<lambda>x. e * norm (y - x) - norm (f y - f x - (A *v y - A *v x)))"
by (force intro: contf continuous_intros)+
have clo1: "closedin (top_of_set S) {x \<in> S. d \<le> norm (y - x)}"
using continuous_closedin_preimage [OF cont1, of "{d..}"] by (simp add: vimage_def Int_def)
have clo2: "closedin (top_of_set S)
{x \<in> S. norm (f y - f x - (A *v y - A *v x)) \<le> e * norm (y - x)}"
using continuous_closedin_preimage [OF cont2, of "{0..}"] by (simp add: vimage_def Int_def)
show ?thesis
by (auto simp: eq not_less matrix_vector_mult_diff_distrib intro: clo1 clo2)
qed
show ?thesis
by (rule lebesgue_closedin [of S]) (force intro: * S clo)+
qed
show ?thesis
by (intro sets sets.Int S sets.countable_UN'' sets.countable_INT'' coQ coA) auto
qed
ultimately show "?T \<in> sets lebesgue"
by simp
let ?M = "(?T - {x \<in> S. matrix (f' x) $ m $ n \<le> b} \<union> ({x \<in> S. matrix (f' x) $ m $ n \<le> b} - ?T))"
let ?\<Theta> = "\<lambda>x v. \<forall>\<xi>>0. \<exists>e>0. \<forall>y \<in> S-{x}. norm (x - y) < e \<longrightarrow> \<bar>v \<bullet> (y - x)\<bar> < \<xi> * norm (x - y)"
have nN: "negligible {x \<in> S. \<exists>v\<noteq>0. ?\<Theta> x v}"
unfolding negligible_eq_zero_density
proof clarsimp
fix x v and r e :: "real"
assume "x \<in> S" "v \<noteq> 0" "r > 0" "e > 0"
and Theta [rule_format]: "?\<Theta> x v"
moreover have "(norm v * e / 2) / CARD('m) ^ CARD('m) > 0"
by (simp add: \<open>v \<noteq> 0\<close> \<open>e > 0\<close>)
ultimately obtain d where "d > 0"
and dless: "\<And>y. \<lbrakk>y \<in> S - {x}; norm (x - y) < d\<rbrakk> \<Longrightarrow>
\<bar>v \<bullet> (y - x)\<bar> < ((norm v * e / 2) / CARD('m) ^ CARD('m)) * norm (x - y)"
by metis
let ?W = "ball x (min d r) \<inter> {y. \<bar>v \<bullet> (y - x)\<bar> < (norm v * e/2 * min d r) / CARD('m) ^ CARD('m)}"
have "open {x. \<bar>v \<bullet> (x - a)\<bar> < b}" for a b
by (intro open_Collect_less continuous_intros)
show "\<exists>d>0. d \<le> r \<and>
(\<exists>U. {x' \<in> S. \<exists>v\<noteq>0. ?\<Theta> x' v} \<inter> ball x d \<subseteq> U \<and>
U \<in> lmeasurable \<and> measure lebesgue U < e * content (ball x d))"
proof (intro exI conjI)
show "0 < min d r" "min d r \<le> r"
using \<open>r > 0\<close> \<open>d > 0\<close> by auto
show "{x' \<in> S. \<exists>v. v \<noteq> 0 \<and> (\<forall>\<xi>>0. \<exists>e>0. \<forall>z\<in>S - {x'}. norm (x' - z) < e \<longrightarrow> \<bar>v \<bullet> (z - x')\<bar> < \<xi> * norm (x' - z))} \<inter> ball x (min d r) \<subseteq> ?W"
proof (clarsimp simp: dist_norm norm_minus_commute)
fix y w
assume "y \<in> S" "w \<noteq> 0"
and less [rule_format]:
"\<forall>\<xi>>0. \<exists>e>0. \<forall>z\<in>S - {y}. norm (y - z) < e \<longrightarrow> \<bar>w \<bullet> (z - y)\<bar> < \<xi> * norm (y - z)"
and d: "norm (y - x) < d" and r: "norm (y - x) < r"
show "\<bar>v \<bullet> (y - x)\<bar> < norm v * e * min d r / (2 * real CARD('m) ^ CARD('m))"
proof (cases "y = x")
case True
with \<open>r > 0\<close> \<open>d > 0\<close> \<open>e > 0\<close> \<open>v \<noteq> 0\<close> show ?thesis
by simp
next
case False
have "\<bar>v \<bullet> (y - x)\<bar> < norm v * e / 2 / real (CARD('m) ^ CARD('m)) * norm (x - y)"
apply (rule dless)
using False \<open>y \<in> S\<close> d by (auto simp: norm_minus_commute)
also have "\<dots> \<le> norm v * e * min d r / (2 * real CARD('m) ^ CARD('m))"
using d r \<open>e > 0\<close> by (simp add: field_simps norm_minus_commute mult_left_mono)
finally show ?thesis .
qed
qed
show "?W \<in> lmeasurable"
by (simp add: fmeasurable_Int_fmeasurable borel_open)
obtain k::'m where True
by metis
obtain T where T: "orthogonal_transformation T" and v: "v = T(norm v *\<^sub>R axis k (1::real))"
using rotation_rightward_line by metis
define b where "b \<equiv> norm v"
have "b > 0"
using \<open>v \<noteq> 0\<close> by (auto simp: b_def)
obtain eqb: "inv T v = b *\<^sub>R axis k (1::real)" and "inj T" "bij T" and invT: "orthogonal_transformation (inv T)"
by (metis UNIV_I b_def T v bij_betw_inv_into_left orthogonal_transformation_inj orthogonal_transformation_bij orthogonal_transformation_inv)
let ?v = "\<chi> i. min d r / CARD('m)"
let ?v' = "\<chi> i. if i = k then (e/2 * min d r) / CARD('m) ^ CARD('m) else min d r"
let ?x' = "inv T x"
let ?W' = "(ball ?x' (min d r) \<inter> {y. \<bar>(y - ?x')$k\<bar> < e * min d r / (2 * CARD('m) ^ CARD('m))})"
have abs: "x - e \<le> y \<and> y \<le> x + e \<longleftrightarrow> abs(y - x) \<le> e" for x y e::real
by auto
have "?W = T ` ?W'"
proof -
have 1: "T ` (ball (inv T x) (min d r)) = ball x (min d r)"
by (simp add: T image_orthogonal_transformation_ball orthogonal_transformation_surj surj_f_inv_f)
have 2: "{y. \<bar>v \<bullet> (y - x)\<bar> < b * e * min d r / (2 * real CARD('m) ^ CARD('m))} =
T ` {y. \<bar>y $ k - ?x' $ k\<bar> < e * min d r / (2 * real CARD('m) ^ CARD('m))}"
proof -
have *: "\<bar>T (b *\<^sub>R axis k 1) \<bullet> (y - x)\<bar> = b * \<bar>inv T y $ k - ?x' $ k\<bar>" for y
proof -
have "\<bar>T (b *\<^sub>R axis k 1) \<bullet> (y - x)\<bar> = \<bar>(b *\<^sub>R axis k 1) \<bullet> inv T (y - x)\<bar>"
by (metis (no_types, hide_lams) b_def eqb invT orthogonal_transformation_def v)
also have "\<dots> = b * \<bar>(axis k 1) \<bullet> inv T (y - x)\<bar>"
using \<open>b > 0\<close> by (simp add: abs_mult)
also have "\<dots> = b * \<bar>inv T y $ k - ?x' $ k\<bar>"
using orthogonal_transformation_linear [OF invT]
by (simp add: inner_axis' linear_diff)
finally show ?thesis
by simp
qed
show ?thesis
using v b_def [symmetric]
using \<open>b > 0\<close> by (simp add: * bij_image_Collect_eq [OF \<open>bij T\<close>] mult_less_cancel_left_pos times_divide_eq_right [symmetric] del: times_divide_eq_right)
qed
show ?thesis
using \<open>b > 0\<close> by (simp add: image_Int \<open>inj T\<close> 1 2 b_def [symmetric])
qed
moreover have "?W' \<in> lmeasurable"
by (auto intro: fmeasurable_Int_fmeasurable)
ultimately have "measure lebesgue ?W = measure lebesgue ?W'"
by (metis measure_orthogonal_image T)
also have "\<dots> \<le> measure lebesgue (cbox (?x' - ?v') (?x' + ?v'))"
proof (rule measure_mono_fmeasurable)
show "?W' \<subseteq> cbox (?x' - ?v') (?x' + ?v')"
apply (clarsimp simp add: mem_box_cart abs dist_norm norm_minus_commute simp del: min_less_iff_conj min.bounded_iff)
by (metis component_le_norm_cart less_eq_real_def le_less_trans vector_minus_component)
qed auto
also have "\<dots> \<le> e/2 * measure lebesgue (cbox (?x' - ?v) (?x' + ?v))"
proof -
have "cbox (?x' - ?v) (?x' + ?v) \<noteq> {}"
using \<open>r > 0\<close> \<open>d > 0\<close> by (auto simp: interval_eq_empty_cart divide_less_0_iff)
with \<open>r > 0\<close> \<open>d > 0\<close> \<open>e > 0\<close> show ?thesis
apply (simp add: content_cbox_if_cart mem_box_cart)
apply (auto simp: prod_nonneg)
apply (simp add: abs if_distrib prod.delta_remove field_simps power_diff split: if_split_asm)
done
qed
also have "\<dots> \<le> e/2 * measure lebesgue (cball ?x' (min d r))"
proof (rule mult_left_mono [OF measure_mono_fmeasurable])
have *: "norm (?x' - y) \<le> min d r"
if y: "\<And>i. \<bar>?x' $ i - y $ i\<bar> \<le> min d r / real CARD('m)" for y
proof -
have "norm (?x' - y) \<le> (\<Sum>i\<in>UNIV. \<bar>(?x' - y) $ i\<bar>)"
by (rule norm_le_l1_cart)
also have "\<dots> \<le> real CARD('m) * (min d r / real CARD('m))"
by (rule sum_bounded_above) (use y in auto)
finally show ?thesis
by simp
qed
show "cbox (?x' - ?v) (?x' + ?v) \<subseteq> cball ?x' (min d r)"
apply (clarsimp simp only: mem_box_cart dist_norm mem_cball intro!: *)
by (simp add: abs_diff_le_iff abs_minus_commute)
qed (use \<open>e > 0\<close> in auto)
also have "\<dots> < e * content (cball ?x' (min d r))"
using \<open>r > 0\<close> \<open>d > 0\<close> \<open>e > 0\<close> by (auto intro: content_cball_pos)
also have "\<dots> = e * content (ball x (min d r))"
using \<open>r > 0\<close> \<open>d > 0\<close> content_ball_conv_unit_ball[of "min d r" "inv T x"]
content_ball_conv_unit_ball[of "min d r" x]
by (simp add: content_cball_conv_ball)
finally show "measure lebesgue ?W < e * content (ball x (min d r))" .
qed
qed
have *: "(\<And>x. (x \<notin> S) \<Longrightarrow> (x \<in> T \<longleftrightarrow> x \<in> U)) \<Longrightarrow> (T - U) \<union> (U - T) \<subseteq> S" for S T U :: "(real,'m) vec set"
by blast
have MN: "?M \<subseteq> {x \<in> S. \<exists>v\<noteq>0. ?\<Theta> x v}"
proof (rule *)
fix x
assume x: "x \<notin> {x \<in> S. \<exists>v\<noteq>0. ?\<Theta> x v}"
show "(x \<in> ?T) \<longleftrightarrow> (x \<in> {x \<in> S. matrix (f' x) $ m $ n \<le> b})"
proof (cases "x \<in> S")
case True
then have x: "\<not> ?\<Theta> x v" if "v \<noteq> 0" for v
using x that by force
show ?thesis
proof (rule iffI; clarsimp)
assume b: "\<forall>e>0. \<exists>d>0. \<exists>A. A $ m $ n < b \<and> (\<forall>i j. A $ i $ j \<in> \<rat>) \<and>
(\<forall>y\<in>S. norm (y - x) < d \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> e * norm (y - x))"
(is "\<forall>e>0. \<exists>d>0. \<exists>A. ?\<Phi> e d A")
then have "\<forall>k. \<exists>d>0. \<exists>A. ?\<Phi> (1 / Suc k) d A"
by (metis (no_types, hide_lams) less_Suc_eq_0_disj of_nat_0_less_iff zero_less_divide_1_iff)
then obtain \<delta> A where \<delta>: "\<And>k. \<delta> k > 0"
and Ab: "\<And>k. A k $ m $ n < b"
and A: "\<And>k y. \<lbrakk>y \<in> S; norm (y - x) < \<delta> k\<rbrakk> \<Longrightarrow>
norm (f y - f x - A k *v (y - x)) \<le> 1/(Suc k) * norm (y - x)"
by metis
have "\<forall>i j. \<exists>a. (\<lambda>n. A n $ i $ j) \<longlonglongrightarrow> a"
proof (intro allI)
fix i j
have vax: "(A n *v axis j 1) $ i = A n $ i $ j" for n
by (metis cart_eq_inner_axis matrix_vector_mul_component)
let ?CA = "{x. Cauchy (\<lambda>n. (A n) *v x)}"
have "subspace ?CA"
unfolding subspace_def convergent_eq_Cauchy [symmetric]
by (force simp: algebra_simps intro: tendsto_intros)
then have CA_eq: "?CA = span ?CA"
by (metis span_eq_iff)
also have "\<dots> = UNIV"
proof -
have "dim ?CA \<le> CARD('m)"
using dim_subset_UNIV[of ?CA]
by auto
moreover have "False" if less: "dim ?CA < CARD('m)"
proof -
obtain d where "d \<noteq> 0" and d: "\<And>y. y \<in> span ?CA \<Longrightarrow> orthogonal d y"
using less by (force intro: orthogonal_to_subspace_exists [of ?CA])
with x [OF \<open>d \<noteq> 0\<close>] obtain \<xi> where "\<xi> > 0"
and \<xi>: "\<And>e. e > 0 \<Longrightarrow> \<exists>y \<in> S - {x}. norm (x - y) < e \<and> \<xi> * norm (x - y) \<le> \<bar>d \<bullet> (y - x)\<bar>"
by (fastforce simp: not_le Bex_def)
obtain \<gamma> z where \<gamma>Sx: "\<And>i. \<gamma> i \<in> S - {x}"
and \<gamma>le: "\<And>i. \<xi> * norm(\<gamma> i - x) \<le> \<bar>d \<bullet> (\<gamma> i - x)\<bar>"
and \<gamma>x: "\<gamma> \<longlonglongrightarrow> x"
and z: "(\<lambda>n. (\<gamma> n - x) /\<^sub>R norm (\<gamma> n - x)) \<longlonglongrightarrow> z"
proof -
have "\<exists>\<gamma>. (\<forall>i. (\<gamma> i \<in> S - {x} \<and>
\<xi> * norm(\<gamma> i - x) \<le> \<bar>d \<bullet> (\<gamma> i - x)\<bar> \<and> norm(\<gamma> i - x) < 1/Suc i) \<and>
norm(\<gamma>(Suc i) - x) < norm(\<gamma> i - x))"
proof (rule dependent_nat_choice)
show "\<exists>y. y \<in> S - {x} \<and> \<xi> * norm (y - x) \<le> \<bar>d \<bullet> (y - x)\<bar> \<and> norm (y - x) < 1 / Suc 0"
using \<xi> [of 1] by (auto simp: dist_norm norm_minus_commute)
next
fix y i
assume "y \<in> S - {x} \<and> \<xi> * norm (y - x) \<le> \<bar>d \<bullet> (y - x)\<bar> \<and> norm (y - x) < 1/Suc i"
then have "min (norm(y - x)) (1/((Suc i) + 1)) > 0"
by auto
then obtain y' where "y' \<in> S - {x}" and y': "norm (x - y') < min (norm (y - x)) (1/((Suc i) + 1))"
"\<xi> * norm (x - y') \<le> \<bar>d \<bullet> (y' - x)\<bar>"
using \<xi> by metis
with \<xi> show "\<exists>y'. (y' \<in> S - {x} \<and> \<xi> * norm (y' - x) \<le> \<bar>d \<bullet> (y' - x)\<bar> \<and>
norm (y' - x) < 1/(Suc (Suc i))) \<and> norm (y' - x) < norm (y - x)"
by (auto simp: dist_norm norm_minus_commute)
qed
then obtain \<gamma> where
\<gamma>Sx: "\<And>i. \<gamma> i \<in> S - {x}"
and \<gamma>le: "\<And>i. \<xi> * norm(\<gamma> i - x) \<le> \<bar>d \<bullet> (\<gamma> i - x)\<bar>"
and \<gamma>conv: "\<And>i. norm(\<gamma> i - x) < 1/(Suc i)"
by blast
let ?f = "\<lambda>i. (\<gamma> i - x) /\<^sub>R norm (\<gamma> i - x)"
have "?f i \<in> sphere 0 1" for i
using \<gamma>Sx by auto
then obtain l \<rho> where "l \<in> sphere 0 1" "strict_mono \<rho>" and l: "(?f \<circ> \<rho>) \<longlonglongrightarrow> l"
using compact_sphere [of "0::(real,'m) vec" 1] unfolding compact_def by meson
show thesis
proof
show "(\<gamma> \<circ> \<rho>) i \<in> S - {x}" "\<xi> * norm ((\<gamma> \<circ> \<rho>) i - x) \<le> \<bar>d \<bullet> ((\<gamma> \<circ> \<rho>) i - x)\<bar>" for i
using \<gamma>Sx \<gamma>le by auto
have "\<gamma> \<longlonglongrightarrow> x"
proof (clarsimp simp add: LIMSEQ_def dist_norm)
fix r :: "real"
assume "r > 0"
with real_arch_invD obtain no where "no \<noteq> 0" "real no > 1/r"
by (metis divide_less_0_1_iff not_less_iff_gr_or_eq of_nat_0_eq_iff reals_Archimedean2)
with \<gamma>conv show "\<exists>no. \<forall>n\<ge>no. norm (\<gamma> n - x) < r"
by (metis \<open>r > 0\<close> add.commute divide_inverse inverse_inverse_eq inverse_less_imp_less less_trans mult.left_neutral nat_le_real_less of_nat_Suc)
qed
with \<open>strict_mono \<rho>\<close> show "(\<gamma> \<circ> \<rho>) \<longlonglongrightarrow> x"
by (metis LIMSEQ_subseq_LIMSEQ)
show "(\<lambda>n. ((\<gamma> \<circ> \<rho>) n - x) /\<^sub>R norm ((\<gamma> \<circ> \<rho>) n - x)) \<longlonglongrightarrow> l"
using l by (auto simp: o_def)
qed
qed
have "isCont (\<lambda>x. (\<bar>d \<bullet> x\<bar> - \<xi>)) z"
by (intro continuous_intros)
from isCont_tendsto_compose [OF this z]
have lim: "(\<lambda>y. \<bar>d \<bullet> ((\<gamma> y - x) /\<^sub>R norm (\<gamma> y - x))\<bar> - \<xi>) \<longlonglongrightarrow> \<bar>d \<bullet> z\<bar> - \<xi>"
by auto
moreover have "\<forall>\<^sub>F i in sequentially. 0 \<le> \<bar>d \<bullet> ((\<gamma> i - x) /\<^sub>R norm (\<gamma> i - x))\<bar> - \<xi>"
proof (rule eventuallyI)
fix n
show "0 \<le> \<bar>d \<bullet> ((\<gamma> n - x) /\<^sub>R norm (\<gamma> n - x))\<bar> - \<xi>"
using \<gamma>le [of n] \<gamma>Sx by (auto simp: abs_mult divide_simps)
qed
ultimately have "\<xi> \<le> \<bar>d \<bullet> z\<bar>"
using tendsto_lowerbound [where a=0] by fastforce
have "Cauchy (\<lambda>n. (A n) *v z)"
proof (clarsimp simp add: Cauchy_def)
fix \<epsilon> :: "real"
assume "0 < \<epsilon>"
then obtain N::nat where "N > 0" and N: "\<epsilon>/2 > 1/N"
by (metis half_gt_zero inverse_eq_divide neq0_conv real_arch_inverse)
show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (A m *v z) (A n *v z) < \<epsilon>"
proof (intro exI allI impI)
fix i j
assume ij: "N \<le> i" "N \<le> j"
let ?V = "\<lambda>i k. A i *v ((\<gamma> k - x) /\<^sub>R norm (\<gamma> k - x))"
have "\<forall>\<^sub>F k in sequentially. dist (\<gamma> k) x < min (\<delta> i) (\<delta> j)"
using \<gamma>x [unfolded tendsto_iff] by (meson min_less_iff_conj \<delta>)
then have even: "\<forall>\<^sub>F k in sequentially. norm (?V i k - ?V j k) - 2 / N \<le> 0"
proof (rule eventually_mono, clarsimp)
fix p
assume p: "dist (\<gamma> p) x < \<delta> i" "dist (\<gamma> p) x < \<delta> j"
let ?C = "\<lambda>k. f (\<gamma> p) - f x - A k *v (\<gamma> p - x)"
have "norm ((A i - A j) *v (\<gamma> p - x)) = norm (?C j - ?C i)"
by (simp add: algebra_simps)
also have "\<dots> \<le> norm (?C j) + norm (?C i)"
using norm_triangle_ineq4 by blast
also have "\<dots> \<le> 1/(Suc j) * norm (\<gamma> p - x) + 1/(Suc i) * norm (\<gamma> p - x)"
by (metis A Diff_iff \<gamma>Sx dist_norm p add_mono)
also have "\<dots> \<le> 1/N * norm (\<gamma> p - x) + 1/N * norm (\<gamma> p - x)"
apply (intro add_mono mult_right_mono)
using ij \<open>N > 0\<close> by (auto simp: field_simps)
also have "\<dots> = 2 / N * norm (\<gamma> p - x)"
by simp
finally have no_le: "norm ((A i - A j) *v (\<gamma> p - x)) \<le> 2 / N * norm (\<gamma> p - x)" .
have "norm (?V i p - ?V j p) =
norm ((A i - A j) *v ((\<gamma> p - x) /\<^sub>R norm (\<gamma> p - x)))"
by (simp add: algebra_simps)
also have "\<dots> = norm ((A i - A j) *v (\<gamma> p - x)) / norm (\<gamma> p - x)"
by (simp add: divide_inverse matrix_vector_mult_scaleR)
also have "\<dots> \<le> 2 / N"
using no_le by (auto simp: field_split_simps)
finally show "norm (?V i p - ?V j p) \<le> 2 / N" .
qed
have "isCont (\<lambda>w. (norm(A i *v w - A j *v w) - 2 / N)) z"
by (intro continuous_intros)
from isCont_tendsto_compose [OF this z]
have lim: "(\<lambda>w. norm (A i *v ((\<gamma> w - x) /\<^sub>R norm (\<gamma> w - x)) -
A j *v ((\<gamma> w - x) /\<^sub>R norm (\<gamma> w - x))) - 2 / N)
\<longlonglongrightarrow> norm (A i *v z - A j *v z) - 2 / N"
by auto
have "dist (A i *v z) (A j *v z) \<le> 2 / N"
using tendsto_upperbound [OF lim even] by (auto simp: dist_norm)
with N show "dist (A i *v z) (A j *v z) < \<epsilon>"
by linarith
qed
qed
then have "d \<bullet> z = 0"
using CA_eq d orthogonal_def by auto
then show False
using \<open>0 < \<xi>\<close> \<open>\<xi> \<le> \<bar>d \<bullet> z\<bar>\<close> by auto
qed
ultimately show ?thesis
using dim_eq_full by fastforce
qed
finally have "?CA = UNIV" .
then have "Cauchy (\<lambda>n. (A n) *v axis j 1)"
by auto
then obtain L where "(\<lambda>n. A n *v axis j 1) \<longlonglongrightarrow> L"
by (auto simp: Cauchy_convergent_iff convergent_def)
then have "(\<lambda>x. (A x *v axis j 1) $ i) \<longlonglongrightarrow> L $ i"
by (rule tendsto_vec_nth)
then show "\<exists>a. (\<lambda>n. A n $ i $ j) \<longlonglongrightarrow> a"
by (force simp: vax)
qed
then obtain B where B: "\<And>i j. (\<lambda>n. A n $ i $ j) \<longlonglongrightarrow> B $ i $ j"
by (auto simp: lambda_skolem)
have lin_df: "linear (f' x)"
and lim_df: "((\<lambda>y. (1 / norm (y - x)) *\<^sub>R (f y - (f x + f' x (y - x)))) \<longlongrightarrow> 0) (at x within S)"
using \<open>x \<in> S\<close> assms by (auto simp: has_derivative_within linear_linear)
moreover
interpret linear "f' x" by fact
have "(matrix (f' x) - B) *v w = 0" for w
proof (rule lemma_partial_derivatives [of "(*v) (matrix (f' x) - B)"])
show "linear ((*v) (matrix (f' x) - B))"
by (rule matrix_vector_mul_linear)
have "((\<lambda>y. ((f x + f' x (y - x)) - f y) /\<^sub>R norm (y - x)) \<longlongrightarrow> 0) (at x within S)"
- using tendsto_minus [OF lim_df] by (simp add: algebra_simps field_split_simps)
+ using tendsto_minus [OF lim_df] by (simp add: field_split_simps)
then show "((\<lambda>y. (matrix (f' x) - B) *v (y - x) /\<^sub>R norm (y - x)) \<longlongrightarrow> 0) (at x within S)"
proof (rule Lim_transform)
have "((\<lambda>y. ((f y + B *v x - (f x + B *v y)) /\<^sub>R norm (y - x))) \<longlongrightarrow> 0) (at x within S)"
proof (clarsimp simp add: Lim_within dist_norm)
fix e :: "real"
assume "e > 0"
then obtain q::nat where "q \<noteq> 0" and qe2: "1/q < e/2"
by (metis divide_pos_pos inverse_eq_divide real_arch_inverse zero_less_numeral)
let ?g = "\<lambda>p. sum (\<lambda>i. sum (\<lambda>j. abs((A p - B)$i$j)) UNIV) UNIV"
have "(\<lambda>k. onorm (\<lambda>y. (A k - B) *v y)) \<longlonglongrightarrow> 0"
proof (rule Lim_null_comparison)
show "\<forall>\<^sub>F k in sequentially. norm (onorm (\<lambda>y. (A k - B) *v y)) \<le> ?g k"
proof (rule eventually_sequentiallyI)
fix k :: "nat"
assume "0 \<le> k"
have "0 \<le> onorm ((*v) (A k - B))"
using matrix_vector_mul_bounded_linear
by (rule onorm_pos_le)
then show "norm (onorm ((*v) (A k - B))) \<le> (\<Sum>i\<in>UNIV. \<Sum>j\<in>UNIV. \<bar>(A k - B) $ i $ j\<bar>)"
by (simp add: onorm_le_matrix_component_sum del: vector_minus_component)
qed
next
show "?g \<longlonglongrightarrow> 0"
using B Lim_null tendsto_rabs_zero_iff by (fastforce intro!: tendsto_null_sum)
qed
with \<open>e > 0\<close> obtain p where "\<And>n. n \<ge> p \<Longrightarrow> \<bar>onorm ((*v) (A n - B))\<bar> < e/2"
unfolding lim_sequentially by (metis diff_zero dist_real_def divide_pos_pos zero_less_numeral)
then have pqe2: "\<bar>onorm ((*v) (A (p + q) - B))\<bar> < e/2" (*17 [`abs (onorm (\y. A (p + q) ** y - B ** y)) < e / &2`]*)
using le_add1 by blast
show "\<exists>d>0. \<forall>y\<in>S. y \<noteq> x \<and> norm (y - x) < d \<longrightarrow>
inverse (norm (y - x)) * norm (f y + B *v x - (f x + B *v y)) < e"
proof (intro exI, safe)
show "0 < \<delta>(p + q)"
by (simp add: \<delta>)
next
fix y
assume y: "y \<in> S" "norm (y - x) < \<delta>(p + q)" and "y \<noteq> x"
have *: "\<lbrakk>norm(b - c) < e - d; norm(y - x - b) \<le> d\<rbrakk> \<Longrightarrow> norm(y - x - c) < e"
for b c d e x and y:: "real^'n"
using norm_triangle_ineq2 [of "y - x - c" "y - x - b"] by simp
have "norm (f y - f x - B *v (y - x)) < e * norm (y - x)"
proof (rule *)
show "norm (f y - f x - A (p + q) *v (y - x)) \<le> norm (y - x) / (Suc (p + q))"
using A [OF y] by simp
have "norm (A (p + q) *v (y - x) - B *v (y - x)) \<le> onorm(\<lambda>x. (A(p + q) - B) *v x) * norm(y - x)"
by (metis linear_linear matrix_vector_mul_linear matrix_vector_mult_diff_rdistrib onorm)
also have "\<dots> < (e/2) * norm (y - x)"
using \<open>y \<noteq> x\<close> pqe2 by auto
also have "\<dots> \<le> (e - 1 / (Suc (p + q))) * norm (y - x)"
proof (rule mult_right_mono)
have "1 / Suc (p + q) \<le> 1 / q"
using \<open>q \<noteq> 0\<close> by (auto simp: field_split_simps)
also have "\<dots> < e/2"
using qe2 by auto
finally show "e / 2 \<le> e - 1 / real (Suc (p + q))"
by linarith
qed auto
finally show "norm (A (p + q) *v (y - x) - B *v (y - x)) < e * norm (y - x) - norm (y - x) / real (Suc (p + q))"
by (simp add: algebra_simps)
qed
then show "inverse (norm (y - x)) * norm (f y + B *v x - (f x + B *v y)) < e"
using \<open>y \<noteq> x\<close> by (simp add: field_split_simps algebra_simps)
qed
qed
then show "((\<lambda>y. (matrix (f' x) - B) *v (y - x) /\<^sub>R
norm (y - x) - (f x + f' x (y - x) - f y) /\<^sub>R norm (y - x)) \<longlongrightarrow> 0)
(at x within S)"
- by (simp add: algebra_simps diff lin_df matrix_vector_mul_linear scalar_mult_eq_scaleR)
+ by (simp add: algebra_simps diff lin_df scalar_mult_eq_scaleR)
qed
qed (use x in \<open>simp; auto simp: not_less\<close>)
ultimately have "f' x = (*v) B"
by (force simp: algebra_simps scalar_mult_eq_scaleR)
show "matrix (f' x) $ m $ n \<le> b"
proof (rule tendsto_upperbound [of "\<lambda>i. (A i $ m $ n)" _ sequentially])
show "(\<lambda>i. A i $ m $ n) \<longlonglongrightarrow> matrix (f' x) $ m $ n"
by (simp add: B \<open>f' x = (*v) B\<close>)
show "\<forall>\<^sub>F i in sequentially. A i $ m $ n \<le> b"
by (simp add: Ab less_eq_real_def)
qed auto
next
fix e :: "real"
assume "x \<in> S" and b: "matrix (f' x) $ m $ n \<le> b" and "e > 0"
then obtain d where "d>0"
and d: "\<And>y. y\<in>S \<Longrightarrow> 0 < dist y x \<and> dist y x < d \<longrightarrow> norm (f y - f x - f' x (y - x)) / (norm (y - x))
< e/2"
using f [OF \<open>x \<in> S\<close>]
by (simp add: Deriv.has_derivative_at_within Lim_within)
(auto simp add: field_simps dest: spec [of _ "e/2"])
let ?A = "matrix(f' x) - (\<chi> i j. if i = m \<and> j = n then e / 4 else 0)"
obtain B where BRats: "\<And>i j. B$i$j \<in> \<rat>" and Bo_e6: "onorm((*v) (?A - B)) < e/6"
using matrix_rational_approximation \<open>e > 0\<close>
by (metis zero_less_divide_iff zero_less_numeral)
show "\<exists>d>0. \<exists>A. A $ m $ n < b \<and> (\<forall>i j. A $ i $ j \<in> \<rat>) \<and>
(\<forall>y\<in>S. norm (y - x) < d \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> e * norm (y - x))"
proof (intro exI conjI ballI allI impI)
show "d>0"
by (rule \<open>d>0\<close>)
show "B $ m $ n < b"
proof -
have "\<bar>matrix ((*v) (?A - B)) $ m $ n\<bar> \<le> onorm ((*v) (?A - B))"
using component_le_onorm [OF matrix_vector_mul_linear, of _ m n] by metis
then show ?thesis
using b Bo_e6 by simp
qed
show "B $ i $ j \<in> \<rat>" for i j
using BRats by auto
show "norm (f y - f x - B *v (y - x)) \<le> e * norm (y - x)"
if "y \<in> S" and y: "norm (y - x) < d" for y
proof (cases "y = x")
case True then show ?thesis
by simp
next
case False
have *: "norm(d' - d) \<le> e/2 \<Longrightarrow> norm(y - (x + d')) < e/2 \<Longrightarrow> norm(y - x - d) \<le> e" for d d' e and x y::"real^'n"
using norm_triangle_le [of "d' - d" "y - (x + d')"] by simp
show ?thesis
proof (rule *)
have split246: "\<lbrakk>norm y \<le> e / 6; norm(x - y) \<le> e / 4\<rbrakk> \<Longrightarrow> norm x \<le> e/2" if "e > 0" for e and x y :: "real^'n"
using norm_triangle_le [of y "x-y" "e/2"] \<open>e > 0\<close> by simp
have "linear (f' x)"
using True f has_derivative_linear by blast
then have "norm (f' x (y - x) - B *v (y - x)) = norm ((matrix (f' x) - B) *v (y - x))"
by (simp add: matrix_vector_mult_diff_rdistrib)
also have "\<dots> \<le> (e * norm (y - x)) / 2"
proof (rule split246)
have "norm ((?A - B) *v (y - x)) / norm (y - x) \<le> onorm(\<lambda>x. (?A - B) *v x)"
by (rule le_onorm) auto
also have "\<dots> < e/6"
by (rule Bo_e6)
finally have "norm ((?A - B) *v (y - x)) / norm (y - x) < e / 6" .
then show "norm ((?A - B) *v (y - x)) \<le> e * norm (y - x) / 6"
by (simp add: field_split_simps False)
have "norm ((matrix (f' x) - B) *v (y - x) - ((?A - B) *v (y - x))) = norm ((\<chi> i j. if i = m \<and> j = n then e / 4 else 0) *v (y - x))"
by (simp add: algebra_simps)
also have "\<dots> = norm((e/4) *\<^sub>R (y - x)$n *\<^sub>R axis m (1::real))"
proof -
have "(\<Sum>j\<in>UNIV. (if i = m \<and> j = n then e / 4 else 0) * (y $ j - x $ j)) * 4 = e * (y $ n - x $ n) * axis m 1 $ i" for i
proof (cases "i=m")
case True then show ?thesis
by (auto simp: if_distrib [of "\<lambda>z. z * _"] cong: if_cong)
next
case False then show ?thesis
by (simp add: axis_def)
qed
then have "(\<chi> i j. if i = m \<and> j = n then e / 4 else 0) *v (y - x) = (e/4) *\<^sub>R (y - x)$n *\<^sub>R axis m (1::real)"
by (auto simp: vec_eq_iff matrix_vector_mult_def)
then show ?thesis
by metis
qed
also have "\<dots> \<le> e * norm (y - x) / 4"
using \<open>e > 0\<close> apply (simp add: norm_mult abs_mult)
by (metis component_le_norm_cart vector_minus_component)
finally show "norm ((matrix (f' x) - B) *v (y - x) - ((?A - B) *v (y - x))) \<le> e * norm (y - x) / 4" .
show "0 < e * norm (y - x)"
by (simp add: False \<open>e > 0\<close>)
qed
finally show "norm (f' x (y - x) - B *v (y - x)) \<le> (e * norm (y - x)) / 2" .
show "norm (f y - (f x + f' x (y - x))) < (e * norm (y - x)) / 2"
using False d [OF \<open>y \<in> S\<close>] y by (simp add: dist_norm field_simps)
qed
qed
qed
qed
qed auto
qed
show "negligible ?M"
using negligible_subset [OF nN MN] .
qed
then show ?thesis
by (simp add: borel_measurable_vimage_halfspace_component_le sets_restrict_space_iff assms)
qed
theorem borel_measurable_det_Jacobian:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
assumes S: "S \<in> sets lebesgue" and f: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
shows "(\<lambda>x. det(matrix(f' x))) \<in> borel_measurable (lebesgue_on S)"
unfolding det_def
by (intro measurable) (auto intro: f borel_measurable_partial_derivatives [OF S])
text\<open>The localisation wrt S uses the same argument for many similar results.\<close>
(*See HOL Light's MEASURABLE_ON_LEBESGUE_MEASURABLE_PREIMAGE_BOREL, etc.*)
theorem borel_measurable_lebesgue_on_preimage_borel:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "S \<in> sets lebesgue"
shows "f \<in> borel_measurable (lebesgue_on S) \<longleftrightarrow>
(\<forall>T. T \<in> sets borel \<longrightarrow> {x \<in> S. f x \<in> T} \<in> sets lebesgue)"
proof -
have "{x. (if x \<in> S then f x else 0) \<in> T} \<in> sets lebesgue \<longleftrightarrow> {x \<in> S. f x \<in> T} \<in> sets lebesgue"
if "T \<in> sets borel" for T
proof (cases "0 \<in> T")
case True
then have "{x \<in> S. f x \<in> T} = {x. (if x \<in> S then f x else 0) \<in> T} \<inter> S"
"{x. (if x \<in> S then f x else 0) \<in> T} = {x \<in> S. f x \<in> T} \<union> -S"
by auto
then show ?thesis
by (metis (no_types, lifting) Compl_in_sets_lebesgue assms sets.Int sets.Un)
next
case False
then have "{x. (if x \<in> S then f x else 0) \<in> T} = {x \<in> S. f x \<in> T}"
by auto
then show ?thesis
by auto
qed
then show ?thesis
unfolding borel_measurable_lebesgue_preimage_borel borel_measurable_if [OF assms, symmetric]
by blast
qed
lemma sets_lebesgue_almost_borel:
assumes "S \<in> sets lebesgue"
obtains B N where "B \<in> sets borel" "negligible N" "B \<union> N = S"
proof -
obtain T N N' where "S = T \<union> N" "N \<subseteq> N'" "N' \<in> null_sets lborel" "T \<in> sets borel"
using sets_completionE [OF assms] by auto
then show thesis
by (metis negligible_iff_null_sets negligible_subset null_sets_completionI that)
qed
lemma double_lebesgue_sets:
assumes S: "S \<in> sets lebesgue" and T: "T \<in> sets lebesgue" and fim: "f ` S \<subseteq> T"
shows "(\<forall>U. U \<in> sets lebesgue \<and> U \<subseteq> T \<longrightarrow> {x \<in> S. f x \<in> U} \<in> sets lebesgue) \<longleftrightarrow>
f \<in> borel_measurable (lebesgue_on S) \<and>
(\<forall>U. negligible U \<and> U \<subseteq> T \<longrightarrow> {x \<in> S. f x \<in> U} \<in> sets lebesgue)"
(is "?lhs \<longleftrightarrow> _ \<and> ?rhs")
unfolding borel_measurable_lebesgue_on_preimage_borel [OF S]
proof (intro iffI allI conjI impI, safe)
fix V :: "'b set"
assume *: "\<forall>U. U \<in> sets lebesgue \<and> U \<subseteq> T \<longrightarrow> {x \<in> S. f x \<in> U} \<in> sets lebesgue"
and "V \<in> sets borel"
then have V: "V \<in> sets lebesgue"
by simp
have "{x \<in> S. f x \<in> V} = {x \<in> S. f x \<in> T \<inter> V}"
using fim by blast
also have "{x \<in> S. f x \<in> T \<inter> V} \<in> sets lebesgue"
using T V * le_inf_iff by blast
finally show "{x \<in> S. f x \<in> V} \<in> sets lebesgue" .
next
fix U :: "'b set"
assume "\<forall>U. U \<in> sets lebesgue \<and> U \<subseteq> T \<longrightarrow> {x \<in> S. f x \<in> U} \<in> sets lebesgue"
"negligible U" "U \<subseteq> T"
then show "{x \<in> S. f x \<in> U} \<in> sets lebesgue"
using negligible_imp_sets by blast
next
fix U :: "'b set"
assume 1 [rule_format]: "(\<forall>T. T \<in> sets borel \<longrightarrow> {x \<in> S. f x \<in> T} \<in> sets lebesgue)"
and 2 [rule_format]: "\<forall>U. negligible U \<and> U \<subseteq> T \<longrightarrow> {x \<in> S. f x \<in> U} \<in> sets lebesgue"
and "U \<in> sets lebesgue" "U \<subseteq> T"
then obtain C N where C: "C \<in> sets borel \<and> negligible N \<and> C \<union> N = U"
using sets_lebesgue_almost_borel
by metis
then have "{x \<in> S. f x \<in> C} \<in> sets lebesgue"
by (blast intro: 1)
moreover have "{x \<in> S. f x \<in> N} \<in> sets lebesgue"
using C \<open>U \<subseteq> T\<close> by (blast intro: 2)
moreover have "{x \<in> S. f x \<in> C \<union> N} = {x \<in> S. f x \<in> C} \<union> {x \<in> S. f x \<in> N}"
by auto
ultimately show "{x \<in> S. f x \<in> U} \<in> sets lebesgue"
using C by auto
qed
subsection\<open>Simplest case of Sard's theorem (we don't need continuity of derivative)\<close>
lemma Sard_lemma00:
fixes P :: "'b::euclidean_space set"
assumes "a \<ge> 0" and a: "a *\<^sub>R i \<noteq> 0" and i: "i \<in> Basis"
and P: "P \<subseteq> {x. a *\<^sub>R i \<bullet> x = 0}"
and "0 \<le> m" "0 \<le> e"
obtains S where "S \<in> lmeasurable"
and "{z. norm z \<le> m \<and> (\<exists>t \<in> P. norm(z - t) \<le> e)} \<subseteq> S"
and "measure lebesgue S \<le> (2 * e) * (2 * m) ^ (DIM('b) - 1)"
proof -
have "a > 0"
using assms by simp
let ?v = "(\<Sum>j\<in>Basis. (if j = i then e else m) *\<^sub>R j)"
show thesis
proof
have "- e \<le> x \<bullet> i" "x \<bullet> i \<le> e"
if "t \<in> P" "norm (x - t) \<le> e" for x t
using \<open>a > 0\<close> that Basis_le_norm [of i "x-t"] P i
by (auto simp: inner_commute algebra_simps)
moreover have "- m \<le> x \<bullet> j" "x \<bullet> j \<le> m"
if "norm x \<le> m" "t \<in> P" "norm (x - t) \<le> e" "j \<in> Basis" and "j \<noteq> i"
for x t j
using that Basis_le_norm [of j x] by auto
ultimately
show "{z. norm z \<le> m \<and> (\<exists>t\<in>P. norm (z - t) \<le> e)} \<subseteq> cbox (-?v) ?v"
by (auto simp: mem_box)
have *: "\<forall>k\<in>Basis. - ?v \<bullet> k \<le> ?v \<bullet> k"
using \<open>0 \<le> m\<close> \<open>0 \<le> e\<close> by (auto simp: inner_Basis)
have 2: "2 ^ DIM('b) = 2 * 2 ^ (DIM('b) - Suc 0)"
by (metis DIM_positive Suc_pred power_Suc)
show "measure lebesgue (cbox (-?v) ?v) \<le> 2 * e * (2 * m) ^ (DIM('b) - 1)"
using \<open>i \<in> Basis\<close>
by (simp add: content_cbox [OF *] prod.distrib prod.If_cases Diff_eq [symmetric] 2)
qed blast
qed
text\<open>As above, but reorienting the vector (HOL Light's @text{GEOM\_BASIS\_MULTIPLE\_TAC})\<close>
lemma Sard_lemma0:
fixes P :: "(real^'n::{finite,wellorder}) set"
assumes "a \<noteq> 0"
and P: "P \<subseteq> {x. a \<bullet> x = 0}" and "0 \<le> m" "0 \<le> e"
obtains S where "S \<in> lmeasurable"
and "{z. norm z \<le> m \<and> (\<exists>t \<in> P. norm(z - t) \<le> e)} \<subseteq> S"
and "measure lebesgue S \<le> (2 * e) * (2 * m) ^ (CARD('n) - 1)"
proof -
obtain T and k::'n where T: "orthogonal_transformation T" and a: "a = T (norm a *\<^sub>R axis k (1::real))"
using rotation_rightward_line by metis
have Tinv [simp]: "T (inv T x) = x" for x
by (simp add: T orthogonal_transformation_surj surj_f_inv_f)
obtain S where S: "S \<in> lmeasurable"
and subS: "{z. norm z \<le> m \<and> (\<exists>t \<in> T-`P. norm(z - t) \<le> e)} \<subseteq> S"
and mS: "measure lebesgue S \<le> (2 * e) * (2 * m) ^ (CARD('n) - 1)"
proof (rule Sard_lemma00 [of "norm a" "axis k (1::real)" "T-`P" m e])
have "norm a *\<^sub>R axis k 1 \<bullet> x = 0" if "T x \<in> P" for x
proof -
have "a \<bullet> T x = 0"
using P that by blast
then show ?thesis
by (metis (no_types, lifting) T a orthogonal_orthogonal_transformation orthogonal_def)
qed
then show "T -` P \<subseteq> {x. norm a *\<^sub>R axis k 1 \<bullet> x = 0}"
by auto
qed (use assms T in auto)
show thesis
proof
show "T ` S \<in> lmeasurable"
using S measurable_orthogonal_image T by blast
have "{z. norm z \<le> m \<and> (\<exists>t\<in>P. norm (z - t) \<le> e)} \<subseteq> T ` {z. norm z \<le> m \<and> (\<exists>t\<in>T -` P. norm (z - t) \<le> e)}"
proof clarsimp
fix x t
assume "norm x \<le> m" "t \<in> P" "norm (x - t) \<le> e"
then have "norm (inv T x) \<le> m"
using orthogonal_transformation_inv [OF T] by (simp add: orthogonal_transformation_norm)
moreover have "\<exists>t\<in>T -` P. norm (inv T x - t) \<le> e"
proof
have "T (inv T x - inv T t) = x - t"
using T linear_diff orthogonal_transformation_def
by (metis (no_types, hide_lams) Tinv)
then have "norm (inv T x - inv T t) = norm (x - t)"
by (metis T orthogonal_transformation_norm)
then show "norm (inv T x - inv T t) \<le> e"
using \<open>norm (x - t) \<le> e\<close> by linarith
next
show "inv T t \<in> T -` P"
using \<open>t \<in> P\<close> by force
qed
ultimately show "x \<in> T ` {z. norm z \<le> m \<and> (\<exists>t\<in>T -` P. norm (z - t) \<le> e)}"
by force
qed
then show "{z. norm z \<le> m \<and> (\<exists>t\<in>P. norm (z - t) \<le> e)} \<subseteq> T ` S"
using image_mono [OF subS] by (rule order_trans)
show "measure lebesgue (T ` S) \<le> 2 * e * (2 * m) ^ (CARD('n) - 1)"
using mS T by (simp add: S measure_orthogonal_image)
qed
qed
text\<open>As above, but translating the sets (HOL Light's @text{GEN\_GEOM\_ORIGIN\_TAC})\<close>
lemma Sard_lemma1:
fixes P :: "(real^'n::{finite,wellorder}) set"
assumes P: "dim P < CARD('n)" and "0 \<le> m" "0 \<le> e"
obtains S where "S \<in> lmeasurable"
and "{z. norm(z - w) \<le> m \<and> (\<exists>t \<in> P. norm(z - w - t) \<le> e)} \<subseteq> S"
and "measure lebesgue S \<le> (2 * e) * (2 * m) ^ (CARD('n) - 1)"
proof -
obtain a where "a \<noteq> 0" "P \<subseteq> {x. a \<bullet> x = 0}"
using lowdim_subset_hyperplane [of P] P span_base by auto
then obtain S where S: "S \<in> lmeasurable"
and subS: "{z. norm z \<le> m \<and> (\<exists>t \<in> P. norm(z - t) \<le> e)} \<subseteq> S"
and mS: "measure lebesgue S \<le> (2 * e) * (2 * m) ^ (CARD('n) - 1)"
by (rule Sard_lemma0 [OF _ _ \<open>0 \<le> m\<close> \<open>0 \<le> e\<close>])
show thesis
proof
show "(+)w ` S \<in> lmeasurable"
by (metis measurable_translation S)
show "{z. norm (z - w) \<le> m \<and> (\<exists>t\<in>P. norm (z - w - t) \<le> e)} \<subseteq> (+)w ` S"
using subS by force
show "measure lebesgue ((+)w ` S) \<le> 2 * e * (2 * m) ^ (CARD('n) - 1)"
by (metis measure_translation mS)
qed
qed
lemma Sard_lemma2:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n::{finite,wellorder}"
assumes mlen: "CARD('m) \<le> CARD('n)" (is "?m \<le> ?n")
and "B > 0" "bounded S"
and derS: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
and rank: "\<And>x. x \<in> S \<Longrightarrow> rank(matrix(f' x)) < CARD('n)"
and B: "\<And>x. x \<in> S \<Longrightarrow> onorm(f' x) \<le> B"
shows "negligible(f ` S)"
proof -
have lin_f': "\<And>x. x \<in> S \<Longrightarrow> linear(f' x)"
using derS has_derivative_linear by blast
show ?thesis
proof (clarsimp simp add: negligible_outer_le)
fix e :: "real"
assume "e > 0"
obtain c where csub: "S \<subseteq> cbox (- (vec c)) (vec c)" and "c > 0"
proof -
obtain b where b: "\<And>x. x \<in> S \<Longrightarrow> norm x \<le> b"
using \<open>bounded S\<close> by (auto simp: bounded_iff)
show thesis
proof
have "- \<bar>b\<bar> - 1 \<le> x $ i \<and> x $ i \<le> \<bar>b\<bar> + 1" if "x \<in> S" for x i
using component_le_norm_cart [of x i] b [OF that] by auto
then show "S \<subseteq> cbox (- vec (\<bar>b\<bar> + 1)) (vec (\<bar>b\<bar> + 1))"
by (auto simp: mem_box_cart)
qed auto
qed
then have box_cc: "box (- (vec c)) (vec c) \<noteq> {}" and cbox_cc: "cbox (- (vec c)) (vec c) \<noteq> {}"
by (auto simp: interval_eq_empty_cart)
obtain d where "d > 0" "d \<le> B"
and d: "(d * 2) * (4 * B) ^ (?n - 1) \<le> e / (2*c) ^ ?m / ?m ^ ?m"
apply (rule that [of "min B (e / (2*c) ^ ?m / ?m ^ ?m / (4 * B) ^ (?n - 1) / 2)"])
using \<open>B > 0\<close> \<open>c > 0\<close> \<open>e > 0\<close>
by (simp_all add: divide_simps min_mult_distrib_right)
have "\<exists>r. 0 < r \<and> r \<le> 1/2 \<and>
(x \<in> S
\<longrightarrow> (\<forall>y. y \<in> S \<and> norm(y - x) < r
\<longrightarrow> norm(f y - f x - f' x (y - x)) \<le> d * norm(y - x)))" for x
proof (cases "x \<in> S")
case True
then obtain r where "r > 0"
and "\<And>y. \<lbrakk>y \<in> S; norm (y - x) < r\<rbrakk>
\<Longrightarrow> norm (f y - f x - f' x (y - x)) \<le> d * norm (y - x)"
using derS \<open>d > 0\<close> by (force simp: has_derivative_within_alt)
then show ?thesis
by (rule_tac x="min r (1/2)" in exI) simp
next
case False
then show ?thesis
by (rule_tac x="1/2" in exI) simp
qed
then obtain r where r12: "\<And>x. 0 < r x \<and> r x \<le> 1/2"
and r: "\<And>x y. \<lbrakk>x \<in> S; y \<in> S; norm(y - x) < r x\<rbrakk>
\<Longrightarrow> norm(f y - f x - f' x (y - x)) \<le> d * norm(y - x)"
by metis
then have ga: "gauge (\<lambda>x. ball x (r x))"
by (auto simp: gauge_def)
obtain \<D> where \<D>: "countable \<D>" and sub_cc: "\<Union>\<D> \<subseteq> cbox (- vec c) (vec c)"
and cbox: "\<And>K. K \<in> \<D> \<Longrightarrow> interior K \<noteq> {} \<and> (\<exists>u v. K = cbox u v)"
and djointish: "pairwise (\<lambda>A B. interior A \<inter> interior B = {}) \<D>"
and covered: "\<And>K. K \<in> \<D> \<Longrightarrow> \<exists>x \<in> S \<inter> K. K \<subseteq> ball x (r x)"
and close: "\<And>u v. cbox u v \<in> \<D> \<Longrightarrow> \<exists>n. \<forall>i::'m. v $ i - u $ i = 2*c / 2^n"
and covers: "S \<subseteq> \<Union>\<D>"
apply (rule covering_lemma [OF csub box_cc ga])
apply (auto simp: Basis_vec_def cart_eq_inner_axis [symmetric])
done
let ?\<mu> = "measure lebesgue"
have "\<exists>T. T \<in> lmeasurable \<and> f ` (K \<inter> S) \<subseteq> T \<and> ?\<mu> T \<le> e / (2*c) ^ ?m * ?\<mu> K"
if "K \<in> \<D>" for K
proof -
obtain u v where uv: "K = cbox u v"
using cbox \<open>K \<in> \<D>\<close> by blast
then have uv_ne: "cbox u v \<noteq> {}"
using cbox that by fastforce
obtain x where x: "x \<in> S \<inter> cbox u v" "cbox u v \<subseteq> ball x (r x)"
using \<open>K \<in> \<D>\<close> covered uv by blast
then have "dim (range (f' x)) < ?n"
using rank_dim_range [of "matrix (f' x)"] x rank[of x]
by (auto simp: matrix_works scalar_mult_eq_scaleR lin_f')
then obtain T where T: "T \<in> lmeasurable"
and subT: "{z. norm(z - f x) \<le> (2 * B) * norm(v - u) \<and> (\<exists>t \<in> range (f' x). norm(z - f x - t) \<le> d * norm(v - u))} \<subseteq> T"
and measT: "?\<mu> T \<le> (2 * (d * norm(v - u))) * (2 * ((2 * B) * norm(v - u))) ^ (?n - 1)"
(is "_ \<le> ?DVU")
apply (rule Sard_lemma1 [of "range (f' x)" "(2 * B) * norm(v - u)" "d * norm(v - u)" "f x"])
using \<open>B > 0\<close> \<open>d > 0\<close> by simp_all
show ?thesis
proof (intro exI conjI)
have "f ` (K \<inter> S) \<subseteq> {z. norm(z - f x) \<le> (2 * B) * norm(v - u) \<and> (\<exists>t \<in> range (f' x). norm(z - f x - t) \<le> d * norm(v - u))}"
unfolding uv
proof (clarsimp simp: mult.assoc, intro conjI)
fix y
assume y: "y \<in> cbox u v" and "y \<in> S"
then have "norm (y - x) < r x"
by (metis dist_norm mem_ball norm_minus_commute subsetCE x(2))
then have le_dyx: "norm (f y - f x - f' x (y - x)) \<le> d * norm (y - x)"
using r [of x y] x \<open>y \<in> S\<close> by blast
have yx_le: "norm (y - x) \<le> norm (v - u)"
proof (rule norm_le_componentwise_cart)
show "norm ((y - x) $ i) \<le> norm ((v - u) $ i)" for i
using x y by (force simp: mem_box_cart dest!: spec [where x=i])
qed
have *: "\<lbrakk>norm(y - x - z) \<le> d; norm z \<le> B; d \<le> B\<rbrakk> \<Longrightarrow> norm(y - x) \<le> 2 * B"
for x y z :: "real^'n::_" and d B
using norm_triangle_ineq2 [of "y - x" z] by auto
show "norm (f y - f x) \<le> 2 * (B * norm (v - u))"
proof (rule * [OF le_dyx])
have "norm (f' x (y - x)) \<le> onorm (f' x) * norm (y - x)"
using onorm [of "f' x" "y-x"] by (meson IntE lin_f' linear_linear x(1))
also have "\<dots> \<le> B * norm (v - u)"
proof (rule mult_mono)
show "onorm (f' x) \<le> B"
using B x by blast
qed (use \<open>B > 0\<close> yx_le in auto)
finally show "norm (f' x (y - x)) \<le> B * norm (v - u)" .
show "d * norm (y - x) \<le> B * norm (v - u)"
using \<open>B > 0\<close> by (auto intro: mult_mono [OF \<open>d \<le> B\<close> yx_le])
qed
show "\<exists>t. norm (f y - f x - f' x t) \<le> d * norm (v - u)"
apply (rule_tac x="y-x" in exI)
using \<open>d > 0\<close> yx_le le_dyx mult_left_mono [where c=d]
by (meson order_trans real_mult_le_cancel_iff2)
qed
with subT show "f ` (K \<inter> S) \<subseteq> T" by blast
show "?\<mu> T \<le> e / (2*c) ^ ?m * ?\<mu> K"
proof (rule order_trans [OF measT])
have "?DVU = (d * 2 * (4 * B) ^ (?n - 1)) * norm (v - u)^?n"
using \<open>c > 0\<close>
- apply (simp add: algebra_simps power_mult_distrib)
+ apply (simp add: algebra_simps)
by (metis Suc_pred power_Suc zero_less_card_finite)
also have "\<dots> \<le> (e / (2*c) ^ ?m / (?m ^ ?m)) * norm(v - u) ^ ?n"
by (rule mult_right_mono [OF d]) auto
also have "\<dots> \<le> e / (2*c) ^ ?m * ?\<mu> K"
proof -
have "u \<in> ball (x) (r x)" "v \<in> ball x (r x)"
using box_ne_empty(1) contra_subsetD [OF x(2)] mem_box(2) uv_ne by fastforce+
moreover have "r x \<le> 1/2"
using r12 by auto
ultimately have "norm (v - u) \<le> 1"
using norm_triangle_half_r [of x u 1 v]
by (metis (no_types, hide_lams) dist_commute dist_norm less_eq_real_def less_le_trans mem_ball)
then have "norm (v - u) ^ ?n \<le> norm (v - u) ^ ?m"
by (simp add: power_decreasing [OF mlen])
also have "\<dots> \<le> ?\<mu> K * real (?m ^ ?m)"
proof -
obtain n where n: "\<And>i. v$i - u$i = 2 * c / 2^n"
using close [of u v] \<open>K \<in> \<D>\<close> uv by blast
have "norm (v - u) ^ ?m \<le> (\<Sum>i\<in>UNIV. \<bar>(v - u) $ i\<bar>) ^ ?m"
by (intro norm_le_l1_cart power_mono) auto
also have "\<dots> \<le> (\<Prod>i\<in>UNIV. v $ i - u $ i) * real CARD('m) ^ CARD('m)"
by (simp add: n field_simps \<open>c > 0\<close> less_eq_real_def)
also have "\<dots> = ?\<mu> K * real (?m ^ ?m)"
by (simp add: uv uv_ne content_cbox_cart)
finally show ?thesis .
qed
finally have *: "1 / real (?m ^ ?m) * norm (v - u) ^ ?n \<le> ?\<mu> K"
by (simp add: field_split_simps)
show ?thesis
using mult_left_mono [OF *, of "e / (2*c) ^ ?m"] \<open>c > 0\<close> \<open>e > 0\<close> by auto
qed
finally show "?DVU \<le> e / (2*c) ^ ?m * ?\<mu> K" .
qed
qed (use T in auto)
qed
then obtain g where meas_g: "\<And>K. K \<in> \<D> \<Longrightarrow> g K \<in> lmeasurable"
and sub_g: "\<And>K. K \<in> \<D> \<Longrightarrow> f ` (K \<inter> S) \<subseteq> g K"
and le_g: "\<And>K. K \<in> \<D> \<Longrightarrow> ?\<mu> (g K) \<le> e / (2*c)^?m * ?\<mu> K"
by metis
have le_e: "?\<mu> (\<Union>i\<in>\<F>. g i) \<le> e"
if "\<F> \<subseteq> \<D>" "finite \<F>" for \<F>
proof -
have "?\<mu> (\<Union>i\<in>\<F>. g i) \<le> (\<Sum>i\<in>\<F>. ?\<mu> (g i))"
using meas_g \<open>\<F> \<subseteq> \<D>\<close> by (auto intro: measure_UNION_le [OF \<open>finite \<F>\<close>])
also have "\<dots> \<le> (\<Sum>K\<in>\<F>. e / (2*c) ^ ?m * ?\<mu> K)"
using \<open>\<F> \<subseteq> \<D>\<close> sum_mono [OF le_g] by (meson le_g subsetCE sum_mono)
also have "\<dots> = e / (2*c) ^ ?m * (\<Sum>K\<in>\<F>. ?\<mu> K)"
by (simp add: sum_distrib_left)
also have "\<dots> \<le> e"
proof -
have "\<F> division_of \<Union>\<F>"
proof (rule division_ofI)
show "K \<subseteq> \<Union>\<F>" "K \<noteq> {}" "\<exists>a b. K = cbox a b" if "K \<in> \<F>" for K
using \<open>K \<in> \<F>\<close> covered cbox \<open>\<F> \<subseteq> \<D>\<close> by (auto simp: Union_upper)
show "interior K \<inter> interior L = {}" if "K \<in> \<F>" and "L \<in> \<F>" and "K \<noteq> L" for K L
by (metis (mono_tags, lifting) \<open>\<F> \<subseteq> \<D>\<close> pairwiseD djointish pairwise_subset that)
qed (use that in auto)
then have "sum ?\<mu> \<F> \<le> ?\<mu> (\<Union>\<F>)"
by (simp add: content_division)
also have "\<dots> \<le> ?\<mu> (cbox (- vec c) (vec c) :: (real, 'm) vec set)"
proof (rule measure_mono_fmeasurable)
show "\<Union>\<F> \<subseteq> cbox (- vec c) (vec c)"
by (meson Sup_subset_mono sub_cc order_trans \<open>\<F> \<subseteq> \<D>\<close>)
qed (use \<open>\<F> division_of \<Union>\<F>\<close> lmeasurable_division in auto)
also have "\<dots> = content (cbox (- vec c) (vec c) :: (real, 'm) vec set)"
by simp
also have "\<dots> \<le> (2 ^ ?m * c ^ ?m)"
using \<open>c > 0\<close> by (simp add: content_cbox_if_cart)
finally have "sum ?\<mu> \<F> \<le> (2 ^ ?m * c ^ ?m)" .
then show ?thesis
- using \<open>e > 0\<close> \<open>c > 0\<close> by (auto simp: field_split_simps mult_less_0_iff)
+ using \<open>e > 0\<close> \<open>c > 0\<close> by (auto simp: field_split_simps)
qed
finally show ?thesis .
qed
show "\<exists>T. f ` S \<subseteq> T \<and> T \<in> lmeasurable \<and> ?\<mu> T \<le> e"
proof (intro exI conjI)
show "f ` S \<subseteq> \<Union> (g ` \<D>)"
using covers sub_g by force
show "\<Union> (g ` \<D>) \<in> lmeasurable"
by (rule fmeasurable_UN_bound [OF \<open>countable \<D>\<close> meas_g le_e])
show "?\<mu> (\<Union> (g ` \<D>)) \<le> e"
by (rule measure_UN_bound [OF \<open>countable \<D>\<close> meas_g le_e])
qed
qed
qed
theorem baby_Sard:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n::{finite,wellorder}"
assumes mlen: "CARD('m) \<le> CARD('n)"
and der: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
and rank: "\<And>x. x \<in> S \<Longrightarrow> rank(matrix(f' x)) < CARD('n)"
shows "negligible(f ` S)"
proof -
let ?U = "\<lambda>n. {x \<in> S. norm(x) \<le> n \<and> onorm(f' x) \<le> real n}"
have "\<And>x. x \<in> S \<Longrightarrow> \<exists>n. norm x \<le> real n \<and> onorm (f' x) \<le> real n"
by (meson linear order_trans real_arch_simple)
then have eq: "S = (\<Union>n. ?U n)"
by auto
have "negligible (f ` ?U n)" for n
proof (rule Sard_lemma2 [OF mlen])
show "0 < real n + 1"
by auto
show "bounded (?U n)"
using bounded_iff by blast
show "(f has_derivative f' x) (at x within ?U n)" if "x \<in> ?U n" for x
using der that by (force intro: has_derivative_subset)
qed (use rank in auto)
then show ?thesis
by (subst eq) (simp add: image_Union negligible_Union_nat)
qed
subsection\<open>A one-way version of change-of-variables not assuming injectivity. \<close>
lemma integral_on_image_ubound_weak:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real"
assumes S: "S \<in> sets lebesgue"
and f: "f \<in> borel_measurable (lebesgue_on (g ` S))"
and nonneg_fg: "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f(g x)"
and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and det_int_fg: "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) integrable_on S"
and meas_gim: "\<And>T. \<lbrakk>T \<subseteq> g ` S; T \<in> sets lebesgue\<rbrakk> \<Longrightarrow> {x \<in> S. g x \<in> T} \<in> sets lebesgue"
shows "f integrable_on (g ` S) \<and>
integral (g ` S) f \<le> integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x))"
(is "_ \<and> _ \<le> ?b")
proof -
let ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar>"
have cont_g: "continuous_on S g"
using der_g has_derivative_continuous_on by blast
have [simp]: "space (lebesgue_on S) = S"
by (simp add: S)
have gS_in_sets_leb: "g ` S \<in> sets lebesgue"
apply (rule differentiable_image_in_sets_lebesgue)
using der_g by (auto simp: S differentiable_def differentiable_on_def)
obtain h where nonneg_h: "\<And>n x. 0 \<le> h n x"
and h_le_f: "\<And>n x. x \<in> S \<Longrightarrow> h n (g x) \<le> f (g x)"
and h_inc: "\<And>n x. h n x \<le> h (Suc n) x"
and h_meas: "\<And>n. h n \<in> borel_measurable lebesgue"
and fin_R: "\<And>n. finite(range (h n))"
and lim: "\<And>x. x \<in> g ` S \<Longrightarrow> (\<lambda>n. h n x) \<longlonglongrightarrow> f x"
proof -
let ?f = "\<lambda>x. if x \<in> g ` S then f x else 0"
have "?f \<in> borel_measurable lebesgue \<and> (\<forall>x. 0 \<le> ?f x)"
by (auto simp: gS_in_sets_leb f nonneg_fg measurable_restrict_space_iff [symmetric])
then show ?thesis
apply (clarsimp simp add: borel_measurable_simple_function_limit_increasing)
apply (rename_tac h)
by (rule_tac h=h in that) (auto split: if_split_asm)
qed
have h_lmeas: "{t. h n (g t) = y} \<inter> S \<in> sets lebesgue" for y n
proof -
have "space (lebesgue_on (UNIV::(real,'n) vec set)) = UNIV"
by simp
then have "((h n) -`{y} \<inter> g ` S) \<in> sets (lebesgue_on (g ` S))"
by (metis Int_commute borel_measurable_vimage h_meas image_eqI inf_top.right_neutral sets_restrict_space space_borel space_completion space_lborel)
then have "({u. h n u = y} \<inter> g ` S) \<in> sets lebesgue"
using gS_in_sets_leb
by (simp add: integral_indicator fmeasurableI2 sets_restrict_space_iff vimage_def)
then have "{x \<in> S. g x \<in> ({u. h n u = y} \<inter> g ` S)} \<in> sets lebesgue"
using meas_gim[of "({u. h n u = y} \<inter> g ` S)"] by force
moreover have "{t. h n (g t) = y} \<inter> S = {x \<in> S. g x \<in> ({u. h n u = y} \<inter> g ` S)}"
by blast
ultimately show ?thesis
by auto
qed
have hint: "h n integrable_on g ` S \<and> integral (g ` S) (h n) \<le> integral S (\<lambda>x. ?D x * h n (g x))"
(is "?INT \<and> ?lhs \<le> ?rhs") for n
proof -
let ?R = "range (h n)"
have hn_eq: "h n = (\<lambda>x. \<Sum>y\<in>?R. y * indicat_real {x. h n x = y} x)"
by (simp add: indicator_def if_distrib fin_R cong: if_cong)
have yind: "(\<lambda>t. y * indicator{x. h n x = y} t) integrable_on (g ` S) \<and>
(integral (g ` S) (\<lambda>t. y * indicator {x. h n x = y} t))
\<le> integral S (\<lambda>t. \<bar>det (matrix (g' t))\<bar> * y * indicator {x. h n x = y} (g t))"
if y: "y \<in> ?R" for y::real
proof (cases "y=0")
case True
then show ?thesis using gS_in_sets_leb integrable_0 by force
next
case False
with that have "y > 0"
using less_eq_real_def nonneg_h by fastforce
have "(\<lambda>x. if x \<in> {t. h n (g t) = y} then ?D x else 0) integrable_on S"
proof (rule measurable_bounded_by_integrable_imp_integrable)
have "(\<lambda>x. ?D x) \<in> borel_measurable (lebesgue_on ({t. h n (g t) = y} \<inter> S))"
apply (intro borel_measurable_abs borel_measurable_det_Jacobian [OF h_lmeas, where f=g])
by (meson der_g IntD2 has_derivative_within_subset inf_le2)
then have "(\<lambda>x. if x \<in> {t. h n (g t) = y} \<inter> S then ?D x else 0) \<in> borel_measurable lebesgue"
by (rule borel_measurable_if_I [OF _ h_lmeas])
then show "(\<lambda>x. if x \<in> {t. h n (g t) = y} then ?D x else 0) \<in> borel_measurable (lebesgue_on S)"
by (simp add: if_if_eq_conj Int_commute borel_measurable_if [OF S, symmetric])
show "(\<lambda>x. ?D x *\<^sub>R f (g x) /\<^sub>R y) integrable_on S"
by (rule integrable_cmul) (use det_int_fg in auto)
show "norm (if x \<in> {t. h n (g t) = y} then ?D x else 0) \<le> ?D x *\<^sub>R f (g x) /\<^sub>R y"
if "x \<in> S" for x
using nonneg_h [of n x] \<open>y > 0\<close> nonneg_fg [of x] h_le_f [of x n] that
by (auto simp: divide_simps mult_left_mono)
qed (use S in auto)
then have int_det: "(\<lambda>t. \<bar>det (matrix (g' t))\<bar>) integrable_on ({t. h n (g t) = y} \<inter> S)"
using integrable_restrict_Int by force
have "(g ` ({t. h n (g t) = y} \<inter> S)) \<in> lmeasurable"
apply (rule measurable_differentiable_image [OF h_lmeas])
apply (blast intro: has_derivative_within_subset [OF der_g])
apply (rule int_det)
done
moreover have "g ` ({t. h n (g t) = y} \<inter> S) = {x. h n x = y} \<inter> g ` S"
by blast
moreover have "measure lebesgue (g ` ({t. h n (g t) = y} \<inter> S))
\<le> integral ({t. h n (g t) = y} \<inter> S) (\<lambda>t. \<bar>det (matrix (g' t))\<bar>)"
apply (rule measure_differentiable_image [OF h_lmeas _ int_det])
apply (blast intro: has_derivative_within_subset [OF der_g])
done
ultimately show ?thesis
using \<open>y > 0\<close> integral_restrict_Int [of S "{t. h n (g t) = y}" "\<lambda>t. \<bar>det (matrix (g' t))\<bar> * y"]
- apply (simp add: integrable_on_indicator integrable_on_cmult_iff integral_indicator)
+ apply (simp add: integrable_on_indicator integral_indicator)
apply (simp add: indicator_def if_distrib cong: if_cong)
done
qed
have hn_int: "h n integrable_on g ` S"
apply (subst hn_eq)
using yind by (force intro: integrable_sum [OF fin_R])
then show ?thesis
proof
have "?lhs = integral (g ` S) (\<lambda>x. \<Sum>y\<in>range (h n). y * indicat_real {x. h n x = y} x)"
by (metis hn_eq)
also have "\<dots> = (\<Sum>y\<in>range (h n). integral (g ` S) (\<lambda>x. y * indicat_real {x. h n x = y} x))"
by (rule integral_sum [OF fin_R]) (use yind in blast)
also have "\<dots> \<le> (\<Sum>y\<in>range (h n). integral S (\<lambda>u. \<bar>det (matrix (g' u))\<bar> * y * indicat_real {x. h n x = y} (g u)))"
using yind by (force intro: sum_mono)
also have "\<dots> = integral S (\<lambda>u. \<Sum>y\<in>range (h n). \<bar>det (matrix (g' u))\<bar> * y * indicat_real {x. h n x = y} (g u))"
proof (rule integral_sum [OF fin_R, symmetric])
fix y assume y: "y \<in> ?R"
with nonneg_h have "y \<ge> 0"
by auto
show "(\<lambda>u. \<bar>det (matrix (g' u))\<bar> * y * indicat_real {x. h n x = y} (g u)) integrable_on S"
proof (rule measurable_bounded_by_integrable_imp_integrable)
have "(\<lambda>x. indicat_real {x. h n x = y} (g x)) \<in> borel_measurable (lebesgue_on S)"
using h_lmeas S
by (auto simp: indicator_vimage [symmetric] borel_measurable_indicator_iff sets_restrict_space_iff)
then show "(\<lambda>u. \<bar>det (matrix (g' u))\<bar> * y * indicat_real {x. h n x = y} (g u)) \<in> borel_measurable (lebesgue_on S)"
by (intro borel_measurable_times borel_measurable_abs borel_measurable_const borel_measurable_det_Jacobian [OF S der_g])
next
fix x
assume "x \<in> S"
have "y * indicat_real {x. h n x = y} (g x) \<le> f (g x)"
by (metis (full_types) \<open>x \<in> S\<close> h_le_f indicator_def mem_Collect_eq mult.right_neutral mult_zero_right nonneg_fg)
with \<open>y \<ge> 0\<close> show "norm (?D x * y * indicat_real {x. h n x = y} (g x)) \<le> ?D x * f(g x)"
by (simp add: abs_mult mult.assoc mult_left_mono)
qed (use S det_int_fg in auto)
qed
also have "\<dots> = integral S (\<lambda>T. \<bar>det (matrix (g' T))\<bar> *
(\<Sum>y\<in>range (h n). y * indicat_real {x. h n x = y} (g T)))"
by (simp add: sum_distrib_left mult.assoc)
also have "\<dots> = ?rhs"
by (metis hn_eq)
finally show "integral (g ` S) (h n) \<le> ?rhs" .
qed
qed
have le: "integral S (\<lambda>T. \<bar>det (matrix (g' T))\<bar> * h n (g T)) \<le> ?b" for n
proof (rule integral_le)
show "(\<lambda>T. \<bar>det (matrix (g' T))\<bar> * h n (g T)) integrable_on S"
proof (rule measurable_bounded_by_integrable_imp_integrable)
have "(\<lambda>T. \<bar>det (matrix (g' T))\<bar> *\<^sub>R h n (g T)) \<in> borel_measurable (lebesgue_on S)"
proof (intro borel_measurable_scaleR borel_measurable_abs borel_measurable_det_Jacobian \<open>S \<in> sets lebesgue\<close>)
have eq: "{x \<in> S. f x \<le> a} = (\<Union>b \<in> (f ` S) \<inter> atMost a. {x. f x = b} \<inter> S)" for f and a::real
by auto
have "finite ((\<lambda>x. h n (g x)) ` S \<inter> {..a})" for a
by (force intro: finite_subset [OF _ fin_R])
with h_lmeas [of n] show "(\<lambda>x. h n (g x)) \<in> borel_measurable (lebesgue_on S)"
apply (simp add: borel_measurable_vimage_halfspace_component_le \<open>S \<in> sets lebesgue\<close> sets_restrict_space_iff eq)
by (metis (mono_tags) SUP_inf sets.finite_UN)
qed (use der_g in blast)
then show "(\<lambda>T. \<bar>det (matrix (g' T))\<bar> * h n (g T)) \<in> borel_measurable (lebesgue_on S)"
by simp
show "norm (?D x * h n (g x)) \<le> ?D x *\<^sub>R f (g x)"
if "x \<in> S" for x
by (simp add: h_le_f mult_left_mono nonneg_h that)
qed (use S det_int_fg in auto)
show "?D x * h n (g x) \<le> ?D x * f (g x)" if "x \<in> S" for x
by (simp add: \<open>x \<in> S\<close> h_le_f mult_left_mono)
show "(\<lambda>x. ?D x * f (g x)) integrable_on S"
using det_int_fg by blast
qed
have "f integrable_on g ` S \<and> (\<lambda>k. integral (g ` S) (h k)) \<longlonglongrightarrow> integral (g ` S) f"
proof (rule monotone_convergence_increasing)
have "\<bar>integral (g ` S) (h n)\<bar> \<le> integral S (\<lambda>x. ?D x * f (g x))" for n
proof -
have "\<bar>integral (g ` S) (h n)\<bar> = integral (g ` S) (h n)"
using hint by (simp add: integral_nonneg nonneg_h)
also have "\<dots> \<le> integral S (\<lambda>x. ?D x * f (g x))"
using hint le by (meson order_trans)
finally show ?thesis .
qed
then show "bounded (range (\<lambda>k. integral (g ` S) (h k)))"
by (force simp: bounded_iff)
qed (use h_inc lim hint in auto)
moreover have "integral (g ` S) (h n) \<le> integral S (\<lambda>x. ?D x * f (g x))" for n
using hint by (blast intro: le order_trans)
ultimately show ?thesis
by (auto intro: Lim_bounded)
qed
lemma integral_on_image_ubound_nonneg:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real"
assumes nonneg_fg: "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f(g x)"
and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and intS: "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) integrable_on S"
shows "f integrable_on (g ` S) \<and> integral (g ` S) f \<le> integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x))"
(is "_ \<and> _ \<le> ?b")
proof -
let ?D = "\<lambda>x. det (matrix (g' x))"
define S' where "S' \<equiv> {x \<in> S. ?D x * f(g x) \<noteq> 0}"
then have der_gS': "\<And>x. x \<in> S' \<Longrightarrow> (g has_derivative g' x) (at x within S')"
by (metis (mono_tags, lifting) der_g has_derivative_within_subset mem_Collect_eq subset_iff)
have "(\<lambda>x. if x \<in> S then \<bar>?D x\<bar> * f (g x) else 0) integrable_on UNIV"
by (simp add: integrable_restrict_UNIV intS)
then have Df_borel: "(\<lambda>x. if x \<in> S then \<bar>?D x\<bar> * f (g x) else 0) \<in> borel_measurable lebesgue"
using integrable_imp_measurable lebesgue_on_UNIV_eq by force
have S': "S' \<in> sets lebesgue"
proof -
from Df_borel borel_measurable_vimage_open [of _ UNIV]
have "{x. (if x \<in> S then \<bar>?D x\<bar> * f (g x) else 0) \<in> T} \<in> sets lebesgue"
if "open T" for T
using that unfolding lebesgue_on_UNIV_eq
by (fastforce simp add: dest!: spec)
then have "{x. (if x \<in> S then \<bar>?D x\<bar> * f (g x) else 0) \<in> -{0}} \<in> sets lebesgue"
using open_Compl by blast
then show ?thesis
by (simp add: S'_def conj_ac split: if_split_asm cong: conj_cong)
qed
then have gS': "g ` S' \<in> sets lebesgue"
proof (rule differentiable_image_in_sets_lebesgue)
show "g differentiable_on S'"
using der_g unfolding S'_def differentiable_def differentiable_on_def
by (blast intro: has_derivative_within_subset)
qed auto
have f: "f \<in> borel_measurable (lebesgue_on (g ` S'))"
proof (clarsimp simp add: borel_measurable_vimage_open)
fix T :: "real set"
assume "open T"
have "{x \<in> g ` S'. f x \<in> T} = g ` {x \<in> S'. f(g x) \<in> T}"
by blast
moreover have "g ` {x \<in> S'. f(g x) \<in> T} \<in> sets lebesgue"
proof (rule differentiable_image_in_sets_lebesgue)
let ?h = "\<lambda>x. \<bar>?D x\<bar> * f (g x) /\<^sub>R \<bar>?D x\<bar>"
have "(\<lambda>x. if x \<in> S' then \<bar>?D x\<bar> * f (g x) else 0) = (\<lambda>x. if x \<in> S then \<bar>?D x\<bar> * f (g x) else 0)"
by (auto simp: S'_def)
also have "\<dots> \<in> borel_measurable lebesgue"
by (rule Df_borel)
finally have *: "(\<lambda>x. \<bar>?D x\<bar> * f (g x)) \<in> borel_measurable (lebesgue_on S')"
by (simp add: borel_measurable_if_D)
have "?h \<in> borel_measurable (lebesgue_on S')"
by (intro * S' der_gS' borel_measurable_det_Jacobian measurable) (blast intro: der_gS')
moreover have "?h x = f(g x)" if "x \<in> S'" for x
using that by (auto simp: S'_def)
ultimately have "(\<lambda>x. f(g x)) \<in> borel_measurable (lebesgue_on S')"
by (metis (no_types, lifting) measurable_lebesgue_cong)
then show "{x \<in> S'. f (g x) \<in> T} \<in> sets lebesgue"
by (simp add: \<open>S' \<in> sets lebesgue\<close> \<open>open T\<close> borel_measurable_vimage_open sets_restrict_space_iff)
show "g differentiable_on {x \<in> S'. f (g x) \<in> T}"
using der_g unfolding S'_def differentiable_def differentiable_on_def
by (blast intro: has_derivative_within_subset)
qed auto
ultimately have "{x \<in> g ` S'. f x \<in> T} \<in> sets lebesgue"
by metis
then show "{x \<in> g ` S'. f x \<in> T} \<in> sets (lebesgue_on (g ` S'))"
by (simp add: \<open>g ` S' \<in> sets lebesgue\<close> sets_restrict_space_iff)
qed
have intS': "(\<lambda>x. \<bar>?D x\<bar> * f (g x)) integrable_on S'"
using intS
by (rule integrable_spike_set) (auto simp: S'_def intro: empty_imp_negligible)
have lebS': "{x \<in> S'. g x \<in> T} \<in> sets lebesgue" if "T \<subseteq> g ` S'" "T \<in> sets lebesgue" for T
proof -
have "g \<in> borel_measurable (lebesgue_on S')"
using der_gS' has_derivative_continuous_on S'
by (blast intro: continuous_imp_measurable_on_sets_lebesgue)
moreover have "{x \<in> S'. g x \<in> U} \<in> sets lebesgue" if "negligible U" "U \<subseteq> g ` S'" for U
proof (intro negligible_imp_sets negligible_differentiable_vimage that)
fix x
assume x: "x \<in> S'"
then have "linear (g' x)"
using der_gS' has_derivative_linear by blast
with x show "inj (g' x)"
by (auto simp: S'_def det_nz_iff_inj)
qed (use der_gS' in auto)
ultimately show ?thesis
using double_lebesgue_sets [OF S' gS' order_refl] that by blast
qed
have int_gS': "f integrable_on g ` S' \<and> integral (g ` S') f \<le> integral S' (\<lambda>x. \<bar>?D x\<bar> * f(g x))"
using integral_on_image_ubound_weak [OF S' f nonneg_fg der_gS' intS' lebS'] S'_def by blast
have "negligible (g ` {x \<in> S. det(matrix(g' x)) = 0})"
proof (rule baby_Sard, simp_all)
fix x
assume x: "x \<in> S \<and> det (matrix (g' x)) = 0"
then show "(g has_derivative g' x) (at x within {x \<in> S. det (matrix (g' x)) = 0})"
by (metis (no_types, lifting) der_g has_derivative_within_subset mem_Collect_eq subsetI)
then show "rank (matrix (g' x)) < CARD('n)"
using det_nz_iff_inj matrix_vector_mul_linear x
by (fastforce simp add: less_rank_noninjective)
qed
then have negg: "negligible (g ` S - g ` {x \<in> S. ?D x \<noteq> 0})"
by (rule negligible_subset) (auto simp: S'_def)
have null: "g ` {x \<in> S. ?D x \<noteq> 0} - g ` S = {}"
by (auto simp: S'_def)
let ?F = "{x \<in> S. f (g x) \<noteq> 0}"
have eq: "g ` S' = g ` ?F \<inter> g ` {x \<in> S. ?D x \<noteq> 0}"
by (auto simp: S'_def image_iff)
show ?thesis
proof
have "((\<lambda>x. if x \<in> g ` ?F then f x else 0) integrable_on g ` {x \<in> S. ?D x \<noteq> 0})"
using int_gS' eq integrable_restrict_Int [where f=f]
by simp
then have "f integrable_on g ` {x \<in> S. ?D x \<noteq> 0}"
by (auto simp: image_iff elim!: integrable_eq)
then show "f integrable_on g ` S"
apply (rule integrable_spike_set [OF _ empty_imp_negligible negligible_subset])
using negg null by auto
have "integral (g ` S) f = integral (g ` {x \<in> S. ?D x \<noteq> 0}) f"
using negg by (auto intro: negligible_subset integral_spike_set)
also have "\<dots> = integral (g ` {x \<in> S. ?D x \<noteq> 0}) (\<lambda>x. if x \<in> g ` ?F then f x else 0)"
by (auto simp: image_iff intro!: integral_cong)
also have "\<dots> = integral (g ` S') f"
using eq integral_restrict_Int by simp
also have "\<dots> \<le> integral S' (\<lambda>x. \<bar>?D x\<bar> * f(g x))"
by (metis int_gS')
also have "\<dots> \<le> ?b"
by (rule integral_subset_le [OF _ intS' intS]) (use nonneg_fg S'_def in auto)
finally show "integral (g ` S) f \<le> ?b" .
qed
qed
lemma absolutely_integrable_on_image_real:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and intS: "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) absolutely_integrable_on S"
shows "f absolutely_integrable_on (g ` S)"
proof -
let ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f (g x)"
let ?N = "{x \<in> S. f (g x) < 0}" and ?P = "{x \<in> S. f (g x) > 0}"
have eq: "{x. (if x \<in> S then ?D x else 0) > 0} = {x \<in> S. ?D x > 0}"
"{x. (if x \<in> S then ?D x else 0) < 0} = {x \<in> S. ?D x < 0}"
by auto
have "?D integrable_on S"
using intS absolutely_integrable_on_def by blast
then have "(\<lambda>x. if x \<in> S then ?D x else 0) integrable_on UNIV"
by (simp add: integrable_restrict_UNIV)
then have D_borel: "(\<lambda>x. if x \<in> S then ?D x else 0) \<in> borel_measurable (lebesgue_on UNIV)"
using integrable_imp_measurable lebesgue_on_UNIV_eq by blast
then have Dlt: "{x \<in> S. ?D x < 0} \<in> sets lebesgue"
unfolding borel_measurable_vimage_halfspace_component_lt
by (drule_tac x=0 in spec) (auto simp: eq)
from D_borel have Dgt: "{x \<in> S. ?D x > 0} \<in> sets lebesgue"
unfolding borel_measurable_vimage_halfspace_component_gt
by (drule_tac x=0 in spec) (auto simp: eq)
have dfgbm: "?D \<in> borel_measurable (lebesgue_on S)"
using intS absolutely_integrable_on_def integrable_imp_measurable by blast
have der_gN: "(g has_derivative g' x) (at x within ?N)" if "x \<in> ?N" for x
using der_g has_derivative_within_subset that by force
have "(\<lambda>x. - f x) integrable_on g ` ?N \<and>
integral (g ` ?N) (\<lambda>x. - f x) \<le> integral ?N (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * - f (g x))"
proof (rule integral_on_image_ubound_nonneg [OF _ der_gN])
have 1: "?D integrable_on {x \<in> S. ?D x < 0}"
using Dlt
by (auto intro: set_lebesgue_integral_eq_integral [OF set_integrable_subset] intS)
have "uminus \<circ> (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * - f (g x)) integrable_on ?N"
by (simp add: o_def mult_less_0_iff empty_imp_negligible integrable_spike_set [OF 1])
then show "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * - f (g x)) integrable_on ?N"
by (simp add: integrable_neg_iff o_def)
qed auto
then have "f integrable_on g ` ?N"
by (simp add: integrable_neg_iff)
moreover have "g ` ?N = {y \<in> g ` S. f y < 0}"
by auto
ultimately have "f integrable_on {y \<in> g ` S. f y < 0}"
by simp
then have N: "f absolutely_integrable_on {y \<in> g ` S. f y < 0}"
by (rule absolutely_integrable_absolutely_integrable_ubound) auto
have der_gP: "(g has_derivative g' x) (at x within ?P)" if "x \<in> ?P" for x
using der_g has_derivative_within_subset that by force
have "f integrable_on g ` ?P \<and> integral (g ` ?P) f \<le> integral ?P ?D"
proof (rule integral_on_image_ubound_nonneg [OF _ der_gP])
have "?D integrable_on {x \<in> S. 0 < ?D x}"
using Dgt
by (auto intro: set_lebesgue_integral_eq_integral [OF set_integrable_subset] intS)
then show "?D integrable_on ?P"
apply (rule integrable_spike_set)
by (auto simp: zero_less_mult_iff empty_imp_negligible)
qed auto
then have "f integrable_on g ` ?P"
by metis
moreover have "g ` ?P = {y \<in> g ` S. f y > 0}"
by auto
ultimately have "f integrable_on {y \<in> g ` S. f y > 0}"
by simp
then have P: "f absolutely_integrable_on {y \<in> g ` S. f y > 0}"
by (rule absolutely_integrable_absolutely_integrable_lbound) auto
have "(\<lambda>x. if x \<in> g ` S \<and> f x < 0 \<or> x \<in> g ` S \<and> 0 < f x then f x else 0) = (\<lambda>x. if x \<in> g ` S then f x else 0)"
by auto
then show ?thesis
using absolutely_integrable_Un [OF N P] absolutely_integrable_restrict_UNIV [symmetric, where f=f]
by simp
qed
proposition absolutely_integrable_on_image:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and intS: "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S"
shows "f absolutely_integrable_on (g ` S)"
apply (rule absolutely_integrable_componentwise [OF absolutely_integrable_on_image_real [OF der_g]])
using absolutely_integrable_component [OF intS] by auto
proposition integral_on_image_ubound:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
assumes"\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f(g x)"
and "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) integrable_on S"
shows "integral (g ` S) f \<le> integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x))"
using integral_on_image_ubound_nonneg [OF assms] by simp
subsection\<open>Change-of-variables theorem\<close>
text\<open>The classic change-of-variables theorem. We have two versions with quite general hypotheses,
the first that the transforming function has a continuous inverse, the second that the base set is
Lebesgue measurable.\<close>
lemma cov_invertible_nonneg_le:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and der_h: "\<And>y. y \<in> T \<Longrightarrow> (h has_derivative h' y) (at y within T)"
and f0: "\<And>y. y \<in> T \<Longrightarrow> 0 \<le> f y"
and hg: "\<And>x. x \<in> S \<Longrightarrow> g x \<in> T \<and> h(g x) = x"
and gh: "\<And>y. y \<in> T \<Longrightarrow> h y \<in> S \<and> g(h y) = y"
and id: "\<And>y. y \<in> T \<Longrightarrow> h' y \<circ> g'(h y) = id"
shows "f integrable_on T \<and> (integral T f) \<le> b \<longleftrightarrow>
(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) integrable_on S \<and>
integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) \<le> b"
(is "?lhs = ?rhs")
proof -
have Teq: "T = g`S" and Seq: "S = h`T"
using hg gh image_iff by fastforce+
have gS: "g differentiable_on S"
by (meson der_g differentiable_def differentiable_on_def)
let ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f (g x)"
show ?thesis
proof
assume ?lhs
then have fT: "f integrable_on T" and intf: "integral T f \<le> b"
by blast+
show ?rhs
proof
let ?fgh = "\<lambda>x. \<bar>det (matrix (h' x))\<bar> * (\<bar>det (matrix (g' (h x)))\<bar> * f (g (h x)))"
have ddf: "?fgh x = f x"
if "x \<in> T" for x
proof -
have "matrix (h' x) ** matrix (g' (h x)) = mat 1"
using that id[OF that] der_g[of "h x"] gh[OF that] left_inverse_linear has_derivative_linear
by (subst matrix_compose[symmetric]) (force simp: matrix_id_mat_1 has_derivative_linear)+
then have "\<bar>det (matrix (h' x))\<bar> * \<bar>det (matrix (g' (h x)))\<bar> = 1"
by (metis abs_1 abs_mult det_I det_mul)
then show ?thesis
by (simp add: gh that)
qed
have "?D integrable_on (h ` T)"
proof (intro set_lebesgue_integral_eq_integral absolutely_integrable_on_image_real)
show "(\<lambda>x. ?fgh x) absolutely_integrable_on T"
proof (subst absolutely_integrable_on_iff_nonneg)
show "(\<lambda>x. ?fgh x) integrable_on T"
using ddf fT integrable_eq by force
qed (simp add: zero_le_mult_iff f0 gh)
qed (use der_h in auto)
with Seq show "(\<lambda>x. ?D x) integrable_on S"
by simp
have "integral S (\<lambda>x. ?D x) \<le> integral T (\<lambda>x. ?fgh x)"
unfolding Seq
proof (rule integral_on_image_ubound)
show "(\<lambda>x. ?fgh x) integrable_on T"
using ddf fT integrable_eq by force
qed (use f0 gh der_h in auto)
also have "\<dots> = integral T f"
by (force simp: ddf intro: integral_cong)
also have "\<dots> \<le> b"
by (rule intf)
finally show "integral S (\<lambda>x. ?D x) \<le> b" .
qed
next
assume R: ?rhs
then have "f integrable_on g ` S"
using der_g f0 hg integral_on_image_ubound_nonneg by blast
moreover have "integral (g ` S) f \<le> integral S (\<lambda>x. ?D x)"
by (rule integral_on_image_ubound [OF f0 der_g]) (use R Teq in auto)
ultimately show ?lhs
using R by (simp add: Teq)
qed
qed
lemma cov_invertible_nonneg_eq:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
assumes "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and "\<And>y. y \<in> T \<Longrightarrow> (h has_derivative h' y) (at y within T)"
and "\<And>y. y \<in> T \<Longrightarrow> 0 \<le> f y"
and "\<And>x. x \<in> S \<Longrightarrow> g x \<in> T \<and> h(g x) = x"
and "\<And>y. y \<in> T \<Longrightarrow> h y \<in> S \<and> g(h y) = y"
and "\<And>y. y \<in> T \<Longrightarrow> h' y \<circ> g'(h y) = id"
shows "((\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) has_integral b) S \<longleftrightarrow> (f has_integral b) T"
using cov_invertible_nonneg_le [OF assms]
by (simp add: has_integral_iff) (meson eq_iff)
lemma cov_invertible_real:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and der_h: "\<And>y. y \<in> T \<Longrightarrow> (h has_derivative h' y) (at y within T)"
and hg: "\<And>x. x \<in> S \<Longrightarrow> g x \<in> T \<and> h(g x) = x"
and gh: "\<And>y. y \<in> T \<Longrightarrow> h y \<in> S \<and> g(h y) = y"
and id: "\<And>y. y \<in> T \<Longrightarrow> h' y \<circ> g'(h y) = id"
shows "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) absolutely_integrable_on S \<and>
integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) = b \<longleftrightarrow>
f absolutely_integrable_on T \<and> integral T f = b"
(is "?lhs = ?rhs")
proof -
have Teq: "T = g`S" and Seq: "S = h`T"
using hg gh image_iff by fastforce+
let ?DP = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)" and ?DN = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> * -f(g x)"
have "+": "(?DP has_integral b) {x \<in> S. f (g x) > 0} \<longleftrightarrow> (f has_integral b) {y \<in> T. f y > 0}" for b
proof (rule cov_invertible_nonneg_eq)
have *: "(\<lambda>x. f (g x)) -` Y \<inter> {x \<in> S. f (g x) > 0}
= ((\<lambda>x. f (g x)) -` Y \<inter> S) \<inter> {x \<in> S. f (g x) > 0}" for Y
by auto
show "(g has_derivative g' x) (at x within {x \<in> S. f (g x) > 0})" if "x \<in> {x \<in> S. f (g x) > 0}" for x
using that der_g has_derivative_within_subset by fastforce
show "(h has_derivative h' y) (at y within {y \<in> T. f y > 0})" if "y \<in> {y \<in> T. f y > 0}" for y
using that der_h has_derivative_within_subset by fastforce
qed (use gh hg id in auto)
have "-": "(?DN has_integral b) {x \<in> S. f (g x) < 0} \<longleftrightarrow> ((\<lambda>x. - f x) has_integral b) {y \<in> T. f y < 0}" for b
proof (rule cov_invertible_nonneg_eq)
have *: "(\<lambda>x. - f (g x)) -` y \<inter> {x \<in> S. f (g x) < 0}
= ((\<lambda>x. f (g x)) -` uminus ` y \<inter> S) \<inter> {x \<in> S. f (g x) < 0}" for y
using image_iff by fastforce
show "(g has_derivative g' x) (at x within {x \<in> S. f (g x) < 0})" if "x \<in> {x \<in> S. f (g x) < 0}" for x
using that der_g has_derivative_within_subset by fastforce
show "(h has_derivative h' y) (at y within {y \<in> T. f y < 0})" if "y \<in> {y \<in> T. f y < 0}" for y
using that der_h has_derivative_within_subset by fastforce
qed (use gh hg id in auto)
show ?thesis
proof
assume LHS: ?lhs
have eq: "{x. (if x \<in> S then ?DP x else 0) > 0} = {x \<in> S. ?DP x > 0}"
"{x. (if x \<in> S then ?DP x else 0) < 0} = {x \<in> S. ?DP x < 0}"
by auto
have "?DP integrable_on S"
using LHS absolutely_integrable_on_def by blast
then have "(\<lambda>x. if x \<in> S then ?DP x else 0) integrable_on UNIV"
by (simp add: integrable_restrict_UNIV)
then have D_borel: "(\<lambda>x. if x \<in> S then ?DP x else 0) \<in> borel_measurable (lebesgue_on UNIV)"
using integrable_imp_measurable lebesgue_on_UNIV_eq by blast
then have SN: "{x \<in> S. ?DP x < 0} \<in> sets lebesgue"
unfolding borel_measurable_vimage_halfspace_component_lt
by (drule_tac x=0 in spec) (auto simp: eq)
from D_borel have SP: "{x \<in> S. ?DP x > 0} \<in> sets lebesgue"
unfolding borel_measurable_vimage_halfspace_component_gt
by (drule_tac x=0 in spec) (auto simp: eq)
have "?DP absolutely_integrable_on {x \<in> S. ?DP x > 0}"
using LHS by (fast intro!: set_integrable_subset [OF _, of _ S] SP)
then have aP: "?DP absolutely_integrable_on {x \<in> S. f (g x) > 0}"
by (rule absolutely_integrable_spike_set) (auto simp: zero_less_mult_iff empty_imp_negligible)
have "?DP absolutely_integrable_on {x \<in> S. ?DP x < 0}"
using LHS by (fast intro!: set_integrable_subset [OF _, of _ S] SN)
then have aN: "?DP absolutely_integrable_on {x \<in> S. f (g x) < 0}"
by (rule absolutely_integrable_spike_set) (auto simp: mult_less_0_iff empty_imp_negligible)
have fN: "f integrable_on {y \<in> T. f y < 0}"
"integral {y \<in> T. f y < 0} f = integral {x \<in> S. f (g x) < 0} ?DP"
using "-" [of "integral {x \<in> S. f(g x) < 0} ?DN"] aN
by (auto simp: set_lebesgue_integral_eq_integral has_integral_iff integrable_neg_iff)
have faN: "f absolutely_integrable_on {y \<in> T. f y < 0}"
apply (rule absolutely_integrable_integrable_bound [where g = "\<lambda>x. - f x"])
using fN by (auto simp: integrable_neg_iff)
have fP: "f integrable_on {y \<in> T. f y > 0}"
"integral {y \<in> T. f y > 0} f = integral {x \<in> S. f (g x) > 0} ?DP"
using "+" [of "integral {x \<in> S. f(g x) > 0} ?DP"] aP
by (auto simp: set_lebesgue_integral_eq_integral has_integral_iff integrable_neg_iff)
have faP: "f absolutely_integrable_on {y \<in> T. f y > 0}"
apply (rule absolutely_integrable_integrable_bound [where g = f])
using fP by auto
have fa: "f absolutely_integrable_on ({y \<in> T. f y < 0} \<union> {y \<in> T. f y > 0})"
by (rule absolutely_integrable_Un [OF faN faP])
show ?rhs
proof
have eq: "((if x \<in> T \<and> f x < 0 \<or> x \<in> T \<and> 0 < f x then 1 else 0) * f x)
= (if x \<in> T then 1 else 0) * f x" for x
by auto
show "f absolutely_integrable_on T"
using fa by (simp add: indicator_def set_integrable_def eq)
have [simp]: "{y \<in> T. f y < 0} \<inter> {y \<in> T. 0 < f y} = {}" for T and f :: "(real^'n::_) \<Rightarrow> real"
by auto
have "integral T f = integral ({y \<in> T. f y < 0} \<union> {y \<in> T. f y > 0}) f"
by (intro empty_imp_negligible integral_spike_set) (auto simp: eq)
also have "\<dots> = integral {y \<in> T. f y < 0} f + integral {y \<in> T. f y > 0} f"
using fN fP by simp
also have "\<dots> = integral {x \<in> S. f (g x) < 0} ?DP + integral {x \<in> S. 0 < f (g x)} ?DP"
by (simp add: fN fP)
also have "\<dots> = integral ({x \<in> S. f (g x) < 0} \<union> {x \<in> S. 0 < f (g x)}) ?DP"
using aP aN by (simp add: set_lebesgue_integral_eq_integral)
also have "\<dots> = integral S ?DP"
by (intro empty_imp_negligible integral_spike_set) auto
also have "\<dots> = b"
using LHS by simp
finally show "integral T f = b" .
qed
next
assume RHS: ?rhs
have eq: "{x. (if x \<in> T then f x else 0) > 0} = {x \<in> T. f x > 0}"
"{x. (if x \<in> T then f x else 0) < 0} = {x \<in> T. f x < 0}"
by auto
have "f integrable_on T"
using RHS absolutely_integrable_on_def by blast
then have "(\<lambda>x. if x \<in> T then f x else 0) integrable_on UNIV"
by (simp add: integrable_restrict_UNIV)
then have D_borel: "(\<lambda>x. if x \<in> T then f x else 0) \<in> borel_measurable (lebesgue_on UNIV)"
using integrable_imp_measurable lebesgue_on_UNIV_eq by blast
then have TN: "{x \<in> T. f x < 0} \<in> sets lebesgue"
unfolding borel_measurable_vimage_halfspace_component_lt
by (drule_tac x=0 in spec) (auto simp: eq)
from D_borel have TP: "{x \<in> T. f x > 0} \<in> sets lebesgue"
unfolding borel_measurable_vimage_halfspace_component_gt
by (drule_tac x=0 in spec) (auto simp: eq)
have aint: "f absolutely_integrable_on {y. y \<in> T \<and> 0 < (f y)}"
"f absolutely_integrable_on {y. y \<in> T \<and> (f y) < 0}"
and intT: "integral T f = b"
using set_integrable_subset [of _ T] TP TN RHS
by blast+
show ?lhs
proof
have fN: "f integrable_on {v \<in> T. f v < 0}"
using absolutely_integrable_on_def aint by blast
then have DN: "(?DN has_integral integral {y \<in> T. f y < 0} (\<lambda>x. - f x)) {x \<in> S. f (g x) < 0}"
using "-" [of "integral {y \<in> T. f y < 0} (\<lambda>x. - f x)"]
by (simp add: has_integral_neg_iff integrable_integral)
have aDN: "?DP absolutely_integrable_on {x \<in> S. f (g x) < 0}"
apply (rule absolutely_integrable_integrable_bound [where g = ?DN])
using DN hg by (fastforce simp: abs_mult integrable_neg_iff)+
have fP: "f integrable_on {v \<in> T. f v > 0}"
using absolutely_integrable_on_def aint by blast
then have DP: "(?DP has_integral integral {y \<in> T. f y > 0} f) {x \<in> S. f (g x) > 0}"
using "+" [of "integral {y \<in> T. f y > 0} f"]
by (simp add: has_integral_neg_iff integrable_integral)
have aDP: "?DP absolutely_integrable_on {x \<in> S. f (g x) > 0}"
apply (rule absolutely_integrable_integrable_bound [where g = ?DP])
using DP hg by (fastforce simp: integrable_neg_iff)+
have eq: "(if x \<in> S then 1 else 0) * ?DP x = (if x \<in> S \<and> f (g x) < 0 \<or> x \<in> S \<and> f (g x) > 0 then 1 else 0) * ?DP x" for x
by force
have "?DP absolutely_integrable_on ({x \<in> S. f (g x) < 0} \<union> {x \<in> S. f (g x) > 0})"
by (rule absolutely_integrable_Un [OF aDN aDP])
then show I: "?DP absolutely_integrable_on S"
by (simp add: indicator_def eq set_integrable_def)
have [simp]: "{y \<in> S. f y < 0} \<inter> {y \<in> S. 0 < f y} = {}" for S and f :: "(real^'n::_) \<Rightarrow> real"
by auto
have "integral S ?DP = integral ({x \<in> S. f (g x) < 0} \<union> {x \<in> S. f (g x) > 0}) ?DP"
by (intro empty_imp_negligible integral_spike_set) auto
also have "\<dots> = integral {x \<in> S. f (g x) < 0} ?DP + integral {x \<in> S. 0 < f (g x)} ?DP"
using aDN aDP by (simp add: set_lebesgue_integral_eq_integral)
also have "\<dots> = - integral {y \<in> T. f y < 0} (\<lambda>x. - f x) + integral {y \<in> T. f y > 0} f"
using DN DP by (auto simp: has_integral_iff)
also have "\<dots> = integral ({x \<in> T. f x < 0} \<union> {x \<in> T. 0 < f x}) f"
by (simp add: fN fP)
also have "\<dots> = integral T f"
by (intro empty_imp_negligible integral_spike_set) auto
also have "\<dots> = b"
using intT by simp
finally show "integral S ?DP = b" .
qed
qed
qed
lemma cv_inv_version3:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and der_h: "\<And>y. y \<in> T \<Longrightarrow> (h has_derivative h' y) (at y within T)"
and hg: "\<And>x. x \<in> S \<Longrightarrow> g x \<in> T \<and> h(g x) = x"
and gh: "\<And>y. y \<in> T \<Longrightarrow> h y \<in> S \<and> g(h y) = y"
and id: "\<And>y. y \<in> T \<Longrightarrow> h' y \<circ> g'(h y) = id"
shows "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and>
integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) = b
\<longleftrightarrow> f absolutely_integrable_on T \<and> integral T f = b"
proof -
let ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)"
have "((\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x) $ i) absolutely_integrable_on S \<and> integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * (f(g x) $ i)) = b $ i) \<longleftrightarrow>
((\<lambda>x. f x $ i) absolutely_integrable_on T \<and> integral T (\<lambda>x. f x $ i) = b $ i)" for i
by (rule cov_invertible_real [OF der_g der_h hg gh id])
then have "?D absolutely_integrable_on S \<and> (?D has_integral b) S \<longleftrightarrow>
f absolutely_integrable_on T \<and> (f has_integral b) T"
unfolding absolutely_integrable_componentwise_iff [where f=f] has_integral_componentwise_iff [of f]
absolutely_integrable_componentwise_iff [where f="?D"] has_integral_componentwise_iff [of ?D]
by (auto simp: all_conj_distrib Basis_vec_def cart_eq_inner_axis [symmetric]
has_integral_iff set_lebesgue_integral_eq_integral)
then show ?thesis
using absolutely_integrable_on_def by blast
qed
lemma cv_inv_version4:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S) \<and> invertible(matrix(g' x))"
and hg: "\<And>x. x \<in> S \<Longrightarrow> continuous_on (g ` S) h \<and> h(g x) = x"
shows "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and>
integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) = b
\<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b"
proof -
have "\<forall>x. \<exists>h'. x \<in> S
\<longrightarrow> (g has_derivative g' x) (at x within S) \<and> linear h' \<and> g' x \<circ> h' = id \<and> h' \<circ> g' x = id"
using der_g matrix_invertible has_derivative_linear by blast
then obtain h' where h':
"\<And>x. x \<in> S
\<Longrightarrow> (g has_derivative g' x) (at x within S) \<and>
linear (h' x) \<and> g' x \<circ> (h' x) = id \<and> (h' x) \<circ> g' x = id"
by metis
show ?thesis
proof (rule cv_inv_version3)
show "\<And>y. y \<in> g ` S \<Longrightarrow> (h has_derivative h' (h y)) (at y within g ` S)"
using h' hg
by (force simp: continuous_on_eq_continuous_within intro!: has_derivative_inverse_within)
qed (use h' hg in auto)
qed
theorem has_absolute_integral_change_of_variables_invertible:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and hg: "\<And>x. x \<in> S \<Longrightarrow> h(g x) = x"
and conth: "continuous_on (g ` S) h"
shows "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and> integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) = b \<longleftrightarrow>
f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b"
(is "?lhs = ?rhs")
proof -
let ?S = "{x \<in> S. invertible (matrix (g' x))}" and ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)"
have *: "?D absolutely_integrable_on ?S \<and> integral ?S ?D = b
\<longleftrightarrow> f absolutely_integrable_on (g ` ?S) \<and> integral (g ` ?S) f = b"
proof (rule cv_inv_version4)
show "(g has_derivative g' x) (at x within ?S) \<and> invertible (matrix (g' x))"
if "x \<in> ?S" for x
using der_g that has_derivative_within_subset that by fastforce
show "continuous_on (g ` ?S) h \<and> h (g x) = x"
if "x \<in> ?S" for x
using that continuous_on_subset [OF conth] by (simp add: hg image_mono)
qed
have "(g has_derivative g' x) (at x within {x \<in> S. rank (matrix (g' x)) < CARD('m)})" if "x \<in> S" for x
by (metis (no_types, lifting) der_g has_derivative_within_subset mem_Collect_eq subsetI that)
then have "negligible (g ` {x \<in> S. \<not> invertible (matrix (g' x))})"
by (auto simp: invertible_det_nz det_eq_0_rank intro: baby_Sard)
then have neg: "negligible {x \<in> g ` S. x \<notin> g ` ?S \<and> f x \<noteq> 0}"
by (auto intro: negligible_subset)
have [simp]: "{x \<in> g ` ?S. x \<notin> g ` S \<and> f x \<noteq> 0} = {}"
by auto
have "?D absolutely_integrable_on ?S \<and> integral ?S ?D = b
\<longleftrightarrow> ?D absolutely_integrable_on S \<and> integral S ?D = b"
apply (intro conj_cong absolutely_integrable_spike_set_eq)
apply(auto simp: integral_spike_set invertible_det_nz empty_imp_negligible neg)
done
moreover
have "f absolutely_integrable_on (g ` ?S) \<and> integral (g ` ?S) f = b
\<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b"
by (auto intro!: conj_cong absolutely_integrable_spike_set_eq integral_spike_set neg)
ultimately
show ?thesis
using "*" by blast
qed
theorem has_absolute_integral_change_of_variables_compact:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes "compact S"
and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and inj: "inj_on g S"
shows "((\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and>
integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) = b
\<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b)"
proof -
obtain h where hg: "\<And>x. x \<in> S \<Longrightarrow> h(g x) = x"
using inj by (metis the_inv_into_f_f)
have conth: "continuous_on (g ` S) h"
by (metis \<open>compact S\<close> continuous_on_inv der_g has_derivative_continuous_on hg)
show ?thesis
by (rule has_absolute_integral_change_of_variables_invertible [OF der_g hg conth])
qed
lemma has_absolute_integral_change_of_variables_compact_family:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes compact: "\<And>n::nat. compact (F n)"
and der_g: "\<And>x. x \<in> (\<Union>n. F n) \<Longrightarrow> (g has_derivative g' x) (at x within (\<Union>n. F n))"
and inj: "inj_on g (\<Union>n. F n)"
shows "((\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on (\<Union>n. F n) \<and>
integral (\<Union>n. F n) (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) = b
\<longleftrightarrow> f absolutely_integrable_on (g ` (\<Union>n. F n)) \<and> integral (g ` (\<Union>n. F n)) f = b)"
proof -
let ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f (g x)"
let ?U = "\<lambda>n. \<Union>m\<le>n. F m"
let ?lift = "vec::real\<Rightarrow>real^1"
have F_leb: "F m \<in> sets lebesgue" for m
by (simp add: compact borel_compact)
have iff: "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f (g x)) absolutely_integrable_on (?U n) \<and>
integral (?U n) (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f (g x)) = b
\<longleftrightarrow> f absolutely_integrable_on (g ` (?U n)) \<and> integral (g ` (?U n)) f = b" for n b and f :: "real^'m::_ \<Rightarrow> real^'k"
proof (rule has_absolute_integral_change_of_variables_compact)
show "compact (?U n)"
by (simp add: compact compact_UN)
show "(g has_derivative g' x) (at x within (?U n))"
if "x \<in> ?U n" for x
using that by (blast intro!: has_derivative_within_subset [OF der_g])
show "inj_on g (?U n)"
using inj by (auto simp: inj_on_def)
qed
show ?thesis
unfolding image_UN
proof safe
assume DS: "?D absolutely_integrable_on (\<Union>n. F n)"
and b: "b = integral (\<Union>n. F n) ?D"
have DU: "\<And>n. ?D absolutely_integrable_on (?U n)"
"(\<lambda>n. integral (?U n) ?D) \<longlonglongrightarrow> integral (\<Union>n. F n) ?D"
using integral_countable_UN [OF DS F_leb] by auto
with iff have fag: "f absolutely_integrable_on g ` (?U n)"
and fg_int: "integral (\<Union>m\<le>n. g ` F m) f = integral (?U n) ?D" for n
by (auto simp: image_UN)
let ?h = "\<lambda>x. if x \<in> (\<Union>m. g ` F m) then norm(f x) else 0"
have "(\<lambda>x. if x \<in> (\<Union>m. g ` F m) then f x else 0) absolutely_integrable_on UNIV"
proof (rule dominated_convergence_absolutely_integrable)
show "(\<lambda>x. if x \<in> (\<Union>m\<le>k. g ` F m) then f x else 0) absolutely_integrable_on UNIV" for k
unfolding absolutely_integrable_restrict_UNIV
using fag by (simp add: image_UN)
let ?nf = "\<lambda>n x. if x \<in> (\<Union>m\<le>n. g ` F m) then norm(f x) else 0"
show "?h integrable_on UNIV"
proof (rule monotone_convergence_increasing [THEN conjunct1])
show "?nf k integrable_on UNIV" for k
using fag
unfolding integrable_restrict_UNIV absolutely_integrable_on_def by (simp add: image_UN)
{ fix n
have "(norm \<circ> ?D) absolutely_integrable_on ?U n"
by (intro absolutely_integrable_norm DU)
then have "integral (g ` ?U n) (norm \<circ> f) = integral (?U n) (norm \<circ> ?D)"
using iff [of n "vec \<circ> norm \<circ> f" "integral (?U n) (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R (?lift \<circ> norm \<circ> f) (g x))"]
unfolding absolutely_integrable_on_1_iff integral_on_1_eq by (auto simp: o_def)
}
moreover have "bounded (range (\<lambda>k. integral (?U k) (norm \<circ> ?D)))"
unfolding bounded_iff
proof (rule exI, clarify)
fix k
show "norm (integral (?U k) (norm \<circ> ?D)) \<le> integral (\<Union>n. F n) (norm \<circ> ?D)"
unfolding integral_restrict_UNIV [of _ "norm \<circ> ?D", symmetric]
proof (rule integral_norm_bound_integral)
show "(\<lambda>x. if x \<in> \<Union> (F ` {..k}) then (norm \<circ> ?D) x else 0) integrable_on UNIV"
"(\<lambda>x. if x \<in> (\<Union>n. F n) then (norm \<circ> ?D) x else 0) integrable_on UNIV"
using DU(1) DS
unfolding absolutely_integrable_on_def o_def integrable_restrict_UNIV by auto
qed auto
qed
ultimately show "bounded (range (\<lambda>k. integral UNIV (?nf k)))"
by (simp add: integral_restrict_UNIV image_UN [symmetric] o_def)
next
show "(\<lambda>k. if x \<in> (\<Union>m\<le>k. g ` F m) then norm (f x) else 0)
\<longlonglongrightarrow> (if x \<in> (\<Union>m. g ` F m) then norm (f x) else 0)" for x
by (force intro: tendsto_eventually eventually_sequentiallyI)
qed auto
next
show "(\<lambda>k. if x \<in> (\<Union>m\<le>k. g ` F m) then f x else 0)
\<longlonglongrightarrow> (if x \<in> (\<Union>m. g ` F m) then f x else 0)" for x
proof clarsimp
fix m y
assume "y \<in> F m"
show "(\<lambda>k. if \<exists>x\<in>{..k}. g y \<in> g ` F x then f (g y) else 0) \<longlonglongrightarrow> f (g y)"
using \<open>y \<in> F m\<close> by (force intro: tendsto_eventually eventually_sequentiallyI [of m])
qed
qed auto
then show fai: "f absolutely_integrable_on (\<Union>m. g ` F m)"
using absolutely_integrable_restrict_UNIV by blast
show "integral ((\<Union>x. g ` F x)) f = integral (\<Union>n. F n) ?D"
proof (rule LIMSEQ_unique)
show "(\<lambda>n. integral (?U n) ?D) \<longlonglongrightarrow> integral (\<Union>x. g ` F x) f"
unfolding fg_int [symmetric]
proof (rule integral_countable_UN [OF fai])
show "g ` F m \<in> sets lebesgue" for m
proof (rule differentiable_image_in_sets_lebesgue [OF F_leb])
show "g differentiable_on F m"
by (meson der_g differentiableI UnionI differentiable_on_def differentiable_on_subset rangeI subsetI)
qed auto
qed
next
show "(\<lambda>n. integral (?U n) ?D) \<longlonglongrightarrow> integral (\<Union>n. F n) ?D"
by (rule DU)
qed
next
assume fs: "f absolutely_integrable_on (\<Union>x. g ` F x)"
and b: "b = integral ((\<Union>x. g ` F x)) f"
have gF_leb: "g ` F m \<in> sets lebesgue" for m
proof (rule differentiable_image_in_sets_lebesgue [OF F_leb])
show "g differentiable_on F m"
using der_g unfolding differentiable_def differentiable_on_def
by (meson Sup_upper UNIV_I UnionI has_derivative_within_subset image_eqI)
qed auto
have fgU: "\<And>n. f absolutely_integrable_on (\<Union>m\<le>n. g ` F m)"
"(\<lambda>n. integral (\<Union>m\<le>n. g ` F m) f) \<longlonglongrightarrow> integral (\<Union>m. g ` F m) f"
using integral_countable_UN [OF fs gF_leb] by auto
with iff have DUn: "?D absolutely_integrable_on ?U n"
and D_int: "integral (?U n) ?D = integral (\<Union>m\<le>n. g ` F m) f" for n
by (auto simp: image_UN)
let ?h = "\<lambda>x. if x \<in> (\<Union>n. F n) then norm(?D x) else 0"
have "(\<lambda>x. if x \<in> (\<Union>n. F n) then ?D x else 0) absolutely_integrable_on UNIV"
proof (rule dominated_convergence_absolutely_integrable)
show "(\<lambda>x. if x \<in> ?U k then ?D x else 0) absolutely_integrable_on UNIV" for k
unfolding absolutely_integrable_restrict_UNIV using DUn by simp
let ?nD = "\<lambda>n x. if x \<in> ?U n then norm(?D x) else 0"
show "?h integrable_on UNIV"
proof (rule monotone_convergence_increasing [THEN conjunct1])
show "?nD k integrable_on UNIV" for k
using DUn
unfolding integrable_restrict_UNIV absolutely_integrable_on_def by (simp add: image_UN)
{ fix n::nat
have "(norm \<circ> f) absolutely_integrable_on (\<Union>m\<le>n. g ` F m)"
apply (rule absolutely_integrable_norm)
using fgU by blast
then have "integral (?U n) (norm \<circ> ?D) = integral (g ` ?U n) (norm \<circ> f)"
using iff [of n "?lift \<circ> norm \<circ> f" "integral (g ` ?U n) (?lift \<circ> norm \<circ> f)"]
unfolding absolutely_integrable_on_1_iff integral_on_1_eq image_UN by (auto simp: o_def)
}
moreover have "bounded (range (\<lambda>k. integral (g ` ?U k) (norm \<circ> f)))"
unfolding bounded_iff
proof (rule exI, clarify)
fix k
show "norm (integral (g ` ?U k) (norm \<circ> f)) \<le> integral (g ` (\<Union>n. F n)) (norm \<circ> f)"
unfolding integral_restrict_UNIV [of _ "norm \<circ> f", symmetric]
proof (rule integral_norm_bound_integral)
show "(\<lambda>x. if x \<in> g ` ?U k then (norm \<circ> f) x else 0) integrable_on UNIV"
"(\<lambda>x. if x \<in> g ` (\<Union>n. F n) then (norm \<circ> f) x else 0) integrable_on UNIV"
using fgU fs
unfolding absolutely_integrable_on_def o_def integrable_restrict_UNIV
by (auto simp: image_UN)
qed auto
qed
ultimately show "bounded (range (\<lambda>k. integral UNIV (?nD k)))"
unfolding integral_restrict_UNIV image_UN [symmetric] o_def by simp
next
show "(\<lambda>k. if x \<in> ?U k then norm (?D x) else 0) \<longlonglongrightarrow> (if x \<in> (\<Union>n. F n) then norm (?D x) else 0)" for x
by (force intro: tendsto_eventually eventually_sequentiallyI)
qed auto
next
show "(\<lambda>k. if x \<in> ?U k then ?D x else 0) \<longlonglongrightarrow> (if x \<in> (\<Union>n. F n) then ?D x else 0)" for x
proof clarsimp
fix n
assume "x \<in> F n"
show "(\<lambda>m. if \<exists>j\<in>{..m}. x \<in> F j then ?D x else 0) \<longlonglongrightarrow> ?D x"
using \<open>x \<in> F n\<close> by (auto intro!: tendsto_eventually eventually_sequentiallyI [of n])
qed
qed auto
then show Dai: "?D absolutely_integrable_on (\<Union>n. F n)"
unfolding absolutely_integrable_restrict_UNIV by simp
show "integral (\<Union>n. F n) ?D = integral ((\<Union>x. g ` F x)) f"
proof (rule LIMSEQ_unique)
show "(\<lambda>n. integral (\<Union>m\<le>n. g ` F m) f) \<longlonglongrightarrow> integral (\<Union>x. g ` F x) f"
by (rule fgU)
show "(\<lambda>n. integral (\<Union>m\<le>n. g ` F m) f) \<longlonglongrightarrow> integral (\<Union>n. F n) ?D"
unfolding D_int [symmetric] by (rule integral_countable_UN [OF Dai F_leb])
qed
qed
qed
theorem has_absolute_integral_change_of_variables:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes S: "S \<in> sets lebesgue"
and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and inj: "inj_on g S"
shows "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and>
integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) = b
\<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b"
proof -
obtain C N where "fsigma C" and N: "N \<in> null_sets lebesgue" and CNS: "C \<union> N = S" and "disjnt C N"
using lebesgue_set_almost_fsigma [OF S] .
then obtain F :: "nat \<Rightarrow> (real^'m::_) set"
where F: "range F \<subseteq> Collect compact" and Ceq: "C = Union(range F)"
using fsigma_Union_compact by metis
have "negligible N"
using N by (simp add: negligible_iff_null_sets)
let ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f (g x)"
have "?D absolutely_integrable_on C \<and> integral C ?D = b
\<longleftrightarrow> f absolutely_integrable_on (g ` C) \<and> integral (g ` C) f = b"
unfolding Ceq
proof (rule has_absolute_integral_change_of_variables_compact_family)
fix n x
assume "x \<in> \<Union>(F ` UNIV)"
then show "(g has_derivative g' x) (at x within \<Union>(F ` UNIV))"
using Ceq \<open>C \<union> N = S\<close> der_g has_derivative_within_subset by blast
next
have "\<Union>(F ` UNIV) \<subseteq> S"
using Ceq \<open>C \<union> N = S\<close> by blast
then show "inj_on g (\<Union>(F ` UNIV))"
using inj by (meson inj_on_subset)
qed (use F in auto)
moreover
have "?D absolutely_integrable_on C \<and> integral C ?D = b
\<longleftrightarrow> ?D absolutely_integrable_on S \<and> integral S ?D = b"
proof (rule conj_cong)
have neg: "negligible {x \<in> C - S. ?D x \<noteq> 0}" "negligible {x \<in> S - C. ?D x \<noteq> 0}"
using CNS by (blast intro: negligible_subset [OF \<open>negligible N\<close>])+
then show "(?D absolutely_integrable_on C) = (?D absolutely_integrable_on S)"
by (rule absolutely_integrable_spike_set_eq)
show "(integral C ?D = b) \<longleftrightarrow> (integral S ?D = b)"
using integral_spike_set [OF neg] by simp
qed
moreover
have "f absolutely_integrable_on (g ` C) \<and> integral (g ` C) f = b
\<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b"
proof (rule conj_cong)
have "g differentiable_on N"
by (metis CNS der_g differentiable_def differentiable_on_def differentiable_on_subset sup.cobounded2)
with \<open>negligible N\<close>
have neg_gN: "negligible (g ` N)"
by (blast intro: negligible_differentiable_image_negligible)
have neg: "negligible {x \<in> g ` C - g ` S. f x \<noteq> 0}"
"negligible {x \<in> g ` S - g ` C. f x \<noteq> 0}"
using CNS by (blast intro: negligible_subset [OF neg_gN])+
then show "(f absolutely_integrable_on g ` C) = (f absolutely_integrable_on g ` S)"
by (rule absolutely_integrable_spike_set_eq)
show "(integral (g ` C) f = b) \<longleftrightarrow> (integral (g ` S) f = b)"
using integral_spike_set [OF neg] by simp
qed
ultimately show ?thesis
by simp
qed
corollary absolutely_integrable_change_of_variables:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes "S \<in> sets lebesgue"
and "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and "inj_on g S"
shows "f absolutely_integrable_on (g ` S)
\<longleftrightarrow> (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S"
using assms has_absolute_integral_change_of_variables by blast
corollary integral_change_of_variables:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes S: "S \<in> sets lebesgue"
and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and inj: "inj_on g S"
and disj: "(f absolutely_integrable_on (g ` S) \<or>
(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S)"
shows "integral (g ` S) f = integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x))"
using has_absolute_integral_change_of_variables [OF S der_g inj] disj
by blast
lemma has_absolute_integral_change_of_variables_1:
fixes f :: "real \<Rightarrow> real^'n::{finite,wellorder}" and g :: "real \<Rightarrow> real"
assumes S: "S \<in> sets lebesgue"
and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_vector_derivative g' x) (at x within S)"
and inj: "inj_on g S"
shows "(\<lambda>x. \<bar>g' x\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and>
integral S (\<lambda>x. \<bar>g' x\<bar> *\<^sub>R f(g x)) = b
\<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b"
proof -
let ?lift = "vec :: real \<Rightarrow> real^1"
let ?drop = "(\<lambda>x::real^1. x $ 1)"
have S': "?lift ` S \<in> sets lebesgue"
by (auto intro: differentiable_image_in_sets_lebesgue [OF S] differentiable_vec)
have "((\<lambda>x. vec (g (x $ 1))) has_derivative (*\<^sub>R) (g' z)) (at (vec z) within ?lift ` S)"
if "z \<in> S" for z
using der_g [OF that]
by (simp add: has_vector_derivative_def has_derivative_vector_1)
then have der': "\<And>x. x \<in> ?lift ` S \<Longrightarrow>
(?lift \<circ> g \<circ> ?drop has_derivative (*\<^sub>R) (g' (?drop x))) (at x within ?lift ` S)"
by (auto simp: o_def)
have inj': "inj_on (vec \<circ> g \<circ> ?drop) (vec ` S)"
using inj by (simp add: inj_on_def)
let ?fg = "\<lambda>x. \<bar>g' x\<bar> *\<^sub>R f(g x)"
have "((\<lambda>x. ?fg x $ i) absolutely_integrable_on S \<and> ((\<lambda>x. ?fg x $ i) has_integral b $ i) S
\<longleftrightarrow> (\<lambda>x. f x $ i) absolutely_integrable_on g ` S \<and> ((\<lambda>x. f x $ i) has_integral b $ i) (g ` S))" for i
using has_absolute_integral_change_of_variables [OF S' der' inj', of "\<lambda>x. ?lift(f (?drop x) $ i)" "?lift (b$i)"]
unfolding integrable_on_1_iff integral_on_1_eq absolutely_integrable_on_1_iff absolutely_integrable_drop absolutely_integrable_on_def
by (auto simp: image_comp o_def integral_vec1_eq has_integral_iff)
then have "?fg absolutely_integrable_on S \<and> (?fg has_integral b) S
\<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> (f has_integral b) (g ` S)"
unfolding has_integral_componentwise_iff [where y=b]
absolutely_integrable_componentwise_iff [where f=f]
absolutely_integrable_componentwise_iff [where f = ?fg]
by (force simp: Basis_vec_def cart_eq_inner_axis)
then show ?thesis
using absolutely_integrable_on_def by blast
qed
corollary absolutely_integrable_change_of_variables_1:
fixes f :: "real \<Rightarrow> real^'n::{finite,wellorder}" and g :: "real \<Rightarrow> real"
assumes S: "S \<in> sets lebesgue"
and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_vector_derivative g' x) (at x within S)"
and inj: "inj_on g S"
shows "(f absolutely_integrable_on g ` S \<longleftrightarrow>
(\<lambda>x. \<bar>g' x\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S)"
using has_absolute_integral_change_of_variables_1 [OF assms] by auto
subsection\<open>Change of variables for integrals: special case of linear function\<close>
lemma has_absolute_integral_change_of_variables_linear:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes "linear g"
shows "(\<lambda>x. \<bar>det (matrix g)\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and>
integral S (\<lambda>x. \<bar>det (matrix g)\<bar> *\<^sub>R f(g x)) = b
\<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b"
proof (cases "det(matrix g) = 0")
case True
then have "negligible(g ` S)"
using assms det_nz_iff_inj negligible_linear_singular_image by blast
with True show ?thesis
by (auto simp: absolutely_integrable_on_def integrable_negligible integral_negligible)
next
case False
then obtain h where h: "\<And>x. x \<in> S \<Longrightarrow> h (g x) = x" "linear h"
using assms det_nz_iff_inj linear_injective_isomorphism by metis
show ?thesis
proof (rule has_absolute_integral_change_of_variables_invertible)
show "(g has_derivative g) (at x within S)" for x
by (simp add: assms linear_imp_has_derivative)
show "continuous_on (g ` S) h"
using continuous_on_eq_continuous_within has_derivative_continuous linear_imp_has_derivative h by blast
qed (use h in auto)
qed
lemma absolutely_integrable_change_of_variables_linear:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes "linear g"
shows "(\<lambda>x. \<bar>det (matrix g)\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S
\<longleftrightarrow> f absolutely_integrable_on (g ` S)"
using assms has_absolute_integral_change_of_variables_linear by blast
lemma absolutely_integrable_on_linear_image:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes "linear g"
shows "f absolutely_integrable_on (g ` S)
\<longleftrightarrow> (f \<circ> g) absolutely_integrable_on S \<or> det(matrix g) = 0"
unfolding assms absolutely_integrable_change_of_variables_linear [OF assms, symmetric] absolutely_integrable_on_scaleR_iff
by (auto simp: set_integrable_def)
lemma integral_change_of_variables_linear:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes "linear g"
and "f absolutely_integrable_on (g ` S) \<or> (f \<circ> g) absolutely_integrable_on S"
shows "integral (g ` S) f = \<bar>det (matrix g)\<bar> *\<^sub>R integral S (f \<circ> g)"
proof -
have "((\<lambda>x. \<bar>det (matrix g)\<bar> *\<^sub>R f (g x)) absolutely_integrable_on S) \<or> (f absolutely_integrable_on g ` S)"
using absolutely_integrable_on_linear_image assms by blast
moreover
have ?thesis if "((\<lambda>x. \<bar>det (matrix g)\<bar> *\<^sub>R f (g x)) absolutely_integrable_on S)" "(f absolutely_integrable_on g ` S)"
using has_absolute_integral_change_of_variables_linear [OF \<open>linear g\<close>] that
by (auto simp: o_def)
ultimately show ?thesis
using absolutely_integrable_change_of_variables_linear [OF \<open>linear g\<close>]
by blast
qed
subsection\<open>Change of variable for measure\<close>
lemma has_measure_differentiable_image:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
assumes "S \<in> sets lebesgue"
and "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
and "inj_on f S"
shows "f ` S \<in> lmeasurable \<and> measure lebesgue (f ` S) = m
\<longleftrightarrow> ((\<lambda>x. \<bar>det (matrix (f' x))\<bar>) has_integral m) S"
using has_absolute_integral_change_of_variables [OF assms, of "\<lambda>x. (1::real^1)" "vec m"]
unfolding absolutely_integrable_on_1_iff integral_on_1_eq integrable_on_1_iff absolutely_integrable_on_def
by (auto simp: has_integral_iff lmeasurable_iff_integrable_on lmeasure_integral)
lemma measurable_differentiable_image_eq:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
assumes "S \<in> sets lebesgue"
and "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
and "inj_on f S"
shows "f ` S \<in> lmeasurable \<longleftrightarrow> (\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on S"
using has_measure_differentiable_image [OF assms]
by blast
lemma measurable_differentiable_image_alt:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
assumes "S \<in> sets lebesgue"
and "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
and "inj_on f S"
shows "f ` S \<in> lmeasurable \<longleftrightarrow> (\<lambda>x. \<bar>det (matrix (f' x))\<bar>) absolutely_integrable_on S"
using measurable_differentiable_image_eq [OF assms]
by (simp only: absolutely_integrable_on_iff_nonneg)
lemma measure_differentiable_image_eq:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
assumes S: "S \<in> sets lebesgue"
and der_f: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
and inj: "inj_on f S"
and intS: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on S"
shows "measure lebesgue (f ` S) = integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
using measurable_differentiable_image_eq [OF S der_f inj]
assms has_measure_differentiable_image by blast
end
diff --git a/src/HOL/Analysis/Complex_Analysis_Basics.thy b/src/HOL/Analysis/Complex_Analysis_Basics.thy
--- a/src/HOL/Analysis/Complex_Analysis_Basics.thy
+++ b/src/HOL/Analysis/Complex_Analysis_Basics.thy
@@ -1,854 +1,854 @@
(* 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>
text \<open>Definitions of analytic and holomorphic functions, limit theorems, complex differentiation\<close>
theory Complex_Analysis_Basics
imports Derivative "HOL-Library.Nonpos_Ints"
begin
subsection\<^marker>\<open>tag unimportant\<close>\<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 fact_cancel:
fixes c :: "'a::real_field"
shows "of_nat (Suc n) * c / (fact (Suc n)) = c / (fact n)"
using of_nat_neq_0 by force
lemma vector_derivative_cnj_within:
assumes "at x within A \<noteq> bot" and "f differentiable at x within A"
shows "vector_derivative (\<lambda>z. cnj (f z)) (at x within A) =
cnj (vector_derivative f (at x within A))" (is "_ = cnj ?D")
proof -
let ?D = "vector_derivative f (at x within A)"
from assms have "(f has_vector_derivative ?D) (at x within A)"
by (subst (asm) vector_derivative_works)
hence "((\<lambda>x. cnj (f x)) has_vector_derivative cnj ?D) (at x within A)"
by (rule has_vector_derivative_cnj)
thus ?thesis using assms by (auto dest: vector_derivative_within)
qed
lemma vector_derivative_cnj:
assumes "f differentiable at x"
shows "vector_derivative (\<lambda>z. cnj (f z)) (at x) = cnj (vector_derivative f (at x))"
using assms by (intro vector_derivative_cnj_within) auto
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)
lemma 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)
lemma 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 "\<not> 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 sum_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\<^marker>\<open>tag important\<close> 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\<^marker>\<open>tag important\<close> 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_on:
"f holomorphic_on s \<Longrightarrow> f differentiable_on s"
unfolding holomorphic_on_def differentiable_on_def
by (simp add: field_differentiable_imp_differentiable)
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]: "((*) 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_balls_imp_entire:
assumes "\<not>bdd_above A" "\<And>r. r \<in> A \<Longrightarrow> f holomorphic_on ball c r"
shows "f holomorphic_on B"
proof (rule holomorphic_on_subset)
show "f holomorphic_on UNIV" unfolding holomorphic_on_def
proof
fix z :: complex
from \<open>\<not>bdd_above A\<close> obtain r where r: "r \<in> A" "r > norm (z - c)"
by (meson bdd_aboveI not_le)
with assms(2) have "f holomorphic_on ball c r" by blast
moreover from r have "z \<in> ball c r" by (auto simp: dist_norm norm_minus_commute)
ultimately show "f field_differentiable at z"
by (auto simp: holomorphic_on_def at_within_open[of _ "ball c r"])
qed
qed auto
lemma holomorphic_on_balls_imp_entire':
assumes "\<And>r. r > 0 \<Longrightarrow> f holomorphic_on ball c r"
shows "f holomorphic_on B"
proof (rule holomorphic_on_balls_imp_entire)
{
fix M :: real
have "\<exists>x. x > max M 0" by (intro gt_ex)
hence "\<exists>x>0. x > M" by auto
}
thus "\<not>bdd_above {(0::real)<..}" unfolding bdd_above_def
by (auto simp: not_le)
qed (insert assms, auto)
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_sum [holomorphic_intros]:
"(\<And>i. i \<in> I \<Longrightarrow> (f i) holomorphic_on s) \<Longrightarrow> (\<lambda>x. sum (\<lambda>i. f i x) I) holomorphic_on s"
unfolding holomorphic_on_def by (metis field_differentiable_sum)
lemma holomorphic_on_prod [holomorphic_intros]:
"(\<And>i. i \<in> I \<Longrightarrow> (f i) holomorphic_on s) \<Longrightarrow> (\<lambda>x. prod (\<lambda>i. f i x) I) holomorphic_on s"
by (induction I rule: infinite_finite_induct) (auto intro: holomorphic_intros)
lemma holomorphic_pochhammer [holomorphic_intros]:
"f holomorphic_on A \<Longrightarrow> (\<lambda>s. pochhammer (f s) n) holomorphic_on A"
by (induction n) (auto intro!: holomorphic_intros simp: pochhammer_Suc)
lemma holomorphic_on_scaleR [holomorphic_intros]:
"f holomorphic_on A \<Longrightarrow> (\<lambda>x. c *\<^sub>R f x) holomorphic_on A"
by (auto simp: scaleR_conv_of_real intro!: holomorphic_intros)
lemma holomorphic_on_Un [holomorphic_intros]:
assumes "f holomorphic_on A" "f holomorphic_on B" "open A" "open B"
shows "f holomorphic_on (A \<union> B)"
using assms by (auto simp: holomorphic_on_def at_within_open[of _ A]
at_within_open[of _ B] at_within_open[of _ "A \<union> B"] open_Un)
lemma holomorphic_on_If_Un [holomorphic_intros]:
assumes "f holomorphic_on A" "g holomorphic_on B" "open A" "open B"
assumes "\<And>z. z \<in> A \<Longrightarrow> z \<in> B \<Longrightarrow> f z = g z"
shows "(\<lambda>z. if z \<in> A then f z else g z) holomorphic_on (A \<union> B)" (is "?h holomorphic_on _")
proof (intro holomorphic_on_Un)
note \<open>f holomorphic_on A\<close>
also have "f holomorphic_on A \<longleftrightarrow> ?h holomorphic_on A"
by (intro holomorphic_cong) auto
finally show \<dots> .
next
note \<open>g holomorphic_on B\<close>
also have "g holomorphic_on B \<longleftrightarrow> ?h holomorphic_on B"
using assms by (intro holomorphic_cong) auto
finally show \<dots> .
qed (insert assms, auto)
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_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 has_field_derivative_transform_within_open at_within_open)
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"
by (rule nonzero_deriv_nonconstant [of f "deriv f \<xi>" \<xi> S])
(use assms in \<open>auto simp: holomorphic_derivI\<close>)
subsection\<open>Analyticity on a set\<close>
definition\<^marker>\<open>tag important\<close> 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)"
named_theorems\<^marker>\<open>tag important\<close> analytic_intros "introduction rules for proving analyticity"
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 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>\<T>) \<longleftrightarrow> (\<forall>T \<in> \<T>. 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 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 [analytic_intros,simp]: "((*) c) analytic_on S"
by (auto simp add: analytic_on_holomorphic)
lemma analytic_on_const [analytic_intros,simp]: "(\<lambda>z. c) analytic_on S"
by (metis analytic_on_def holomorphic_on_const zero_less_one)
lemma analytic_on_ident [analytic_intros,simp]: "(\<lambda>x. x) analytic_on S"
by (simp add: analytic_on_def gt_ex)
lemma analytic_on_id [analytic_intros]: "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 [analytic_intros]:
"f analytic_on S \<Longrightarrow> (\<lambda>z. -(f z)) analytic_on S"
by (metis analytic_on_holomorphic holomorphic_on_minus)
lemma analytic_on_add [analytic_intros]:
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 [analytic_intros]:
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 [analytic_intros]:
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 [analytic_intros]:
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 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 [analytic_intros]:
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 [analytic_intros]:
"f analytic_on S \<Longrightarrow> (\<lambda>z. (f z) ^ n) analytic_on S"
by (induct n) (auto simp: analytic_on_mult)
lemma analytic_on_sum [analytic_intros]:
"(\<And>i. i \<in> I \<Longrightarrow> (f i) analytic_on S) \<Longrightarrow> (\<lambda>x. sum (\<lambda>i. f i x) I) analytic_on S"
- by (induct I rule: infinite_finite_induct) (auto simp: analytic_on_const analytic_on_add)
+ by (induct I rule: infinite_finite_induct) (auto simp: 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 deriv_chain image_subset_iff)
also have "... = deriv id w"
proof (rule complex_derivative_transform_within_open [where s=S])
show "g \<circ> f holomorphic_on S"
by (rule assms holomorphic_on_compose_gen holomorphic_intros)+
qed (use assms in auto)
also have "... = 1"
by simp
finally show ?thesis .
qed
subsection\<^marker>\<open>tag unimportant\<close>\<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: holomorphic_on_subset)
done
next
assume ?rhs
then show ?lhs
by (force simp add: analytic_at)
qed
subsection\<^marker>\<open>tag unimportant\<close>\<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)
+by (auto simp: complex_derivative_mult_at)
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)
+by (auto simp: complex_derivative_mult_at)
subsection\<^marker>\<open>tag unimportant\<close>\<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 "(\<lambda>n. f n x) \<longlonglongrightarrow> l"
by (rule tf)
next show "\<And>e. e > 0 \<Longrightarrow> \<forall>\<^sub>F n in sequentially. \<forall>x\<in>S. \<forall>h. cmod (f' n x * h - g' x * h) \<le> e * cmod h"
unfolding eventually_sequentially by (blast intro: **)
qed (metis has_field_derivative_def df)
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 sum_distrib_left)
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 "\<And>e. e>0 \<Longrightarrow> \<forall>\<^sub>F n in sequentially. \<forall>x\<in>S. \<forall>h. cmod ((\<Sum>i<n. h * f' i x) - g' x * h) \<le> e * cmod h"
unfolding eventually_sequentially by (blast intro: **)
qed
qed
subsection\<^marker>\<open>tag unimportant\<close> \<open>Taylor on Complex Numbers\<close>
lemma sum_Suc_reindex:
fixes f :: "nat \<Rightarrow> 'a::ab_group_add"
shows "sum f {0..n} = f 0 - f (Suc n) + sum (\<lambda>i. f (Suc i)) {0..n}"
by (induct n) auto
lemma field_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> norm (f (Suc n) x) \<le> B"
and w: "w \<in> S"
and z: "z \<in> S"
shows "norm(f 0 z - (\<Sum>i\<le>n. f i w * (z-w) ^ i / (fact i)))
\<le> B * norm(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 "norm (f (Suc n) u) * norm (z - u) ^ n \<le> norm (f (Suc n) u) * norm (u - z) ^ n"
by (metis norm_minus_commute order_refl)
also have "... \<le> norm (f (Suc n) u) * norm (z - w) ^ n"
by (metis mult_left_mono norm_ge_zero power_mono segment_bound [OF u])
also have "... \<le> B * norm (z - w) ^ n"
by (metis norm_ge_zero zero_le_power mult_right_mono B [OF us])
finally have "norm (f (Suc n) u) * norm (z - u) ^ n \<le> B * norm (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 sum_zero_power) simp
finally have "norm (f 0 z - (\<Sum>i\<le>n. f i w * (z - w) ^ i / (fact i)))
\<le> norm ((\<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 * norm (z - w) ^ n / (fact n) * norm (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]
+ apply (auto simp: DERIV_subset [OF sum_deriv wzs]
norm_divide norm_mult norm_power divide_le_cancel cmod_bound)
done
also have "... \<le> B * norm (z - w) ^ Suc n / (fact n)"
by (simp add: algebra_simps norm_minus_commute)
finally show ?thesis .
qed
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"
using assms by (rule field_Taylor)
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_withinI, 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 sum_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 sum_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
end
diff --git a/src/HOL/Analysis/Complex_Transcendental.thy b/src/HOL/Analysis/Complex_Transcendental.thy
--- a/src/HOL/Analysis/Complex_Transcendental.thy
+++ b/src/HOL/Analysis/Complex_Transcendental.thy
@@ -1,4058 +1,4058 @@
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 "HOL-Library.Periodic_Fun"
begin
subsection\<open>Möbius transformations\<close>
(* TODO: Figure out what to do with Möbius transformations *)
definition\<^marker>\<open>tag important\<close> "moebius a b c d = (\<lambda>z. (a*z+b) / (c*z+d :: 'a :: field))"
theorem 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 (simp add: 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)"
proof (cases "Im z = 0 \<or> Re z = 0")
case True
with assms abs_square_less_1 show ?thesis
by (force simp add: Re_power2 Im_power2 cmod_def)
next
case False
with cmod_diff_real_less [of "1 - z\<^sup>2" "1"] show ?thesis
by (simp add: norm_power Im_power2)
qed
subsection\<^marker>\<open>tag unimportant\<close>\<open>The Exponential Function\<close>
lemma norm_exp_i_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
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)
lemma holomorphic_on_exp' [holomorphic_intros]:
"f holomorphic_on s \<Longrightarrow> (\<lambda>x. exp (f x)) holomorphic_on s"
using holomorphic_on_compose[OF _ holomorphic_on_exp] by (simp add: o_def)
subsection\<open>Euler and de Moivre formulas\<close>
text\<open>The sine series times \<^term>\<open>i\<close>\<close>
lemma sin_i_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_i_eq [of z]]
by (simp add: field_simps scaleR_conv_of_real)
ultimately show ?thesis
using sums_unique2 by blast
qed
corollary\<^marker>\<open>tag unimportant\<close> 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)
theorem 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 Complex_eq)
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\<^marker>\<open>tag unimportant\<close>\<open>Relationships between real and complex trigonometric and hyperbolic functions\<close>
lemma real_sin_eq [simp]: "Re(sin(of_real x)) = sin x"
by (simp add: sin_of_real)
lemma real_cos_eq [simp]: "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)"
by (metis exp_Euler [symmetric] exp_of_nat_mult mult.left_commute)
lemma exp_cnj: "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)
lemma holomorphic_on_sin' [holomorphic_intros]:
assumes "f holomorphic_on A"
shows "(\<lambda>x. sin (f x)) holomorphic_on A"
using holomorphic_on_compose[OF assms holomorphic_on_sin] by (simp add: o_def)
lemma holomorphic_on_cos' [holomorphic_intros]:
assumes "f holomorphic_on A"
shows "(\<lambda>x. cos (f x)) holomorphic_on A"
using holomorphic_on_compose[OF assms holomorphic_on_cos] by (simp add: o_def)
subsection\<^marker>\<open>tag unimportant\<close>\<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: Complex_eq 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)"
(is "?lhs = ?rhs")
proof
assume "exp z = 1"
then have "Re z = 0"
by (metis exp_eq_one_iff norm_exp_eq_Re norm_one)
with \<open>?lhs\<close> show ?rhs
by (metis Re_exp complex_Re_of_int cos_one_2pi_int exp_zero mult.commute mult_numeral_1 numeral_One of_int_mult of_int_numeral)
next
assume ?rhs then show ?lhs
using Im_exp Re_exp complex_eq_iff
by (simp add: cos_one_2pi_int cos_one_sin_zero mult.commute)
qed
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 unfolding Ints_def exp_eq by auto
also have "... = 1"
by simp
finally show ?thesis .
qed
lemma exp_plus_2pin [simp]: "exp (z + \<i> * (of_int n * (of_real pi * 2))) = exp z"
by (simp add: exp_eq)
lemma exp_integer_2pi_plus1:
assumes "n \<in> \<int>"
shows "exp(((2 * n + 1) * pi) * \<i>) = - 1"
proof -
from assms obtain n' where [simp]: "n = of_int n'"
by (auto simp: Ints_def)
have "exp(((2 * n + 1) * pi) * \<i>) = exp (pi * \<i>)"
using assms by (subst exp_eq) (auto intro!: exI[of _ n'] simp: algebra_simps)
also have "... = - 1"
by simp
finally show ?thesis .
qed
lemma inj_on_exp_pi:
fixes z::complex shows "inj_on exp (ball z pi)"
proof (clarsimp simp: inj_on_def exp_eq)
fix y n
assume "dist z (y + 2 * of_int n * of_real pi * \<i>) < pi"
"dist z y < pi"
then have "dist y (y + 2 * of_int n * of_real pi * \<i>) < pi+pi"
using dist_commute_lessI dist_triangle_less_add by blast
then have "norm (2 * of_int n * of_real pi * \<i>) < 2*pi"
by (simp add: dist_norm)
then show "n = 0"
by (auto simp: norm_mult)
qed
lemma cmod_add_squared:
fixes r1 r2::real
assumes "r1 \<ge> 0" "r2 \<ge> 0"
shows "(cmod (r1 * exp (\<i> * \<theta>1) + r2 * exp (\<i> * \<theta>2)))\<^sup>2 = r1\<^sup>2 + r2\<^sup>2 + 2 * r1 * r2 * cos (\<theta>1 - \<theta>2)" (is "(cmod (?z1 + ?z2))\<^sup>2 = ?rhs")
proof -
have "(cmod (?z1 + ?z2))\<^sup>2 = (?z1 + ?z2) * cnj (?z1 + ?z2)"
by (rule complex_norm_square)
also have "\<dots> = (?z1 * cnj ?z1 + ?z2 * cnj ?z2) + (?z1 * cnj ?z2 + cnj ?z1 * ?z2)"
by (simp add: algebra_simps)
also have "\<dots> = (norm ?z1)\<^sup>2 + (norm ?z2)\<^sup>2 + 2 * Re (?z1 * cnj ?z2)"
unfolding complex_norm_square [symmetric] cnj_add_mult_eq_Re by simp
also have "\<dots> = ?rhs"
by (simp add: norm_mult) (simp add: exp_Euler complex_is_Real_iff [THEN iffD1] cos_diff algebra_simps)
finally show ?thesis
using of_real_eq_iff by blast
qed
lemma cmod_diff_squared:
fixes r1 r2::real
assumes "r1 \<ge> 0" "r2 \<ge> 0"
shows "(cmod (r1 * exp (\<i> * \<theta>1) - r2 * exp (\<i> * \<theta>2)))\<^sup>2 = r1\<^sup>2 + r2\<^sup>2 - 2*r1*r2*cos (\<theta>1 - \<theta>2)" (is "(cmod (?z1 - ?z2))\<^sup>2 = ?rhs")
proof -
have "exp (\<i> * (\<theta>2 + pi)) = - exp (\<i> * \<theta>2)"
by (simp add: exp_Euler cos_plus_pi sin_plus_pi)
then have "(cmod (?z1 - ?z2))\<^sup>2 = cmod (?z1 + r2 * exp (\<i> * (\<theta>2 + pi))) ^2"
by simp
also have "\<dots> = r1\<^sup>2 + r2\<^sup>2 + 2*r1*r2*cos (\<theta>1 - (\<theta>2 + pi))"
using assms cmod_add_squared by blast
also have "\<dots> = ?rhs"
by (simp add: add.commute diff_add_eq_diff_diff_swap)
finally show ?thesis .
qed
lemma polar_convergence:
fixes R::real
assumes "\<And>j. r j > 0" "R > 0"
shows "((\<lambda>j. r j * exp (\<i> * \<theta> j)) \<longlonglongrightarrow> (R * exp (\<i> * \<Theta>))) \<longleftrightarrow>
(r \<longlonglongrightarrow> R) \<and> (\<exists>k. (\<lambda>j. \<theta> j - of_int (k j) * (2 * pi)) \<longlonglongrightarrow> \<Theta>)" (is "(?z \<longlonglongrightarrow> ?Z) = ?rhs")
proof
assume L: "?z \<longlonglongrightarrow> ?Z"
have rR: "r \<longlonglongrightarrow> R"
using tendsto_norm [OF L] assms by (auto simp: norm_mult abs_of_pos)
moreover obtain k where "(\<lambda>j. \<theta> j - of_int (k j) * (2 * pi)) \<longlonglongrightarrow> \<Theta>"
proof -
have "cos (\<theta> j - \<Theta>) = ((r j)\<^sup>2 + R\<^sup>2 - (norm(?z j - ?Z))\<^sup>2) / (2 * R * r j)" for j
apply (subst cmod_diff_squared)
using assms by (auto simp: field_split_simps less_le)
moreover have "(\<lambda>j. ((r j)\<^sup>2 + R\<^sup>2 - (norm(?z j - ?Z))\<^sup>2) / (2 * R * r j)) \<longlonglongrightarrow> ((R\<^sup>2 + R\<^sup>2 - (norm(?Z - ?Z))\<^sup>2) / (2 * R * R))"
by (intro L rR tendsto_intros) (use \<open>R > 0\<close> in force)
moreover have "((R\<^sup>2 + R\<^sup>2 - (norm(?Z - ?Z))\<^sup>2) / (2 * R * R)) = 1"
using \<open>R > 0\<close> by (simp add: power2_eq_square field_split_simps)
ultimately have "(\<lambda>j. cos (\<theta> j - \<Theta>)) \<longlonglongrightarrow> 1"
by auto
then show ?thesis
using that cos_diff_limit_1 by blast
qed
ultimately show ?rhs
by metis
next
assume R: ?rhs
show "?z \<longlonglongrightarrow> ?Z"
proof (rule tendsto_mult)
show "(\<lambda>x. complex_of_real (r x)) \<longlonglongrightarrow> of_real R"
using R by (auto simp: tendsto_of_real_iff)
obtain k where "(\<lambda>j. \<theta> j - of_int (k j) * (2 * pi)) \<longlonglongrightarrow> \<Theta>"
using R by metis
then have "(\<lambda>j. complex_of_real (\<theta> j - of_int (k j) * (2 * pi))) \<longlonglongrightarrow> of_real \<Theta>"
using tendsto_of_real_iff by force
then have "(\<lambda>j. exp (\<i> * of_real (\<theta> j - of_int (k j) * (2 * pi)))) \<longlonglongrightarrow> exp (\<i> * \<Theta>)"
using tendsto_mult [OF tendsto_const] isCont_exp isCont_tendsto_compose by blast
moreover have "exp (\<i> * of_real (\<theta> j - of_int (k j) * (2 * pi))) = exp (\<i> * \<theta> j)" for j
unfolding exp_eq
by (rule_tac x="- k j" in exI) (auto simp: algebra_simps)
ultimately show "(\<lambda>j. exp (\<i> * \<theta> j)) \<longlonglongrightarrow> exp (\<i> * \<Theta>)"
by auto
qed
qed
lemma sin_cos_eq_iff: "sin y = sin x \<and> cos y = cos x \<longleftrightarrow> (\<exists>n::int. y = x + 2 * pi * n)"
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 = 2 * pi * n"
using cos_one_2pi_int by auto }
then show ?thesis
apply (auto simp: sin_add cos_add)
apply (metis add.commute diff_add_cancel)
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)
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)
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 safe
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"]
by (simp add: sin_diff algebra_simps equation_minus_iff)
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)"
by (rule iff_exI) (auto simp: algebra_simps intro: injD [OF inj_of_real [where 'a = complex]])
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)"
by (rule iff_exI) (auto simp: algebra_simps intro: injD [OF inj_of_real [where 'a = complex]])
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)"
by (rule iff_exI) (auto simp: algebra_simps intro: injD [OF inj_of_real [where 'a = complex]])
also have "... = ?rhs"
by (auto simp: algebra_simps)
finally show ?thesis .
qed
lemma dist_exp_i_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 complex_sin_eq:
fixes w :: complex
shows "sin w = sin z \<longleftrightarrow> (\<exists>n \<in> \<int>. w = z + of_real(2*n*pi) \<or> w = -z + of_real((2*n + 1)*pi))"
(is "?lhs = ?rhs")
proof
assume ?lhs
then have "sin w - sin z = 0"
by (auto simp: algebra_simps)
then have "sin ((w - z) / 2)*cos ((w + z) / 2) = 0"
by (auto simp: sin_diff_sin)
then consider "sin ((w - z) / 2) = 0" | "cos ((w + z) / 2) = 0"
using mult_eq_0_iff by blast
then show ?rhs
proof cases
case 1
then show ?thesis
by (simp add: sin_eq_0 algebra_simps) (metis Ints_of_int of_real_of_int_eq)
next
case 2
then show ?thesis
by (simp add: cos_eq_0 algebra_simps) (metis Ints_of_int of_real_of_int_eq)
qed
next
assume ?rhs
then obtain n::int where w: "w = z + of_real (2* of_int n*pi) \<or>
w = -z + of_real ((2* of_int n + 1)*pi)"
using Ints_cases by blast
then show ?lhs
using Periodic_Fun.sin.plus_of_int [of z n]
apply (auto simp: algebra_simps)
by (metis (no_types, hide_lams) add_diff_cancel_left add_diff_cancel_left' add_minus_cancel
mult.commute sin.plus_of_int sin_minus sin_plus_pi)
qed
lemma complex_cos_eq:
fixes w :: complex
shows "cos w = cos z \<longleftrightarrow> (\<exists>n \<in> \<int>. w = z + of_real(2*n*pi) \<or> w = -z + of_real(2*n*pi))"
(is "?lhs = ?rhs")
proof
assume ?lhs
then have "cos w - cos z = 0"
by (auto simp: algebra_simps)
then have "sin ((w + z) / 2) * sin ((z - w) / 2) = 0"
by (auto simp: cos_diff_cos)
then consider "sin ((w + z) / 2) = 0" | "sin ((z - w) / 2) = 0"
using mult_eq_0_iff by blast
then show ?rhs
proof cases
case 1
then show ?thesis
apply (simp add: sin_eq_0 algebra_simps)
by (metis Ints_of_int of_real_of_int_eq)
next
case 2
then show ?thesis
apply (clarsimp simp: sin_eq_0 algebra_simps)
by (metis Ints_of_int add_minus_cancel distrib_right mult_of_int_commute mult_zero_right of_int_0 of_int_add of_real_of_int_eq)
qed
next
assume ?rhs
then obtain n::int where w: "w = z + of_real (2* of_int n*pi) \<or>
w = -z + of_real(2*n*pi)"
using Ints_cases by (metis of_int_mult of_int_numeral)
then show ?lhs
using Periodic_Fun.cos.plus_of_int [of z n]
apply (simp add: algebra_simps)
by (metis cos.plus_of_int cos_minus minus_add_cancel mult.commute)
qed
lemma sin_eq:
"sin x = sin y \<longleftrightarrow> (\<exists>n \<in> \<int>. x = y + 2*n*pi \<or> x = -y + (2*n + 1)*pi)"
using complex_sin_eq [of x y]
by (simp only: sin_of_real Re_complex_of_real of_real_add [symmetric] of_real_minus [symmetric] of_real_mult [symmetric] of_real_eq_iff)
lemma cos_eq:
"cos x = cos y \<longleftrightarrow> (\<exists>n \<in> \<int>. x = y + 2*n*pi \<or> x = -y + 2*n*pi)"
using complex_cos_eq [of x y]
by (simp only: cos_of_real Re_complex_of_real of_real_add [symmetric] of_real_minus [symmetric] of_real_mult [symmetric] of_real_eq_iff)
lemma sinh_complex:
fixes z :: complex
shows "(exp z - inverse (exp z)) / 2 = -\<i> * sin(\<i> * z)"
by (simp add: sin_exp_eq field_split_simps exp_minus)
lemma sin_i_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_i_times)
lemma cosh_complex:
fixes z :: complex
shows "(exp z + inverse (exp z)) / 2 = cos(\<i> * z)"
by (simp add: cos_exp_eq field_split_simps exp_minus 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 field_split_simps exp_minus exp_of_real)
lemmas cos_i_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 Complex_eq)
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 sin_squared_eq)
- apply (simp add: power2_eq_square algebra_simps field_split_simps)
+ apply (simp add: power2_eq_square field_split_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 Complex_eq)
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 cos_squared_eq)
- apply (simp add: power2_eq_square algebra_simps field_split_simps)
+ apply (simp add: power2_eq_square field_split_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
by (metis exp_eq_one_iff exp_le_cancel_iff mult_2 norm_cos_le norm_ge_zero norm_one norm_triangle_mono)
qed
lemma sinh_conv_sin: "sinh z = -\<i> * sin (\<i>*z)"
by (simp add: sinh_field_def sin_i_times exp_minus)
lemma cosh_conv_cos: "cosh z = cos (\<i>*z)"
by (simp add: cosh_field_def cos_i_times exp_minus)
lemma tanh_conv_tan: "tanh z = -\<i> * tan (\<i>*z)"
by (simp add: tanh_def sinh_conv_sin cosh_conv_cos tan_def)
lemma sin_conv_sinh: "sin z = -\<i> * sinh (\<i>*z)"
by (simp add: sinh_conv_sin)
lemma cos_conv_cosh: "cos z = cosh (\<i>*z)"
by (simp add: cosh_conv_cos)
lemma tan_conv_tanh: "tan z = -\<i> * tanh (\<i>*z)"
by (simp add: tan_def sin_conv_sinh cos_conv_cosh tanh_def)
lemma sinh_complex_eq_iff:
"sinh (z :: complex) = sinh w \<longleftrightarrow>
(\<exists>n\<in>\<int>. z = w - 2 * \<i> * of_real n * of_real pi \<or>
z = -(2 * complex_of_real n + 1) * \<i> * complex_of_real pi - w)" (is "_ = ?rhs")
proof -
have "sinh z = sinh w \<longleftrightarrow> sin (\<i> * z) = sin (\<i> * w)"
by (simp add: sinh_conv_sin)
also have "\<dots> \<longleftrightarrow> ?rhs"
by (subst complex_sin_eq) (force simp: field_simps complex_eq_iff)
finally show ?thesis .
qed
subsection\<^marker>\<open>tag unimportant\<close>\<open>Taylor series for complex exponential, sine and cosine\<close>
declare power_Suc [simp del]
lemma Taylor_exp_field:
fixes z::"'a::{banach,real_normed_field}"
shows "norm (exp z - (\<Sum>i\<le>n. z ^ i / fact i)) \<le> exp (norm z) * (norm z ^ Suc n) / fact n"
proof (rule field_Taylor[of _ n "\<lambda>k. exp" "exp (norm z)" 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: "x \<in> closed_segment 0 z"
have "norm (exp x) \<le> exp (norm x)"
by (rule norm_exp)
also have "norm x \<le> norm z"
using x by (auto simp: closed_segment_def intro!: mult_left_le_one_le)
finally show "norm (exp x) \<le> exp (norm z)"
by simp
qed auto
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 (clarsimp 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)
qed auto
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)/2 \<le> u"
by simp
have *: "(cmod (exp (\<i> * (u * z))) + cmod (exp (- (\<i> * (u * z)))) ) / 2 \<le> exp \<bar>Im z\<bar>"
proof (rule mono)
show "cmod (exp (\<i> * (u * z))) \<le> exp \<bar>Im z\<bar>"
apply (simp add: abs_if mult_left_le_one_le assms not_less)
by (meson assms(1) mult_nonneg_nonneg neg_le_0_iff_le order_trans)
show "cmod (exp (- (\<i> * (u * z)))) \<le> exp \<bar>Im z\<bar>"
apply (simp add: abs_if mult_left_le_one_le assms)
by (meson \<open>0 \<le> u\<close> less_eq_real_def mult_nonneg_nonpos neg_0_le_iff_le order_trans)
qed
have "cmod (sin (u *\<^sub>R z)) = cmod (exp (\<i> * (u * z)) - exp (- (\<i> * (u * z)))) / 2"
by (auto simp: scaleR_conv_of_real norm_mult norm_power sin_exp_eq norm_divide)
also have "... \<le> (cmod (exp (\<i> * (u * z))) + cmod (exp (- (\<i> * (u * z)))) ) / 2"
by (intro divide_right_mono norm_triangle_ineq4) simp
also have "... \<le> exp \<bar>Im z\<bar>"
by (rule *)
finally show "cmod (sin (u *\<^sub>R z)) \<le> exp \<bar>Im z\<bar>" .
have "cmod (cos (u *\<^sub>R z)) = cmod (exp (\<i> * (u * z)) + exp (- (\<i> * (u * z)))) / 2"
by (auto simp: scaleR_conv_of_real norm_mult norm_power cos_exp_eq norm_divide)
also have "... \<le> (cmod (exp (\<i> * (u * z))) + cmod (exp (- (\<i> * (u * z)))) ) / 2"
by (intro divide_right_mono norm_triangle_ineq) simp
also have "... \<le> exp \<bar>Im z\<bar>"
by (rule *)
finally show "cmod (cos (u *\<^sub>R z)) \<le> exp \<bar>Im z\<bar>" .
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: sum_distrib_right in_Reals_norm Re_exp atMost_nat_numeral fact_numeral)
apply (simp only: pos_le_divide_eq [symmetric])
done
lemma e_less_272: "exp 1 < (272/100::real)"
using e_approx_32
by (simp add: abs_if split: if_split_asm)
lemma ln_272_gt_1: "ln (272/100) > (1::real)"
by (metis e_less_272 exp_less_cancel_iff exp_ln_iff less_trans ln_exp)
text\<open>Apparently redundant. But many arguments involve integers.\<close>
lemma ln3_gt_1: "ln 3 > (1::real)"
by (simp add: less_trans [OF ln_272_gt_1])
subsection\<open>The argument of a complex number (HOL Light version)\<close>
definition\<^marker>\<open>tag important\<close> is_Arg :: "[complex,real] \<Rightarrow> bool"
where "is_Arg z r \<equiv> z = of_real(norm z) * exp(\<i> * of_real r)"
definition\<^marker>\<open>tag important\<close> Arg2pi :: "complex \<Rightarrow> real"
where "Arg2pi z \<equiv> if z = 0 then 0 else THE t. 0 \<le> t \<and> t < 2*pi \<and> is_Arg z t"
lemma is_Arg_2pi_iff: "is_Arg z (r + of_int k * (2 * pi)) \<longleftrightarrow> is_Arg z r"
by (simp add: algebra_simps is_Arg_def)
lemma is_Arg_eqI:
assumes r: "is_Arg z r" and s: "is_Arg z s" and rs: "abs (r-s) < 2*pi" and "z \<noteq> 0"
shows "r = s"
proof -
have zr: "z = (cmod z) * exp (\<i> * r)" and zs: "z = (cmod z) * exp (\<i> * s)"
using r s by (auto simp: is_Arg_def)
with \<open>z \<noteq> 0\<close> have eq: "exp (\<i> * r) = exp (\<i> * s)"
by (metis mult_eq_0_iff mult_left_cancel)
have "\<i> * r = \<i> * s"
by (rule exp_complex_eqI) (use rs in \<open>auto simp: eq exp_complex_eqI\<close>)
then show ?thesis
by simp
qed
text\<open>This function returns the angle of a complex number from its representation in polar coordinates.
Due to periodicity, its range is arbitrary. \<^term>\<open>Arg2pi\<close> follows HOL Light in adopting the interval \<open>[0,2\<pi>)\<close>.
But we have the same periodicity issue with logarithms, and it is usual to adopt the same interval
for the complex logarithm and argument functions. Further on down, we shall define both functions for the interval \<open>(-\<pi>,\<pi>]\<close>.
The present version is provided for compatibility.\<close>
lemma Arg2pi_0 [simp]: "Arg2pi(0) = 0"
by (simp add: Arg2pi_def)
lemma Arg2pi_unique_lemma:
assumes z: "is_Arg z t"
and z': "is_Arg z 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' is_Arg_def)
then have "exp (\<i> * of_real t') = exp (\<i> * of_real t)"
by (metis nz mult_left_cancel mult_zero_left z is_Arg_def)
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"
by (cases n) (use t t' in \<open>auto simp: mult_less_0_iff algebra_simps\<close>)
then show "t' = t"
by (simp add: n)
qed
lemma Arg2pi: "0 \<le> Arg2pi z \<and> Arg2pi z < 2*pi \<and> is_Arg z (Arg2pi z)"
proof (cases "z=0")
case True then show ?thesis
by (simp add: Arg2pi_def is_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: "is_Arg z t"
unfolding is_Arg_def
apply (rule complex_eqI)
using t False ReIm
apply (auto simp: exp_Euler sin_of_real cos_of_real field_split_simps)
done
show ?thesis
apply (simp add: Arg2pi_def False)
apply (rule theI [where a=t])
using t z False
apply (auto intro: Arg2pi_unique_lemma)
done
qed
corollary\<^marker>\<open>tag unimportant\<close>
shows Arg2pi_ge_0: "0 \<le> Arg2pi z"
and Arg2pi_lt_2pi: "Arg2pi z < 2*pi"
and Arg2pi_eq: "z = of_real(norm z) * exp(\<i> * of_real(Arg2pi z))"
using Arg2pi is_Arg_def by auto
lemma complex_norm_eq_1_exp: "norm z = 1 \<longleftrightarrow> exp(\<i> * of_real (Arg2pi z)) = z"
by (metis Arg2pi_eq cis_conv_exp mult.left_neutral norm_cis of_real_1)
lemma Arg2pi_unique: "\<lbrakk>of_real r * exp(\<i> * of_real a) = z; 0 < r; 0 \<le> a; a < 2*pi\<rbrakk> \<Longrightarrow> Arg2pi z = a"
by (rule Arg2pi_unique_lemma [unfolded is_Arg_def, OF _ Arg2pi_eq]) (use Arg2pi [of z] in \<open>auto simp: norm_mult\<close>)
lemma Arg2pi_minus: "z \<noteq> 0 \<Longrightarrow> Arg2pi (-z) = (if Arg2pi z < pi then Arg2pi z + pi else Arg2pi z - pi)"
apply (rule Arg2pi_unique [of "norm z"])
apply (rule complex_eqI)
using Arg2pi_ge_0 [of z] Arg2pi_eq [of z] Arg2pi_lt_2pi [of z]
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 Arg2pi_times_of_real [simp]:
assumes "0 < r" shows "Arg2pi (of_real r * z) = Arg2pi z"
proof (cases "z=0")
case False
show ?thesis
by (rule Arg2pi_unique [of "r * norm z"]) (use Arg2pi False assms is_Arg_def in auto)
qed auto
lemma Arg2pi_times_of_real2 [simp]: "0 < r \<Longrightarrow> Arg2pi (z * of_real r) = Arg2pi z"
by (metis Arg2pi_times_of_real mult.commute)
lemma Arg2pi_divide_of_real [simp]: "0 < r \<Longrightarrow> Arg2pi (z / of_real r) = Arg2pi z"
by (metis Arg2pi_times_of_real2 less_numeral_extra(3) nonzero_eq_divide_eq of_real_eq_0_iff)
lemma Arg2pi_le_pi: "Arg2pi z \<le> pi \<longleftrightarrow> 0 \<le> Im z"
proof (cases "z=0")
case False
have "0 \<le> Im z \<longleftrightarrow> 0 \<le> Im (of_real (cmod z) * exp (\<i> * complex_of_real (Arg2pi z)))"
by (metis Arg2pi_eq)
also have "... = (0 \<le> Im (exp (\<i> * complex_of_real (Arg2pi z))))"
using False by (simp add: zero_le_mult_iff)
also have "... \<longleftrightarrow> Arg2pi z \<le> pi"
by (simp add: Im_exp) (metis Arg2pi_ge_0 Arg2pi_lt_2pi sin_lt_zero sin_ge_zero not_le)
finally show ?thesis
by blast
qed auto
lemma Arg2pi_lt_pi: "0 < Arg2pi z \<and> Arg2pi z < pi \<longleftrightarrow> 0 < Im z"
proof (cases "z=0")
case False
have "0 < Im z \<longleftrightarrow> 0 < Im (of_real (cmod z) * exp (\<i> * complex_of_real (Arg2pi z)))"
by (metis Arg2pi_eq)
also have "... = (0 < Im (exp (\<i> * complex_of_real (Arg2pi z))))"
using False by (simp add: zero_less_mult_iff)
also have "... \<longleftrightarrow> 0 < Arg2pi z \<and> Arg2pi z < pi"
using Arg2pi_ge_0 Arg2pi_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 auto
lemma Arg2pi_eq_0: "Arg2pi z = 0 \<longleftrightarrow> z \<in> \<real> \<and> 0 \<le> Re z"
proof (cases "z=0")
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 (Arg2pi z)))"
by (metis Arg2pi_eq)
also have "... \<longleftrightarrow> z \<in> \<real> \<and> 0 \<le> Re (exp (\<i> * complex_of_real (Arg2pi z)))"
using False by (simp add: zero_le_mult_iff)
also have "... \<longleftrightarrow> Arg2pi z = 0"
proof -
have [simp]: "Arg2pi z = 0 \<Longrightarrow> z \<in> \<real>"
using Arg2pi_eq [of z] by (auto simp: Reals_def)
moreover have "\<lbrakk>z \<in> \<real>; 0 \<le> cos (Arg2pi z)\<rbrakk> \<Longrightarrow> Arg2pi z = 0"
by (metis Arg2pi_lt_pi Arg2pi_ge_0 Arg2pi_le_pi cos_pi complex_is_Real_iff leD less_linear less_minus_one_simps(2) minus_minus neg_less_eq_nonneg order_refl)
ultimately show ?thesis
by (auto simp: Re_exp)
qed
finally show ?thesis
by blast
qed auto
corollary\<^marker>\<open>tag unimportant\<close> Arg2pi_gt_0:
assumes "z \<notin> \<real>\<^sub>\<ge>\<^sub>0"
shows "Arg2pi z > 0"
using Arg2pi_eq_0 Arg2pi_ge_0 assms dual_order.strict_iff_order
unfolding nonneg_Reals_def by fastforce
lemma Arg2pi_eq_pi: "Arg2pi z = pi \<longleftrightarrow> z \<in> \<real> \<and> Re z < 0"
using Arg2pi_le_pi [of z] Arg2pi_lt_pi [of z] Arg2pi_eq_0 [of z] Arg2pi_ge_0 [of z]
by (fastforce simp: complex_is_Real_iff)
lemma Arg2pi_eq_0_pi: "Arg2pi z = 0 \<or> Arg2pi z = pi \<longleftrightarrow> z \<in> \<real>"
using Arg2pi_eq_0 Arg2pi_eq_pi not_le by auto
lemma Arg2pi_of_real: "Arg2pi (of_real r) = (if r<0 then pi else 0)"
using Arg2pi_eq_0_pi Arg2pi_eq_pi by fastforce
lemma Arg2pi_real: "z \<in> \<real> \<Longrightarrow> Arg2pi z = (if 0 \<le> Re z then 0 else pi)"
using Arg2pi_eq_0 Arg2pi_eq_0_pi by auto
lemma Arg2pi_inverse: "Arg2pi(inverse z) = (if z \<in> \<real> then Arg2pi z else 2*pi - Arg2pi z)"
proof (cases "z=0")
case False
show ?thesis
apply (rule Arg2pi_unique [of "inverse (norm z)"])
using Arg2pi_eq False Arg2pi_ge_0 [of z] Arg2pi_lt_2pi [of z] Arg2pi_eq_0 [of z]
by (auto simp: Arg2pi_real in_Reals_norm exp_diff field_simps)
qed auto
lemma Arg2pi_eq_iff:
assumes "w \<noteq> 0" "z \<noteq> 0"
shows "Arg2pi w = Arg2pi z \<longleftrightarrow> (\<exists>x. 0 < x & w = of_real x * z)"
using assms Arg2pi_eq [of z] Arg2pi_eq [of w]
apply auto
apply (rule_tac x="norm w / norm z" in exI)
apply (simp add: field_split_simps)
by (metis mult.commute mult.left_commute)
lemma Arg2pi_inverse_eq_0: "Arg2pi(inverse z) = 0 \<longleftrightarrow> Arg2pi z = 0"
by (metis Arg2pi_eq_0 Arg2pi_inverse inverse_inverse_eq)
lemma Arg2pi_divide:
assumes "w \<noteq> 0" "z \<noteq> 0" "Arg2pi w \<le> Arg2pi z"
shows "Arg2pi(z / w) = Arg2pi z - Arg2pi w"
apply (rule Arg2pi_unique [of "norm(z / w)"])
using assms Arg2pi_eq Arg2pi_ge_0 [of w] Arg2pi_lt_2pi [of z]
apply (auto simp: exp_diff norm_divide field_simps)
done
lemma Arg2pi_le_div_sum:
assumes "w \<noteq> 0" "z \<noteq> 0" "Arg2pi w \<le> Arg2pi z"
shows "Arg2pi z = Arg2pi w + Arg2pi(z / w)"
by (simp add: Arg2pi_divide assms)
lemma Arg2pi_le_div_sum_eq:
assumes "w \<noteq> 0" "z \<noteq> 0"
shows "Arg2pi w \<le> Arg2pi z \<longleftrightarrow> Arg2pi z = Arg2pi w + Arg2pi(z / w)"
using assms by (auto simp: Arg2pi_ge_0 intro: Arg2pi_le_div_sum)
lemma Arg2pi_diff:
assumes "w \<noteq> 0" "z \<noteq> 0"
shows "Arg2pi w - Arg2pi z = (if Arg2pi z \<le> Arg2pi w then Arg2pi(w / z) else Arg2pi(w/z) - 2*pi)"
using assms Arg2pi_divide Arg2pi_inverse [of "w/z"] Arg2pi_eq_0_pi
by (force simp add: Arg2pi_ge_0 Arg2pi_divide not_le split: if_split_asm)
lemma Arg2pi_add:
assumes "w \<noteq> 0" "z \<noteq> 0"
shows "Arg2pi w + Arg2pi z = (if Arg2pi w + Arg2pi z < 2*pi then Arg2pi(w * z) else Arg2pi(w * z) + 2*pi)"
using assms Arg2pi_diff [of "w*z" z] Arg2pi_le_div_sum_eq [of z "w*z"]
apply (auto simp: Arg2pi_ge_0 Arg2pi_divide not_le)
apply (metis Arg2pi_lt_2pi add.commute)
apply (metis (no_types) Arg2pi add.commute diff_0 diff_add_cancel diff_less_eq diff_minus_eq_add not_less)
done
lemma Arg2pi_times:
assumes "w \<noteq> 0" "z \<noteq> 0"
shows "Arg2pi (w * z) = (if Arg2pi w + Arg2pi z < 2*pi then Arg2pi w + Arg2pi z
else (Arg2pi w + Arg2pi z) - 2*pi)"
using Arg2pi_add [OF assms]
by auto
lemma Arg2pi_cnj_eq_inverse: "z\<noteq>0 \<Longrightarrow> Arg2pi (cnj z) = Arg2pi (inverse z)"
- apply (simp add: Arg2pi_eq_iff field_split_simps complex_norm_square [symmetric] mult.commute)
+ apply (simp add: Arg2pi_eq_iff field_split_simps complex_norm_square [symmetric])
by (metis of_real_power zero_less_norm_iff zero_less_power)
lemma Arg2pi_cnj: "Arg2pi(cnj z) = (if z \<in> \<real> then Arg2pi z else 2*pi - Arg2pi z)"
proof (cases "z=0")
case False
then show ?thesis
by (simp add: Arg2pi_cnj_eq_inverse Arg2pi_inverse)
qed auto
lemma Arg2pi_exp: "0 \<le> Im z \<Longrightarrow> Im z < 2*pi \<Longrightarrow> Arg2pi(exp z) = Im z"
by (rule Arg2pi_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 Arg2pi cis.ctr cis_conv_exp unfolding Complex_eq is_Arg_def 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\<^marker>\<open>tag unimportant\<close>\<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 \<noteq> 0 \<Longrightarrow> tan field_differentiable at z"
unfolding field_differentiable_def
using DERIV_tan by blast
lemma field_differentiable_within_tan: "cos z \<noteq> 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 \<noteq> 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 \<noteq> 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 \<noteq> 0) \<Longrightarrow> tan holomorphic_on s"
by (simp add: field_differentiable_within_tan holomorphic_on_def)
subsection\<open>The principal branch of the Complex logarithm\<close>
instantiation complex :: ln
begin
definition\<^marker>\<open>tag important\<close> ln_complex :: "complex \<Rightarrow> complex"
where "ln_complex \<equiv> \<lambda>z. THE w. exp w = z & -pi < Im(w) & Im(w) \<le> pi"
text\<open>NOTE: within this scope, the constant Ln is not yet available!\<close>
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 field_split_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 field_split_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\<^marker>\<open>tag unimportant\<close>\<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\<^marker>\<open>tag unimportant\<close> 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\<^marker>\<open>tag unimportant\<close> 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_Reals_eq: "\<lbrakk>x \<in> \<real>; Re x > 0\<rbrakk> \<Longrightarrow> ln x = of_real (ln (Re x))"
using Ln_of_real by force
lemma Ln_1 [simp]: "ln 1 = (0::complex)"
proof -
have "ln (exp 0) = (0::complex)"
by (simp add: del: exp_zero)
then show ?thesis
by simp
qed
lemma Ln_eq_zero_iff [simp]: "x \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> ln x = 0 \<longleftrightarrow> x = 1" for x::complex
by auto (metis exp_Ln exp_zero nonpos_Reals_zero_I)
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\<^marker>\<open>tag unimportant\<close> 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\<^marker>\<open>tag unimportant\<close> exists_complex_root:
fixes z :: complex
assumes "n \<noteq> 0" obtains w where "z = w ^ n"
proof (cases "z=0")
case False
then show ?thesis
by (rule_tac w = "exp(Ln z / n)" in that) (simp add: assms exp_of_nat_mult [symmetric])
qed (use assms in auto)
corollary\<^marker>\<open>tag unimportant\<close> 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\<^marker>\<open>tag unimportant\<close>\<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 [simp]: "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)
let ?U = "{w. -pi < Im(w) \<and> Im(w) < pi}"
have 1: "open ?U"
by (simp add: open_Collect_conj open_halfspace_Im_gt open_halfspace_Im_lt)
have 2: "\<And>x. x \<in> ?U \<Longrightarrow> (exp has_derivative blinfun_apply(Blinfun ((*) (exp x)))) (at x)"
apply (simp add: bounded_linear_Blinfun_apply bounded_linear_mult_right)
using DERIV_exp has_field_derivative_def by blast
have 3: "continuous_on ?U (\<lambda>x. Blinfun ((*) (exp x)))"
unfolding blinfun_mult_right.abs_eq [symmetric] by (intro continuous_intros)
have 4: "Ln z \<in> ?U"
by (auto simp: mpi_less_Im_Ln *)
have 5: "Blinfun ((*) (inverse z)) o\<^sub>L Blinfun ((*) (exp (Ln z))) = id_blinfun"
by (rule blinfun_eqI) (simp add: bounded_linear_mult_right bounded_linear_Blinfun_apply)
obtain U' V g g' where "open U'" and sub: "U' \<subseteq> ?U"
and "Ln z \<in> U'" "open V" "z \<in> V"
and hom: "homeomorphism U' V exp g"
and g: "\<And>y. y \<in> V \<Longrightarrow> (g has_derivative (g' y)) (at y)"
and g': "\<And>y. y \<in> V \<Longrightarrow> g' y = inv ((*) (exp (g y)))"
and bij: "\<And>y. y \<in> V \<Longrightarrow> bij ((*) (exp (g y)))"
using inverse_function_theorem [OF 1 2 3 4 5]
by (simp add: bounded_linear_Blinfun_apply bounded_linear_mult_right) blast
show "(Ln has_field_derivative inverse(z)) (at z)"
unfolding has_field_derivative_def
proof (rule has_derivative_transform_within_open)
show g_eq_Ln: "g y = Ln y" if "y \<in> V" for y
proof -
obtain x where "y = exp x" "x \<in> U'"
using hom homeomorphism_image1 that \<open>y \<in> V\<close> by blast
then show ?thesis
using sub hom homeomorphism_apply1 by fastforce
qed
have "0 \<notin> V"
by (meson exp_not_eq_zero hom homeomorphism_def)
then have "\<And>y. y \<in> V \<Longrightarrow> g' y = inv ((*) y)"
by (metis exp_Ln g' g_eq_Ln)
then have g': "g' z = (\<lambda>x. x/z)"
by (metis (no_types, hide_lams) bij \<open>z \<in> V\<close> bij_inv_eq_iff exp_Ln g_eq_Ln nonzero_mult_div_cancel_left znz)
show "(g has_derivative (*) (inverse z)) (at z)"
using g [OF \<open>z \<in> V\<close>] g'
by (simp add: \<open>z \<in> V\<close> field_class.field_divide_inverse has_derivative_imp_has_field_derivative has_field_derivative_imp_has_derivative)
qed (auto simp: \<open>z \<in> V\<close> \<open>open V\<close>)
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,continuous_intros]:
"\<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 [continuous_intros]: "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 continuous_on_Ln' [continuous_intros]:
"continuous_on S f \<Longrightarrow> (\<And>z. z \<in> S \<Longrightarrow> f z \<notin> \<real>\<^sub>\<le>\<^sub>0) \<Longrightarrow> continuous_on S (\<lambda>x. Ln (f x))"
by (rule continuous_on_compose2[OF continuous_on_Ln, of "UNIV - nonpos_Reals" S f]) auto
lemma holomorphic_on_Ln [holomorphic_intros]: "(\<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)
lemma holomorphic_on_Ln' [holomorphic_intros]:
"(\<And>z. z \<in> A \<Longrightarrow> f z \<notin> \<real>\<^sub>\<le>\<^sub>0) \<Longrightarrow> f holomorphic_on A \<Longrightarrow> (\<lambda>z. Ln (f z)) holomorphic_on A"
using holomorphic_on_compose_gen[OF _ holomorphic_on_Ln, of f A "- \<real>\<^sub>\<le>\<^sub>0"]
by (auto simp: o_def)
lemma tendsto_Ln [tendsto_intros]:
fixes L F
assumes "(f \<longlongrightarrow> L) F" "L \<notin> \<real>\<^sub>\<le>\<^sub>0"
shows "((\<lambda>x. Ln (f x)) \<longlongrightarrow> Ln L) F"
proof -
have "nhds L \<ge> filtermap f F"
using assms(1) by (simp add: filterlim_def)
moreover have "\<forall>\<^sub>F y in nhds L. y \<in> - \<real>\<^sub>\<le>\<^sub>0"
using eventually_nhds_in_open[of "- \<real>\<^sub>\<le>\<^sub>0" L] assms by (auto simp: open_Compl)
ultimately have "\<forall>\<^sub>F y in filtermap f F. y \<in> - \<real>\<^sub>\<le>\<^sub>0" by (rule filter_leD)
moreover have "continuous_on (-\<real>\<^sub>\<le>\<^sub>0) Ln" by (rule continuous_on_Ln) auto
ultimately show ?thesis using continuous_on_tendsto_compose[of "- \<real>\<^sub>\<le>\<^sub>0" Ln f L F] assms
by (simp add: eventually_filtermap)
qed
lemma divide_ln_mono:
fixes x y::real
assumes "3 \<le> x" "x \<le> y"
shows "x / ln x \<le> y / ln y"
proof (rule exE [OF complex_mvt_line [of x y "\<lambda>z. z / Ln z" "\<lambda>z. 1/(Ln z) - 1/(Ln z)^2"]];
clarsimp simp add: closed_segment_Reals closed_segment_eq_real_ivl assms)
show "\<And>u. \<lbrakk>x \<le> u; u \<le> y\<rbrakk> \<Longrightarrow> ((\<lambda>z. z / Ln z) has_field_derivative 1 / Ln u - 1 / (Ln u)\<^sup>2) (at u)"
using \<open>3 \<le> x\<close> by (force intro!: derivative_eq_intros simp: field_simps power_eq_if)
show "x / ln x \<le> y / ln y"
if "Re (y / Ln y) - Re (x / Ln x) = (Re (1 / Ln u) - Re (1 / (Ln u)\<^sup>2)) * (y - x)"
and x: "x \<le> u" "u \<le> y" for u
proof -
have eq: "y / ln y = (1 / ln u - 1 / (ln u)\<^sup>2) * (y - x) + x / ln x"
using that \<open>3 \<le> x\<close> by (auto simp: Ln_Reals_eq in_Reals_norm group_add_class.diff_eq_eq)
show ?thesis
using exp_le \<open>3 \<le> x\<close> x by (simp add: eq) (simp add: power_eq_if divide_simps ln_ge_iff)
qed
qed
theorem 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
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
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_series': "cmod z < 1 \<Longrightarrow> (\<lambda>n. - ((-z)^n) / of_nat n) sums ln (1 + z)"
by (drule Ln_series) (simp add: power_minus')
lemma ln_series':
assumes "abs (x::real) < 1"
shows "(\<lambda>n. - ((-x)^n) / of_nat n) sums ln (1 + x)"
proof -
from assms have "(\<lambda>n. - ((-of_real x)^n) / of_nat n) sums ln (1 + complex_of_real x)"
by (intro Ln_series') simp_all
also have "(\<lambda>n. - ((-of_real x)^n) / of_nat n) = (\<lambda>n. complex_of_real (- ((-x)^n) / of_nat n))"
by (rule ext) simp
also from assms have "ln (1 + complex_of_real x) = of_real (ln (1 + x))"
by (subst Ln_of_real [symmetric]) simp_all
finally show ?thesis by (subst (asm) sums_of_real_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: field_split_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
subsection\<^marker>\<open>tag unimportant\<close>\<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)"
using cos_lt_zero_pi [of "-(Im w)"] cos_lt_zero_pi [of "(Im w)"]
apply (auto simp: Re_exp zero_less_mult_iff cos_gt_zero_pi abs_if split: if_split_asm)
apply (metis cos_minus cos_pi_half divide_minus_left less_irrefl linorder_neqE_linordered_idom nonzero_mult_div_cancel_right zero_neq_numeral)+
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 (simp add: Im_exp zero_less_mult_iff)
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\<^marker>\<open>tag unimportant\<close>\<open>More Properties of Ln\<close>
lemma cnj_Ln: assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0" shows "cnj(Ln z) = Ln(cnj z)"
proof (cases "z=0")
case False
show ?thesis
proof (rule exp_complex_eqI)
have "\<bar>Im (cnj (Ln z)) - Im (Ln (cnj z))\<bar> \<le> \<bar>Im (cnj (Ln z))\<bar> + \<bar>Im (Ln (cnj z))\<bar>"
by (rule abs_triangle_ineq4)
also have "... < pi + pi"
proof -
have "\<bar>Im (cnj (Ln z))\<bar> < pi"
by (simp add: False Im_Ln_less_pi abs_if assms minus_less_iff mpi_less_Im_Ln)
moreover have "\<bar>Im (Ln (cnj z))\<bar> \<le> pi"
by (meson abs_le_iff complex_cnj_zero_iff less_eq_real_def minus_less_iff False Im_Ln_le_pi mpi_less_Im_Ln)
ultimately show ?thesis
by simp
qed
finally show "\<bar>Im (cnj (Ln z)) - Im (Ln (cnj z))\<bar> < 2 * pi"
by simp
show "exp (cnj (Ln z)) = exp (Ln (cnj z))"
by (metis False complex_cnj_zero_iff exp_Ln exp_cnj)
qed
qed (use assms in auto)
lemma Ln_inverse: assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0" shows "Ln(inverse z) = -(Ln z)"
proof (cases "z=0")
case False
show ?thesis
proof (rule exp_complex_eqI)
have "\<bar>Im (Ln (inverse z)) - Im (- Ln z)\<bar> \<le> \<bar>Im (Ln (inverse z))\<bar> + \<bar>Im (- Ln z)\<bar>"
by (rule abs_triangle_ineq4)
also have "... < pi + pi"
proof -
have "\<bar>Im (Ln (inverse z))\<bar> < pi"
by (simp add: False Im_Ln_less_pi abs_if assms minus_less_iff mpi_less_Im_Ln)
moreover have "\<bar>Im (- Ln z)\<bar> \<le> pi"
using False Im_Ln_le_pi mpi_less_Im_Ln by fastforce
ultimately show ?thesis
by simp
qed
finally show "\<bar>Im (Ln (inverse z)) - Im (- Ln z)\<bar> < 2 * pi"
by simp
show "exp (Ln (inverse z)) = exp (- Ln z)"
by (simp add: False exp_minus)
qed
qed (use assms in auto)
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\<^marker>\<open>tag unimportant\<close> 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\<^marker>\<open>tag unimportant\<close> 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\<^marker>\<open>tag unimportant\<close> Ln_times_Reals:
"\<lbrakk>r \<in> Reals; Re r > 0; z \<noteq> 0\<rbrakk> \<Longrightarrow> Ln(r * z) = ln (Re r) + Ln(z)"
using Ln_Reals_eq Ln_times_of_real by fastforce
corollary\<^marker>\<open>tag unimportant\<close> 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)
corollary\<^marker>\<open>tag unimportant\<close> Ln_prod:
fixes f :: "'a \<Rightarrow> complex"
assumes "finite A" "\<And>x. x \<in> A \<Longrightarrow> f x \<noteq> 0"
shows "\<exists>n. Ln (prod f A) = (\<Sum>x \<in> A. Ln (f x) + (of_int (n x) * (2 * pi)) * \<i>)"
using assms
proof (induction A)
case (insert x A)
then obtain n where n: "Ln (prod f A) = (\<Sum>x\<in>A. Ln (f x) + of_real (of_int (n x) * (2 * pi)) * \<i>)"
by auto
define D where "D \<equiv> Im (Ln (f x)) + Im (Ln (prod f A))"
define q::int where "q \<equiv> (if D \<le> -pi then 1 else if D > pi then -1 else 0)"
have "prod f A \<noteq> 0" "f x \<noteq> 0"
by (auto simp: insert.hyps insert.prems)
with insert.hyps pi_ge_zero show ?case
by (rule_tac x="n(x:=q)" in exI) (force simp: Ln_times q_def D_def n intro!: sum.cong)
qed auto
lemma Ln_minus:
assumes "z \<noteq> 0"
shows "Ln(-z) = (if Im(z) \<le> 0 \<and> \<not>(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"
by (subst Ln_minus) (use assms z in auto)
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 (simp add: Ln_times) auto
lemma Ln_of_nat [simp]: "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>The Argument of a Complex Number\<close>
text\<open>Finally: it's is defined for the same interval as the complex logarithm: \<open>(-\<pi>,\<pi>]\<close>.\<close>
definition\<^marker>\<open>tag important\<close> Arg :: "complex \<Rightarrow> real" where "Arg z \<equiv> (if z = 0 then 0 else Im (Ln z))"
lemma Arg_of_real: "Arg (of_real r) = (if r<0 then pi else 0)"
by (simp add: Im_Ln_eq_pi Arg_def)
lemma mpi_less_Arg: "-pi < Arg z"
and Arg_le_pi: "Arg z \<le> pi"
by (auto simp: Arg_def mpi_less_Im_Ln Im_Ln_le_pi)
lemma
assumes "z \<noteq> 0"
shows Arg_eq: "z = of_real(norm z) * exp(\<i> * Arg z)"
using assms exp_Ln exp_eq_polar
by (auto simp: Arg_def)
lemma is_Arg_Arg: "z \<noteq> 0 \<Longrightarrow> is_Arg z (Arg z)"
by (simp add: Arg_eq is_Arg_def)
lemma Argument_exists:
assumes "z \<noteq> 0" and R: "R = {r-pi<..r+pi}"
obtains s where "is_Arg z s" "s\<in>R"
proof -
let ?rp = "r - Arg z + pi"
define k where "k \<equiv> \<lfloor>?rp / (2 * pi)\<rfloor>"
have "(Arg z + of_int k * (2 * pi)) \<in> R"
using floor_divide_lower [of "2*pi" ?rp] floor_divide_upper [of "2*pi" ?rp]
by (auto simp: k_def algebra_simps R)
then show ?thesis
using Arg_eq \<open>z \<noteq> 0\<close> is_Arg_2pi_iff is_Arg_def that by blast
qed
lemma Argument_exists_unique:
assumes "z \<noteq> 0" and R: "R = {r-pi<..r+pi}"
obtains s where "is_Arg z s" "s\<in>R" "\<And>t. \<lbrakk>is_Arg z t; t\<in>R\<rbrakk> \<Longrightarrow> s=t"
proof -
obtain s where s: "is_Arg z s" "s\<in>R"
using Argument_exists [OF assms] .
moreover have "\<And>t. \<lbrakk>is_Arg z t; t\<in>R\<rbrakk> \<Longrightarrow> s=t"
using assms s by (auto simp: is_Arg_eqI)
ultimately show thesis
using that by blast
qed
lemma Argument_Ex1:
assumes "z \<noteq> 0" and R: "R = {r-pi<..r+pi}"
shows "\<exists>!s. is_Arg z s \<and> s \<in> R"
using Argument_exists_unique [OF assms] by metis
lemma Arg_divide:
assumes "w \<noteq> 0" "z \<noteq> 0"
shows "is_Arg (z / w) (Arg z - Arg w)"
using Arg_eq [of z] Arg_eq [of w] Arg_eq [of "norm(z / w)"] assms
by (auto simp: is_Arg_def norm_divide field_simps exp_diff Arg_of_real)
lemma Arg_unique_lemma:
assumes z: "is_Arg z t"
and z': "is_Arg z t'"
and t: "- pi < t" "t \<le> pi"
and t': "- pi < t'" "t' \<le> pi"
and nz: "z \<noteq> 0"
shows "t' = t"
using Arg2pi_unique_lemma [of "- (inverse z)"]
proof -
have "pi - t' = pi - t"
proof (rule Arg2pi_unique_lemma [of "- (inverse z)"])
have "- (inverse z) = - (inverse (of_real(norm z) * exp(\<i> * t)))"
by (metis is_Arg_def z)
also have "... = (cmod (- inverse z)) * exp (\<i> * (pi - t))"
by (auto simp: field_simps exp_diff norm_divide)
finally show "is_Arg (- inverse z) (pi - t)"
unfolding is_Arg_def .
have "- (inverse z) = - (inverse (of_real(norm z) * exp(\<i> * t')))"
by (metis is_Arg_def z')
also have "... = (cmod (- inverse z)) * exp (\<i> * (pi - t'))"
by (auto simp: field_simps exp_diff norm_divide)
finally show "is_Arg (- inverse z) (pi - t')"
unfolding is_Arg_def .
qed (use assms in auto)
then show ?thesis
by simp
qed
lemma complex_norm_eq_1_exp_eq: "norm z = 1 \<longleftrightarrow> exp(\<i> * (Arg z)) = z"
by (metis Arg_eq exp_not_eq_zero exp_zero mult.left_neutral norm_zero of_real_1 norm_exp_i_times)
lemma Arg_unique: "\<lbrakk>of_real r * exp(\<i> * a) = z; 0 < r; -pi < a; a \<le> pi\<rbrakk> \<Longrightarrow> Arg z = a"
by (rule Arg_unique_lemma [unfolded is_Arg_def, OF _ Arg_eq])
(use mpi_less_Arg Arg_le_pi in \<open>auto simp: norm_mult\<close>)
lemma Arg_minus:
assumes "z \<noteq> 0"
shows "Arg (-z) = (if Arg z \<le> 0 then Arg z + pi else Arg z - pi)"
proof -
have [simp]: "cmod z * cos (Arg z) = Re z"
using assms Arg_eq [of z] by (metis Re_exp exp_Ln norm_exp_eq_Re Arg_def)
have [simp]: "cmod z * sin (Arg z) = Im z"
using assms Arg_eq [of z] by (metis Im_exp exp_Ln norm_exp_eq_Re Arg_def)
show ?thesis
apply (rule Arg_unique [of "norm z"])
apply (rule complex_eqI)
using mpi_less_Arg [of z] Arg_le_pi [of z] assms
apply (auto simp: Re_exp Im_exp cos_diff sin_diff cis_conv_exp [symmetric])
done
qed
lemma Arg_times_of_real [simp]:
assumes "0 < r" shows "Arg (of_real r * z) = Arg z"
proof (cases "z=0")
case True
then show ?thesis
by (simp add: Arg_def)
next
case False
with Arg_eq assms show ?thesis
by (auto simp: mpi_less_Arg Arg_le_pi intro!: Arg_unique [of "r * norm z"])
qed
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_less_0: "0 \<le> Arg z \<longleftrightarrow> 0 \<le> Im z"
using Im_Ln_le_pi Im_Ln_pos_le
by (simp add: Arg_def)
lemma Arg_eq_pi: "Arg z = pi \<longleftrightarrow> Re z < 0 \<and> Im z = 0"
by (auto simp: Arg_def Im_Ln_eq_pi)
lemma Arg_lt_pi: "0 < Arg z \<and> Arg z < pi \<longleftrightarrow> 0 < Im z"
using Arg_less_0 [of z] Im_Ln_pos_lt
by (auto simp: order.order_iff_strict Arg_def)
lemma Arg_eq_0: "Arg z = 0 \<longleftrightarrow> z \<in> \<real> \<and> 0 \<le> Re z"
using complex_is_Real_iff
by (simp add: Arg_def Im_Ln_eq_0) (metis less_eq_real_def of_real_Re of_real_def scale_zero_left)
corollary\<^marker>\<open>tag unimportant\<close> Arg_ne_0: assumes "z \<notin> \<real>\<^sub>\<ge>\<^sub>0" shows "Arg z \<noteq> 0"
using assms by (auto simp: nonneg_Reals_def Arg_eq_0)
lemma Arg_eq_pi_iff: "Arg z = pi \<longleftrightarrow> z \<in> \<real> \<and> Re z < 0"
proof (cases "z=0")
case False
then show ?thesis
using Arg_eq_0 [of "-z"] Arg_eq_pi complex_is_Real_iff by blast
qed (simp add: Arg_def)
lemma Arg_eq_0_pi: "Arg z = 0 \<or> Arg z = pi \<longleftrightarrow> z \<in> \<real>"
using Arg_eq_pi_iff Arg_eq_0 by force
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_inverse: "Arg(inverse z) = (if z \<in> \<real> then Arg z else - Arg z)"
proof (cases "z \<in> \<real>")
case True
then show ?thesis
by simp (metis Arg2pi_inverse Arg2pi_real Arg_real Reals_inverse)
next
case False
then have "Arg z < pi" "z \<noteq> 0"
using Arg_eq_0_pi Arg_le_pi by (auto simp: less_eq_real_def)
then show ?thesis
apply (simp add: False)
apply (rule Arg_unique [of "inverse (norm z)"])
using False mpi_less_Arg [of z] Arg_eq [of z]
apply (auto simp: exp_minus field_simps)
done
qed
lemma Arg_eq_iff:
assumes "w \<noteq> 0" "z \<noteq> 0"
shows "Arg w = Arg z \<longleftrightarrow> (\<exists>x. 0 < x \<and> 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: field_split_simps)
by (metis mult.commute mult.left_commute)
lemma Arg_inverse_eq_0: "Arg(inverse z) = 0 \<longleftrightarrow> Arg z = 0"
by (metis Arg_eq_0 Arg_inverse inverse_inverse_eq)
lemma Arg_cnj_eq_inverse: "z\<noteq>0 \<Longrightarrow> Arg (cnj z) = Arg (inverse z)"
- apply (simp add: Arg_eq_iff field_split_simps complex_norm_square [symmetric] mult.commute)
+ apply (simp add: Arg_eq_iff field_split_simps complex_norm_square [symmetric])
by (metis of_real_power zero_less_norm_iff zero_less_power)
lemma Arg_cnj: "Arg(cnj z) = (if z \<in> \<real> then Arg z else - Arg z)"
by (metis Arg_cnj_eq_inverse Arg_inverse Reals_0 complex_cnj_zero)
lemma Arg_exp: "-pi < Im z \<Longrightarrow> Im z \<le> pi \<Longrightarrow> Arg(exp z) = Im z"
by (rule Arg_unique [of "exp(Re z)"]) (auto simp: exp_eq_polar)
lemma Ln_Arg: "z\<noteq>0 \<Longrightarrow> Ln(z) = ln(norm z) + \<i> * Arg(z)"
by (metis Arg_def Re_Ln complex_eq)
lemma continuous_at_Arg:
assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0"
shows "continuous (at z) Arg"
proof -
have [simp]: "(\<lambda>z. Im (Ln z)) \<midarrow>z\<rightarrow> Arg z"
using Arg_def assms continuous_at by fastforce
show ?thesis
unfolding continuous_at
proof (rule Lim_transform_within_open)
show "\<And>w. \<lbrakk>w \<in> - \<real>\<^sub>\<le>\<^sub>0; w \<noteq> z\<rbrakk> \<Longrightarrow> Im (Ln w) = Arg w"
by (metis Arg_def Compl_iff nonpos_Reals_zero_I)
qed (use assms in auto)
qed
lemma continuous_within_Arg: "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> continuous (at z within S) Arg"
using continuous_at_Arg continuous_at_imp_continuous_within by blast
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. But it's always an integer.\<close>
lemma is_Arg_exp_Im: "is_Arg (exp z) (Im z)"
using exp_eq_polar is_Arg_def norm_exp_eq_Re by auto
lemma is_Arg_exp_diff_2pi:
assumes "is_Arg (exp z) \<theta>"
shows "\<exists>k. Im z - of_int k * (2 * pi) = \<theta>"
proof (intro exI is_Arg_eqI)
let ?k = "\<lfloor>(Im z - \<theta>) / (2 * pi)\<rfloor>"
show "is_Arg (exp z) (Im z - real_of_int ?k * (2 * pi))"
by (metis diff_add_cancel is_Arg_2pi_iff is_Arg_exp_Im)
show "\<bar>Im z - real_of_int ?k * (2 * pi) - \<theta>\<bar> < 2 * pi"
using floor_divide_upper [of "2*pi" "Im z - \<theta>"] floor_divide_lower [of "2*pi" "Im z - \<theta>"]
by (auto simp: algebra_simps abs_if)
qed (auto simp: is_Arg_exp_Im assms)
lemma Arg_exp_diff_2pi: "\<exists>k. Im z - of_int k * (2 * pi) = Arg (exp z)"
using is_Arg_exp_diff_2pi [OF is_Arg_Arg] by auto
lemma unwinding_in_Ints: "(z - Ln(exp z)) / (of_real(2*pi) * \<i>) \<in> \<int>"
using Arg_exp_diff_2pi [of z]
by (force simp: Ints_def image_def field_simps Arg_def intro!: complex_eqI)
definition\<^marker>\<open>tag important\<close> unwinding :: "complex \<Rightarrow> int" where
"unwinding z \<equiv> THE k. of_int k = (z - Ln(exp z)) / (of_real(2*pi) * \<i>)"
lemma unwinding: "of_int (unwinding z) = (z - Ln(exp z)) / (of_real(2*pi) * \<i>)"
using unwinding_in_Ints [of z]
unfolding unwinding_def Ints_def by force
lemma unwinding_2pi: "(2*pi) * \<i> * unwinding(z) = z - Ln(exp z)"
by (simp add: unwinding)
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\<^marker>\<open>tag unimportant\<close>\<open>Relation between Ln and Arg2pi, and hence continuity of Arg2pi\<close>
lemma Arg2pi_Ln:
assumes "0 < Arg2pi z" shows "Arg2pi 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(Arg2pi z))"
using Arg2pi [of z]
by (metis is_Arg_def abs_norm_cancel nonzero_mult_div_cancel_left norm_of_real zero_less_norm_iff)
then have "- z / of_real(norm z) = exp (\<i> * (of_real (Arg2pi 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 (Arg2pi z) - pi)))"
by simp
also have "... = \<i> * (of_real(Arg2pi z) - pi)"
using Arg2pi [of z] assms pi_not_less_zero
by auto
finally have "Arg2pi 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_Arg2pi:
assumes "z \<notin> \<real>\<^sub>\<ge>\<^sub>0"
shows "continuous (at z) Arg2pi"
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]: "Im x \<noteq> 0 \<Longrightarrow> Im (Ln (- x)) + pi = Arg2pi x" for x
using Arg2pi_Ln by (simp add: Arg2pi_gt_0 complex_nonneg_Reals_iff)
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> Arg2pi z"
using "*" by (simp add: Arg2pi_Ln Arg2pi_gt_0 assms continuous_within)
show ?thesis
unfolding continuous_at
proof (rule Lim_transform_within_open)
show "\<And>x. \<lbrakk>x \<in> - \<real>\<^sub>\<ge>\<^sub>0; x \<noteq> z\<rbrakk> \<Longrightarrow> Im (Ln (- x)) + pi = Arg2pi x"
by (auto simp add: Arg2pi_Ln [OF Arg2pi_gt_0] complex_nonneg_Reals_iff)
qed (use assms in auto)
qed
text\<open>Relation between Arg2pi and arctangent in upper halfplane\<close>
lemma Arg2pi_arctan_upperhalf:
assumes "0 < Im z"
shows "Arg2pi z = pi/2 - arctan(Re z / Im z)"
proof (cases "z = 0")
case False
show ?thesis
proof (rule Arg2pi_unique [of "norm z"])
show "(cmod z) * exp (\<i> * (pi / 2 - arctan (Re z / Im z))) = z"
apply (auto simp: exp_Euler cos_diff sin_diff)
using assms 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)
done
qed (use False arctan [of "Re z / Im z"] in auto)
qed (use assms in auto)
lemma Arg2pi_eq_Im_Ln:
assumes "0 \<le> Im z" "0 < Re z"
shows "Arg2pi z = Im (Ln z)"
proof (cases "Im z = 0")
case True then show ?thesis
using Arg2pi_eq_0 Ln_in_Reals assms(2) complex_is_Real_iff by auto
next
case False
then have *: "Arg2pi z > 0"
using Arg2pi_gt_0 complex_is_Real_iff by blast
then have "z \<noteq> 0"
by auto
with * assms False show ?thesis
by (subst Arg2pi_Ln) (auto simp: Ln_minus)
qed
lemma continuous_within_upperhalf_Arg2pi:
assumes "z \<noteq> 0"
shows "continuous (at z within {z. 0 \<le> Im z}) Arg2pi"
proof (cases "z \<in> \<real>\<^sub>\<ge>\<^sub>0")
case False then show ?thesis
using continuous_at_Arg2pi 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]: "Arg2pi z = 0" "Im (Ln z) = 0"
by (auto simp: Arg2pi_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>Arg2pi x\<bar> < e"
apply (rule_tac x="min d (Re z / 2)" in exI)
using z d
apply (auto simp: Arg2pi_eq_Im_Ln)
done
qed
qed
lemma continuous_on_upperhalf_Arg2pi: "continuous_on ({z. 0 \<le> Im z} - {0}) Arg2pi"
unfolding continuous_on_eq_continuous_within
by (metis DiffE Diff_subset continuous_within_subset continuous_within_upperhalf_Arg2pi insertCI)
lemma open_Arg2pi2pi_less_Int:
assumes "0 \<le> s" "t \<le> 2*pi"
shows "open ({y. s < Arg2pi y} \<inter> {y. Arg2pi y < t})"
proof -
have 1: "continuous_on (UNIV - \<real>\<^sub>\<ge>\<^sub>0) Arg2pi"
using continuous_at_Arg2pi 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 < Arg2pi y} \<inter> {y. Arg2pi y < t} \<subseteq> - \<real>\<^sub>\<ge>\<^sub>0"
using assms by (auto simp: Arg2pi_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_Arg2pi2pi_gt: "open {z. t < Arg2pi z}"
proof (cases "t < 0")
case True then have "{z. t < Arg2pi z} = UNIV"
using Arg2pi_ge_0 less_le_trans by auto
then show ?thesis
by simp
next
case False then show ?thesis
using open_Arg2pi2pi_less_Int [of t "2*pi"] Arg2pi_lt_2pi
by auto
qed
lemma closed_Arg2pi2pi_le: "closed {z. Arg2pi z \<le> t}"
using open_Arg2pi2pi_gt [of t]
by (simp add: closed_def Set.Collect_neg_eq [symmetric] not_le)
subsection\<^marker>\<open>tag unimportant\<close>\<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 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 powr_complexpow [simp]:
fixes x::complex shows "x \<noteq> 0 \<Longrightarrow> x powr (of_nat n) = x^n"
by (induct n) (auto simp: ac_simps powr_add)
lemma powr_complexnumeral [simp]:
fixes x::complex shows "x \<noteq> 0 \<Longrightarrow> x powr (numeral n) = x ^ (numeral n)"
by (metis of_nat_numeral powr_complexpow)
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:
assumes "w \<in> \<real>" "z \<in> \<real>" "0 < Re w"
shows "w powr z = exp(Re z * ln(Re w))"
proof -
have "w \<noteq> 0"
using assms by auto
with assms show ?thesis
by (simp add: powr_def Ln_Reals_eq of_real_exp)
qed
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 powr_of_int:
fixes z::complex and n::int
assumes "z\<noteq>(0::complex)"
shows "z powr of_int n = (if n\<ge>0 then z^nat n else inverse (z^nat (-n)))"
by (metis assms not_le of_int_of_nat powr_complexpow powr_minus)
lemma powr_Reals_eq: "\<lbrakk>x \<in> \<real>; y \<in> \<real>; Re x \<ge> 0\<rbrakk> \<Longrightarrow> x powr y = of_real (Re x powr Re y)"
by (metis of_real_Re powr_of_real)
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 Re_powr_le: "r \<in> \<real>\<^sub>\<ge>\<^sub>0 \<Longrightarrow> Re (r powr z) \<le> Re r powr Re z"
by (auto simp: powr_def nonneg_Reals_def order_trans [OF complex_Re_le_cmod])
lemma
fixes w::complex
shows Reals_powr [simp]: "\<lbrakk>w \<in> \<real>\<^sub>\<ge>\<^sub>0; z \<in> \<real>\<rbrakk> \<Longrightarrow> w powr z \<in> \<real>"
and nonneg_Reals_powr [simp]: "\<lbrakk>w \<in> \<real>\<^sub>\<ge>\<^sub>0; z \<in> \<real>\<rbrakk> \<Longrightarrow> w powr z \<in> \<real>\<^sub>\<ge>\<^sub>0"
by (auto simp: nonneg_Reals_def Reals_def powr_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
assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0"
shows "((\<lambda>z. z powr s) has_field_derivative (s * z powr (s - 1))) (at z)"
proof (cases "z=0")
case False
show ?thesis
unfolding powr_def
proof (rule has_field_derivative_transform_within)
show "((\<lambda>z. exp (s * Ln z)) has_field_derivative s * (if z = 0 then 0 else exp ((s - 1) * Ln z)))
(at z)"
apply (intro derivative_eq_intros | simp add: assms)+
by (simp add: False divide_complex_def exp_diff left_diff_distrib')
qed (use False in auto)
qed (use assms in auto)
declare has_field_derivative_powr[THEN DERIV_chain2, derivative_intros]
lemma has_field_derivative_powr_of_int:
fixes z :: complex
assumes gderiv:"(g has_field_derivative gd) (at z within s)" and "g z\<noteq>0"
shows "((\<lambda>z. g z powr of_int n) has_field_derivative (n * g z powr (of_int n - 1) * gd)) (at z within s)"
proof -
define dd where "dd = of_int n * g z powr (of_int (n - 1)) * gd"
obtain e where "e>0" and e_dist:"\<forall>y\<in>s. dist z y < e \<longrightarrow> g y \<noteq> 0"
using DERIV_continuous[OF gderiv,THEN continuous_within_avoid] \<open>g z\<noteq>0\<close> by auto
have ?thesis when "n\<ge>0"
proof -
define dd' where "dd' = of_int n * g z ^ (nat n - 1) * gd"
have "dd=dd'"
proof (cases "n=0")
case False
then have "n-1 \<ge>0" using \<open>n\<ge>0\<close> by auto
then have "g z powr (of_int (n - 1)) = g z ^ (nat n - 1)"
using powr_of_int[OF \<open>g z\<noteq>0\<close>,of "n-1"] by (simp add: nat_diff_distrib')
then show ?thesis unfolding dd_def dd'_def by simp
qed (simp add:dd_def dd'_def)
then have "((\<lambda>z. g z powr of_int n) has_field_derivative dd) (at z within s)
\<longleftrightarrow> ((\<lambda>z. g z powr of_int n) has_field_derivative dd') (at z within s)"
by simp
also have "... \<longleftrightarrow> ((\<lambda>z. g z ^ nat n) has_field_derivative dd') (at z within s)"
proof (rule has_field_derivative_cong_eventually)
show "\<forall>\<^sub>F x in at z within s. g x powr of_int n = g x ^ nat n"
unfolding eventually_at
apply (rule exI[where x=e])
using powr_of_int that \<open>e>0\<close> e_dist by (simp add: dist_commute)
qed (use powr_of_int \<open>g z\<noteq>0\<close> that in simp)
also have "..." unfolding dd'_def using gderiv that
by (auto intro!: derivative_eq_intros)
finally have "((\<lambda>z. g z powr of_int n) has_field_derivative dd) (at z within s)" .
then show ?thesis unfolding dd_def by simp
qed
moreover have ?thesis when "n<0"
proof -
define dd' where "dd' = of_int n / g z ^ (nat (1 - n)) * gd"
have "dd=dd'"
proof -
have "g z powr of_int (n - 1) = inverse (g z ^ nat (1-n))"
using powr_of_int[OF \<open>g z\<noteq>0\<close>,of "n-1"] that by auto
then show ?thesis
unfolding dd_def dd'_def by (simp add: divide_inverse)
qed
then have "((\<lambda>z. g z powr of_int n) has_field_derivative dd) (at z within s)
\<longleftrightarrow> ((\<lambda>z. g z powr of_int n) has_field_derivative dd') (at z within s)"
by simp
also have "... \<longleftrightarrow> ((\<lambda>z. inverse (g z ^ nat (-n))) has_field_derivative dd') (at z within s)"
proof (rule has_field_derivative_cong_eventually)
show "\<forall>\<^sub>F x in at z within s. g x powr of_int n = inverse (g x ^ nat (- n))"
unfolding eventually_at
apply (rule exI[where x=e])
using powr_of_int that \<open>e>0\<close> e_dist by (simp add: dist_commute)
qed (use powr_of_int \<open>g z\<noteq>0\<close> that in simp)
also have "..."
proof -
have "nat (- n) + nat (1 - n) - Suc 0 = nat (- n) + nat (- n)"
by auto
then show ?thesis
unfolding dd'_def using gderiv that \<open>g z\<noteq>0\<close>
by (auto intro!: derivative_eq_intros simp add:field_split_simps power_add[symmetric])
qed
finally have "((\<lambda>z. g z powr of_int n) has_field_derivative dd) (at z within s)" .
then show ?thesis unfolding dd_def by simp
qed
ultimately show ?thesis by force
qed
lemma field_differentiable_powr_of_int:
fixes z :: complex
assumes gderiv:"g field_differentiable (at z within s)" and "g z\<noteq>0"
shows "(\<lambda>z. g z powr of_int n) field_differentiable (at z within s)"
using has_field_derivative_powr_of_int assms(2) field_differentiable_def gderiv by blast
lemma holomorphic_on_powr_of_int [holomorphic_intros]:
assumes "f holomorphic_on s" "\<forall>z\<in>s. f z\<noteq>0"
shows "(\<lambda>z. (f z) powr of_int n) holomorphic_on s"
proof (cases "n\<ge>0")
case True
then have "?thesis \<longleftrightarrow> (\<lambda>z. (f z) ^ nat n) holomorphic_on s"
apply (rule_tac holomorphic_cong)
using assms(2) by (auto simp add:powr_of_int)
moreover have "(\<lambda>z. (f z) ^ nat n) holomorphic_on s"
using assms(1) by (auto intro:holomorphic_intros)
ultimately show ?thesis by auto
next
case False
then have "?thesis \<longleftrightarrow> (\<lambda>z. inverse (f z) ^ nat (-n)) holomorphic_on s"
apply (rule_tac holomorphic_cong)
using assms(2) by (auto simp add:powr_of_int power_inverse)
moreover have "(\<lambda>z. inverse (f z) ^ nat (-n)) holomorphic_on s"
using assms by (auto intro!:holomorphic_intros)
ultimately show ?thesis by auto
qed
lemma has_field_derivative_powr_right [derivative_intros]:
"w \<noteq> 0 \<Longrightarrow> ((\<lambda>z. w powr z) has_field_derivative Ln w * w powr z) (at z)"
unfolding powr_def by (intro derivative_eq_intros | simp)+
lemma field_differentiable_powr_right [derivative_intros]:
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 [holomorphic_intros]:
assumes "f holomorphic_on s"
shows "(\<lambda>z. w powr (f z)) holomorphic_on s"
proof (cases "w = 0")
case False
with assms show ?thesis
unfolding holomorphic_on_def field_differentiable_def
by (metis (full_types) DERIV_chain' has_field_derivative_powr_right)
qed simp
lemma holomorphic_on_divide_gen [holomorphic_intros]:
assumes f: "f holomorphic_on s" and g: "g holomorphic_on s" and 0: "\<And>z z'. \<lbrakk>z \<in> s; z' \<in> s\<rbrakk> \<Longrightarrow> g z = 0 \<longleftrightarrow> g z' = 0"
shows "(\<lambda>z. f z / g z) holomorphic_on s"
proof (cases "\<exists>z\<in>s. g z = 0")
case True
with 0 have "g z = 0" if "z \<in> s" for z
using that by blast
then show ?thesis
using g holomorphic_transform by auto
next
case False
with 0 have "g z \<noteq> 0" if "z \<in> s" for z
using that by blast
with holomorphic_on_divide show ?thesis
using f g by blast
qed
lemma norm_powr_real_powr:
"w \<in> \<real> \<Longrightarrow> 0 \<le> Re w \<Longrightarrow> cmod (w powr z) = Re w powr Re z"
by (metis dual_order.order_iff_strict norm_powr_real norm_zero of_real_0 of_real_Re powr_def)
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 "a \<notin> \<real>\<^sub>\<le>\<^sub>0 \<or> (a = 0 \<and> Re b > 0)" and "(f \<longlongrightarrow> a) F" "(g \<longlongrightarrow> b) F"
shows "((\<lambda>z. f z powr g z) \<longlongrightarrow> a powr b) F"
using assms tendsto_powr_complex tendsto_powr_complex_0 by fastforce
lemma tendsto_neg_powr_complex_of_real:
assumes "filterlim f at_top F" and "Re s < 0"
shows "((\<lambda>x. complex_of_real (f x) powr s) \<longlongrightarrow> 0) F"
proof -
have "((\<lambda>x. norm (complex_of_real (f x) powr s)) \<longlongrightarrow> 0) F"
proof (rule Lim_transform_eventually)
from assms(1) have "eventually (\<lambda>x. f x \<ge> 0) F"
by (auto simp: filterlim_at_top)
thus "eventually (\<lambda>x. f x powr Re s = norm (of_real (f x) powr s)) F"
by eventually_elim (simp add: norm_powr_real_powr)
from assms show "((\<lambda>x. f x powr Re s) \<longlongrightarrow> 0) F"
by (intro tendsto_neg_powr)
qed
thus ?thesis by (simp add: tendsto_norm_zero_iff)
qed
lemma tendsto_neg_powr_complex_of_nat:
assumes "filterlim f at_top F" and "Re s < 0"
shows "((\<lambda>x. of_nat (f x) powr s) \<longlongrightarrow> 0) F"
proof -
have "((\<lambda>x. of_real (real (f x)) powr s) \<longlongrightarrow> 0) F" using assms(2)
by (intro filterlim_compose[OF _ tendsto_neg_powr_complex_of_real]
filterlim_compose[OF _ assms(1)] filterlim_real_sequentially filterlim_ident) auto
thus ?thesis by simp
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\<^marker>\<open>tag unimportant\<close>\<open>Some Limits involving Logarithms\<close>
lemma lim_Ln_over_power:
fixes s::complex
assumes "0 < Re s"
shows "(\<lambda>n. Ln (of_nat n) / of_nat n powr s) \<longlonglongrightarrow> 0"
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"
have "2 / (e * (Re s)\<^sup>2) > 0"
using e assms by simp
with x have "x > 0"
by linarith
then have "x * 2 \<le> e * (x\<^sup>2 * (Re s)\<^sup>2)"
using e assms x by (auto simp: power2_eq_square field_simps)
also have "... < e * (2 + (x * (Re s * 2) + x\<^sup>2 * (Re s)\<^sup>2))"
using e assms \<open>x > 0\<close>
by (auto simp: power2_eq_square field_simps add_pos_pos)
finally show "0 < e * 2 + (e * Re s * 2 - 2) * x + e * (Re s)\<^sup>2 * x\<^sup>2"
by (auto simp: algebra_simps)
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 obtain xo where "xo > 0" and xo: "\<And>x. x \<ge> xo \<Longrightarrow> x < e * exp (Re s * x)"
using e by (auto simp: field_simps)
have "norm (Ln (of_nat n) / of_nat n powr s) < e" if "n \<ge> nat \<lceil>exp xo\<rceil>" for n
using e xo [of "ln n"] that
apply (auto simp: norm_divide norm_powr_real field_split_simps)
apply (metis exp_less_mono exp_ln not_le of_nat_0_less_iff)
done
then show "\<exists>no. \<forall>n\<ge>no. norm (Ln (of_nat n) / of_nat n powr s) < e"
by blast
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 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 [tendsto_intros]: "((\<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_log_over_n [tendsto_intros]:
"(\<lambda>n. log k n/n) \<longlonglongrightarrow> 0"
proof -
have *: "log k n/n = (1/ln k) * (ln n / n)" for n
unfolding log_def by auto
have "(\<lambda>n. (1/ln k) * (ln n / n)) \<longlonglongrightarrow> (1/ln k) * 0"
by (intro tendsto_intros)
then show ?thesis
unfolding * by auto
qed
lemma lim_1_over_complex_power:
assumes "0 < Re s"
shows "(\<lambda>n. 1 / of_nat n powr s) \<longlonglongrightarrow> 0"
proof (rule Lim_null_comparison)
have "\<forall>n>0. 3 \<le> n \<longrightarrow> 1 \<le> ln (real_of_nat n)"
using ln_272_gt_1
by (force intro: order_trans [of _ "ln (272/100)"])
then show "\<forall>\<^sub>F x in sequentially. cmod (1 / of_nat x powr s) \<le> cmod (Ln (of_nat x) / of_nat x powr s)"
by (auto simp: norm_divide field_split_simps eventually_sequentially)
show "(\<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)
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 field_split_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: field_split_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>k. Ln (of_nat k) \<noteq> 0 \<longrightarrow> no \<le> k \<longrightarrow> 1 < r * cmod (Ln (of_nat k))"
using n \<open>0 < r\<close>
by (rule_tac x=n in exI) (force simp: field_split_simps intro: less_le_trans)
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
lemma lim_ln1_over_ln: "(\<lambda>n. ln(Suc n) / ln n) \<longlonglongrightarrow> 1"
proof (rule Lim_transform_eventually)
have "(\<lambda>n. ln(1 + 1/n) / ln n) \<longlonglongrightarrow> 0"
proof (rule Lim_transform_bound)
show "(inverse o real) \<longlonglongrightarrow> 0"
by (metis comp_def lim_inverse_n lim_explicit)
show "\<forall>\<^sub>F n in sequentially. norm (ln (1 + 1 / n) / ln n) \<le> norm ((inverse \<circ> real) n)"
proof
fix n::nat
assume n: "3 \<le> n"
then have "ln 3 \<le> ln n" and ln0: "0 \<le> ln n"
by auto
with ln3_gt_1 have "1/ ln n \<le> 1"
by (simp add: field_split_simps)
moreover have "ln (1 + 1 / real n) \<le> 1/n"
by (simp add: ln_add_one_self_le_self)
ultimately have "ln (1 + 1 / real n) * (1 / ln n) \<le> (1/n) * 1"
by (intro mult_mono) (use n in auto)
then show "norm (ln (1 + 1 / n) / ln n) \<le> norm ((inverse \<circ> real) n)"
by (simp add: field_simps ln0)
qed
qed
then show "(\<lambda>n. 1 + ln(1 + 1/n) / ln n) \<longlonglongrightarrow> 1"
by (metis (full_types) add.right_neutral tendsto_add_const_iff)
show "\<forall>\<^sub>F k in sequentially. 1 + ln (1 + 1 / k) / ln k = ln(Suc k) / ln k"
by (simp add: field_split_simps ln_div eventually_sequentiallyI [of 2])
qed
lemma lim_ln_over_ln1: "(\<lambda>n. ln n / ln(Suc n)) \<longlonglongrightarrow> 1"
proof -
have "(\<lambda>n. inverse (ln(Suc n) / ln n)) \<longlonglongrightarrow> inverse 1"
by (rule tendsto_inverse [OF lim_ln1_over_ln]) auto
then show ?thesis
by simp
qed
subsection\<^marker>\<open>tag unimportant\<close>\<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_div_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")
case False
then show ?thesis
using assms csqrt_exp_Ln Ln_inverse exp_minus
by (simp add: csqrt_exp_Ln Ln_inverse exp_minus)
qed auto
lemma cnj_csqrt:
assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0"
shows "cnj(csqrt z) = csqrt(cnj z)"
proof (cases "z=0")
case False
then show ?thesis
by (simp add: assms cnj_Ln csqrt_exp_Ln exp_cnj)
qed auto
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: field_split_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
proof (rule has_field_derivative_transform_within)
show "\<And>x. dist x z < cmod z \<Longrightarrow> exp (Ln x / 2) = csqrt x"
by (metis csqrt_exp_Ln dist_0_norm less_irrefl)
qed (use z in auto)
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\<^marker>\<open>tag unimportant\<close> 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
have *: "\<And>e. \<lbrakk>0 < e\<rbrakk>
\<Longrightarrow> \<forall>x'\<in>\<real> \<inter> {w. 0 \<le> Re w}. cmod x' < e^2 \<longrightarrow> cmod (csqrt x') < e"
by (auto simp: Reals_def real_less_lsqrt)
have "Im z = 0" "Re z < 0 \<or> z = 0"
using True cnj.code complex_cnj_zero_iff by (auto simp: Complex_eq complex_nonpos_Reals_iff) fastforce
with * show ?thesis
apply (auto simp: continuous_within_closed_nontrivial [OF closed_Real_halfspace_Re_ge])
apply (auto simp: continuous_within_eps_delta)
using zero_less_power by blast
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\<^marker>\<open>tag important\<close> 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
i_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 add_diff_cancel_left' complex_i_mult_minus diff_0 i_squared imaginary_unit.simps
less_asym neg_equal_iff_equal)
have "z \<noteq> -\<i>" using assms
by auto
then have zz: "1 + z * z \<noteq> 0"
by (metis abs_one assms i_squared imaginary_unit.simps less_irrefl minus_unique square_eq_iff)
have nz1: "1 - \<i>*z \<noteq> 0"
using assms by (force simp add: i_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 (intro derivative_eq_intros | simp add: nz0 *)+
using nz0 nz1 zz
- apply (simp add: algebra_simps field_split_simps power2_eq_square)
+ apply (simp add: field_split_simps power2_eq_square)
apply algebra
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)
theorem 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 (((2*Suc n+1) / (Suc n)) /
((2*Suc n-1) / (Suc n)))"
by (simp add: h_def norm_mult norm_power norm_divide field_split_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: field_split_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: field_split_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: field_split_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 strict_mono_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)
+ hence u: "norm u < 1" by (simp)
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: strict_mono_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: 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
from this[of 0] have "c = 0" by (simp add: G_def g_def)
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: strict_mono_def power_mult g_def h_def)
qed
text \<open>A quickly-converging series for the logarithm, based on the arctangent.\<close>
theorem 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\<^marker>\<open>tag unimportant\<close> \<open>Real arctangent\<close>
lemma Im_Arctan_of_real [simp]: "Im (Arctan (of_real x)) = 0"
proof -
have ne: "1 + x\<^sup>2 \<noteq> 0"
by (metis power_one sum_power2_eq_zero_iff zero_neq_one)
have "Re (Ln ((1 - \<i> * x) * inverse (1 + \<i> * x))) = 0"
apply (rule norm_exp_imaginary)
apply (subst exp_Ln)
using ne apply (simp_all add: cmod_def complex_eq_iff)
apply (auto simp: field_split_simps)
apply algebra
done
then show ?thesis
unfolding Arctan_def divide_complex_def by (simp add: complex_eq_iff)
qed
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: field_split_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: field_split_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 -
have 1: "\<And>x. x \<in> \<real> \<Longrightarrow> cmod (inverse (1 + x\<^sup>2)) \<le> 1"
by (simp add: norm_divide divide_simps in_Reals_norm complex_is_Real_iff power2_eq_square)
have "cmod (Arctan w - Arctan z) \<le> 1 * cmod (w-z)" if "w \<in> \<real>" "z \<in> \<real>" for 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])
using 1 that apply (auto simp: Reals_def)
done
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"
proof (rule tendsto_eq_rhs)
show "(\<lambda>k. 1 / real (k * 2 + 1) * x ^ (k * 2 + 1)) \<longlonglongrightarrow> 0 * 0"
using assms
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)
qed simp
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: field_split_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\<^marker>\<open>tag unimportant\<close> \<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)
lemma pi_gt3: "pi > 3"
using pi_approx by simp
subsection\<open>Inverse Sine\<close>
definition\<^marker>\<open>tag important\<close> 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 leim: "(cmod (1 - z\<^sup>2) + (1 - Re (z\<^sup>2))) / 2 \<le> (Im z)\<^sup>2"
using le0 sqrt_le_D by fastforce
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 leim cmod_power2 [of z] norm_triangle_ineq2 [of "z^2" 1]
by (simp add: norm_power Re_power2 norm_minus_commute [of 1])
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 1: "\<i> * z + csqrt (1 - z\<^sup>2) \<notin> \<real>\<^sub>\<le>\<^sub>0"
by (metis isCont_Arcsin_lemma assms complex_nonpos_Reals_iff)
have 2: "1 - z\<^sup>2 \<notin> \<real>\<^sub>\<le>\<^sub>0"
by (simp add: one_minus_z2_notin_nonpos_Reals assms)
show ?thesis
using assms unfolding Arcsin_def by (intro isCont_Ln' isCont_csqrt' continuous_intros 1 2)
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_i_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_i_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"
using assms by (auto simp: abs_if simp del: eq_divide_eq_numeral1 split: if_split_asm)
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 one_minus_z2_notin_nonpos_Reals by force
then have "cos (Arcsin z) \<noteq> 0"
by (metis diff_0_right power_zero_numeral sin_squared_eq)
then show ?thesis
by (rule has_field_derivative_inverse_basic [OF DERIV_sin _ _ open_ball [of z 1]]) (auto intro: isCont_Arcsin assms)
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\<^marker>\<open>tag important\<close> 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 Arcsin_body_lemma complex_i_mult_minus diff_minus_eq_add power2_minus right_minus_eq)
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
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 "x\<^sup>2 \<noteq> -1" for x::real
by (sos "((R<1 + (([~1] * A=0) + (R<1 * (R<1 * [x__]^2)))))")
with assms have "(cos (Arccos z))\<^sup>2 \<noteq> 1"
by (auto simp: complex_eq_iff Re_power2 Im_power2 abs_square_eq_1)
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)"
by (rule has_field_derivative_inverse_basic [OF DERIV_cos _ _ open_ball [of z 1]])
(auto intro: isCont_Arccos assms)
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\<^marker>\<open>tag unimportant\<close>\<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 Im_Arccos_bound: "\<bar>Im (Arccos w)\<bar> \<le> cmod w"
proof -
have "(Im (Arccos w))\<^sup>2 \<le> (cmod (cos (Arccos w)))\<^sup>2 - (cos (Re (Arccos w)))\<^sup>2"
using norm_cos_squared [of "Arccos w"] real_le_abs_sinh [of "Im (Arccos w)"]
apply (simp only: abs_le_square_iff)
apply (simp add: field_split_simps)
done
also have "... \<le> (cmod w)\<^sup>2"
by (auto simp: cmod_power2)
finally show ?thesis
using abs_le_square_iff by force
qed
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)
lemma norm_Arccos_bounded:
fixes w :: complex
shows "norm (Arccos w) \<le> pi + norm w"
proof -
have Re: "(Re (Arccos w))\<^sup>2 \<le> pi\<^sup>2" "(Im (Arccos w))\<^sup>2 \<le> (cmod w)\<^sup>2"
using Re_Arccos_bound [of w] Im_Arccos_bound [of w] abs_le_square_iff by force+
have "Arccos w \<bullet> Arccos w \<le> pi\<^sup>2 + (cmod w)\<^sup>2"
using Re by (simp add: dot_square_norm cmod_power2 [of "Arccos w"])
then have "cmod (Arccos w) \<le> pi + cmod (cos (Arccos w))"
apply (simp add: norm_le_square)
by (metis dot_square_norm norm_ge_zero norm_le_square pi_ge_zero triangle_lemma)
then show "cmod (Arccos w) \<le> pi + cmod w"
by auto
qed
subsection\<^marker>\<open>tag unimportant\<close>\<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)
+ by (simp add: 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)
+ by (simp add: 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\<^marker>\<open>tag unimportant\<close>\<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\<^marker>\<open>tag unimportant\<close> 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\<^marker>\<open>tag unimportant\<close>\<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\<^marker>\<open>tag unimportant\<close> 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\<^marker>\<open>tag unimportant\<close>\<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\<^marker>\<open>tag unimportant\<close>\<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
lemma sinh_ln_complex: "x \<noteq> 0 \<Longrightarrow> sinh (ln x :: complex) = (x - inverse x) / 2"
by (simp add: sinh_def exp_minus scaleR_conv_of_real exp_of_real)
lemma cosh_ln_complex: "x \<noteq> 0 \<Longrightarrow> cosh (ln x :: complex) = (x + inverse x) / 2"
by (simp add: cosh_def exp_minus scaleR_conv_of_real)
lemma tanh_ln_complex: "x \<noteq> 0 \<Longrightarrow> tanh (ln x :: complex) = (x ^ 2 - 1) / (x ^ 2 + 1)"
by (simp add: tanh_def sinh_ln_complex cosh_ln_complex divide_simps power2_eq_square)
subsection\<open>Roots of unity\<close>
theorem 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)
+ by (simp)
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: mod_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 field_split_simps mult_ac of_real_numeral *)
+ by (simp add: exp_eq field_split_simps *)
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
diff --git a/src/HOL/Analysis/Convex_Euclidean_Space.thy b/src/HOL/Analysis/Convex_Euclidean_Space.thy
--- a/src/HOL/Analysis/Convex_Euclidean_Space.thy
+++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy
@@ -1,2274 +1,2274 @@
(* Title: HOL/Analysis/Convex_Euclidean_Space.thy
Author: L C Paulson, University of Cambridge
Author: Robert Himmelmann, TU Muenchen
Author: Bogdan Grechuk, University of Edinburgh
Author: Armin Heller, TU Muenchen
Author: Johannes Hoelzl, TU Muenchen
*)
section \<open>Convex Sets and Functions on (Normed) Euclidean Spaces\<close>
theory Convex_Euclidean_Space
imports
Convex
Topology_Euclidean_Space
begin
subsection\<^marker>\<open>tag unimportant\<close> \<open>Topological Properties of Convex Sets and Functions\<close>
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 "(+) 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>Relative interior of a set\<close>
definition\<^marker>\<open>tag important\<close> "rel_interior S =
{x. \<exists>T. openin (top_of_set (affine hull S)) T \<and> x \<in> T \<and> T \<subseteq> S}"
lemma rel_interior_mono:
"\<lbrakk>S \<subseteq> T; affine hull S = affine hull T\<rbrakk>
\<Longrightarrow> (rel_interior S) \<subseteq> (rel_interior T)"
by (auto simp: rel_interior_def)
lemma rel_interior_maximal:
"\<lbrakk>T \<subseteq> S; openin(top_of_set (affine hull S)) T\<rbrakk> \<Longrightarrow> T \<subseteq> (rel_interior S)"
by (auto simp: rel_interior_def)
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: 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: open_contains_ball)
apply (rule_tac x = "ball x e" in exI, 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: open_contains_cball)
apply (rule_tac x = "ball x e" in exI)
apply (simp add: subset_trans [OF ball_subset_cball], 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: 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, 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: rel_interior_def)
lemma rel_interior_subset: "rel_interior S \<subseteq> S"
by (auto simp: rel_interior_def)
lemma rel_interior_subset_closure: "rel_interior S \<subseteq> closure S"
using rel_interior_subset by (auto simp: closure_def)
lemma interior_subset_rel_interior: "interior S \<subseteq> rel_interior S"
by (auto simp: 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_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: 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: 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: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[THEN subsetD])
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_atLeast [simp]:
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: 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: 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_atMost [simp]:
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: 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: 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 [simp]: "interior {a..b} = {a<..<b :: real}"
proof-
have "{a..b} = {a..} \<inter> {..b}" by auto
also have "interior \<dots> = {a<..} \<inter> {..<b}"
- by (simp add: interior_real_atLeast interior_real_atMost)
+ by (simp)
also have "\<dots> = {a<..<b}" by auto
finally show ?thesis .
qed
lemma interior_atLeastLessThan [simp]:
fixes a::real shows "interior {a..<b} = {a<..<b}"
by (metis atLeastLessThan_def greaterThanLessThan_def interior_atLeastAtMost_real interior_Int interior_interior interior_real_atLeast)
lemma interior_lessThanAtMost [simp]:
fixes a::real shows "interior {a<..b} = {a<..<b}"
by (metis atLeastAtMost_def greaterThanAtMost_def interior_atLeastAtMost_real interior_Int
interior_interior interior_real_atLeast)
lemma interior_greaterThanLessThan_real [simp]: "interior {a<..<b} = {a<..<b :: real}"
by (metis interior_atLeastAtMost_real interior_interior)
lemma frontier_real_atMost [simp]:
fixes a :: real
shows "frontier {..a} = {a}"
unfolding frontier_def by (auto simp: interior_real_atMost)
lemma frontier_real_atLeast [simp]: "frontier {a..} = {a::real}"
by (auto simp: frontier_def)
lemma frontier_real_greaterThan [simp]: "frontier {a<..} = {a::real}"
by (auto simp: interior_open frontier_def)
lemma frontier_real_lessThan [simp]: "frontier {..<a} = {a::real}"
by (auto simp: interior_open frontier_def)
lemma rel_interior_real_box [simp]:
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 [simp]:
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_atLeast interior_rel_interior_gen[of "{a..}"]
by (auto split: if_split_asm)
qed
subsubsection \<open>Relative open sets\<close>
definition\<^marker>\<open>tag important\<close> "rel_open S \<longleftrightarrow> rel_interior S = S"
lemma rel_open: "rel_open S \<longleftrightarrow> openin (top_of_set (affine hull S)) S"
unfolding rel_open_def rel_interior_def
apply auto
using openin_subopen[of "top_of_set (affine hull S)" S]
apply auto
done
lemma openin_rel_interior: "openin (top_of_set (affine hull S)) (rel_interior S)"
apply (simp add: rel_interior_def)
apply (subst openin_subopen, blast)
done
lemma openin_set_rel_interior:
"openin (top_of_set 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 = ((+) 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 closed_affine_hull [iff]:
fixes S :: "'n::euclidean_space set"
shows "closed (affine hull S)"
by (metis affine_affine_hull affine_closed)
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], 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: 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 (simp add: norm_minus_commute) (simp add: field_simps)
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 simp
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(top_of_set (affine hull s)) s"
using rel_open rel_open_def by blast
lemma rel_interior_openin:
"openin(top_of_set (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\<^marker>\<open>tag unimportant\<close>\<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 ((+) 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"
proof -
interpret f: bounded_linear f by fact
have "affine {x. f x \<in> affine hull f ` s}"
unfolding affine_def
by (auto simp: f.scaleR f.add affine_affine_hull[unfolded affine_def, rule_format])
moreover have "affine {x. x \<in> f ` (affine hull s)}"
using affine_affine_hull[unfolded affine_def, of s]
unfolding affine_def by (auto simp: f.scaleR [symmetric] f.add [symmetric])
ultimately show ?thesis
by (auto simp: hull_inc elim!: hull_induct)
qed
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_superset 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_superset
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\<^marker>\<open>tag unimportant\<close> \<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)"
proof (clarsimp simp: open_contains_cball convex_hull_explicit)
fix T and u :: "'a\<Rightarrow>real"
assume obt: "finite T" "T\<subseteq>S" "\<forall>x\<in>T. 0 \<le> u x" "sum u T = 1"
from assms[unfolded open_contains_cball] obtain b
where b: "\<And>x. x\<in>S \<Longrightarrow> 0 < b x \<and> cball x (b x) \<subseteq> S" by metis
have "b ` T \<noteq> {}"
using obt by auto
define i where "i = b ` T"
let ?\<Phi> = "\<lambda>y. \<exists>F. finite F \<and> F \<subseteq> S \<and> (\<exists>u. (\<forall>x\<in>F. 0 \<le> u x) \<and> sum u F = 1 \<and> (\<Sum>v\<in>F. u v *\<^sub>R v) = y)"
let ?a = "\<Sum>v\<in>T. u v *\<^sub>R v"
show "\<exists>e > 0. cball ?a e \<subseteq> {y. ?\<Phi> y}"
proof (intro exI subsetI conjI)
show "0 < Min i"
unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] \<open>b ` T\<noteq>{}\<close>]
using b \<open>T\<subseteq>S\<close> by auto
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"
by (simp add: i_def obt(1))
then have "x + (y - ?a) \<in> cball x (b x)"
using y unfolding mem_cball dist_norm by auto
moreover have "x \<in> S"
using \<open>x\<in>T\<close> \<open>T\<subseteq>S\<close> by auto
ultimately have "x + (y - ?a) \<in> S"
using y b by blast
}
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)) *\<^sub>R v) = y"
unfolding sum.reindex[OF *] o_def using obt(4)
by (simp add: sum.distrib sum_subtractf scaleR_left.sum[symmetric] scaleR_right_distrib)
ultimately show "y \<in> {y. ?\<Phi> y}"
proof (intro CollectI exI conjI)
show "finite ((\<lambda>v. v + (y - ?a)) ` T)"
by (simp add: obt(1))
show "sum (\<lambda>v. u (v - (y - ?a))) ((\<lambda>v. v + (y - ?a)) ` T) = 1"
unfolding sum.reindex[OF *] o_def using obt(4) by auto
qed (use obt(1, 3) in auto)
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"
by force
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)
with assms show ?thesis
by (simp add: * compact_Times compact_continuous_image)
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, 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: 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, 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\<^marker>\<open>tag unimportant\<close> \<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 (intro impI ballI, 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 field_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: 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>]
proof (intro CollectI conjI exI)
show "u + w \<ge> 0" "v - w \<ge> 0"
using obt(1) w by auto
qed (use obt in auto)
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>]
proof (intro CollectI conjI exI)
show "u - w \<ge> 0" "v + w \<ge> 0"
using obt(1) w by auto
qed (use obt in auto)
ultimately show ?thesis by auto
qed
qed auto
qed
qed auto
qed (auto simp: 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 \<open>finite S\<close>], 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: 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\<^marker>\<open>tag important\<close> 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_in_set: "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_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
have "norm (v *\<^sub>R z - y) < norm y"
if "0 < v" and "v \<le> inner y z / inner z z" for v
unfolding norm_lt using z assms that
by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ \<open>0<v\<close>])
then show ?thesis
using assms z
by (rule_tac x = "inner y z / inner z z" in exI) 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: 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: 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: 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 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 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])
proposition closest_point_in_rel_interior:
assumes "closed S" "S \<noteq> {}" and x: "x \<in> affine hull S"
shows "closest_point S x \<in> rel_interior S \<longleftrightarrow> x \<in> rel_interior S"
proof (cases "x \<in> S")
case True
then show ?thesis
by (simp add: closest_point_self)
next
case False
then have "False" if asm: "closest_point S x \<in> rel_interior S"
proof -
obtain e where "e > 0" and clox: "closest_point S x \<in> S"
and e: "cball (closest_point S x) e \<inter> affine hull S \<subseteq> S"
using asm mem_rel_interior_cball by blast
then have clo_notx: "closest_point S x \<noteq> x"
using \<open>x \<notin> S\<close> by auto
define y where "y \<equiv> closest_point S x -
(min 1 (e / norm(closest_point S x - x))) *\<^sub>R (closest_point S x - x)"
have "x - y = (1 - min 1 (e / norm (closest_point S x - x))) *\<^sub>R (x - closest_point S x)"
by (simp add: y_def algebra_simps)
then have "norm (x - y) = abs ((1 - min 1 (e / norm (closest_point S x - x)))) * norm(x - closest_point S x)"
by simp
also have "\<dots> < norm(x - closest_point S x)"
using clo_notx \<open>e > 0\<close>
by (auto simp: mult_less_cancel_right2 field_split_simps)
finally have no_less: "norm (x - y) < norm (x - closest_point S x)" .
have "y \<in> affine hull S"
unfolding y_def
by (meson affine_affine_hull clox hull_subset mem_affine_3_minus2 subsetD x)
moreover have "dist (closest_point S x) y \<le> e"
using \<open>e > 0\<close> by (auto simp: y_def min_mult_distrib_right)
ultimately have "y \<in> S"
using subsetD [OF e] by simp
then have "dist x (closest_point S x) \<le> dist x y"
by (simp add: closest_point_le \<open>closed S\<close>)
with no_less show False
by (simp add: dist_norm)
qed
moreover have "x \<notin> rel_interior S"
using rel_interior_subset False by blast
ultimately show ?thesis by blast
qed
subsubsection\<^marker>\<open>tag unimportant\<close> \<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
proof (intro exI bexI conjI ballI)
show "(y - z) \<bullet> z < (y - z) \<bullet> y"
by (metis \<open>y \<in> S\<close> assms(4) diff_gt_0_iff_gt inner_commute inner_diff_left inner_gt_zero_iff right_minus_eq)
show "(y - z) \<bullet> y \<le> (y - z) \<bullet> x" if "x \<in> S" for x
proof (rule ccontr)
have *: "\<And>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> (y - z) \<bullet> y \<le> (y - z) \<bullet> 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: inner_diff)
then show False
using *[of v] by (auto simp: dist_commute algebra_simps)
qed
qed (use \<open>y \<in> S\<close> in 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
by (simp add: gt_ex)
next
case False
obtain y where "y \<in> S" and y: "\<And>x. x \<in> S \<Longrightarrow> dist z y \<le> dist z x"
by (metis distance_attains_inf[OF assms(2) False])
show ?thesis
proof (intro exI conjI ballI)
show "(y - z) \<bullet> z < inner (y - z) z + (norm (y - z))\<^sup>2 / 2"
using \<open>y\<in>S\<close> \<open>z\<notin>S\<close> by auto
next
fix x
assume "x \<in> S"
have "False" if *: "0 < inner (z - y) (x - y)"
proof -
obtain u where "u > 0" "u \<le> 1" "dist (y + u *\<^sub>R (x - y)) z < dist y z"
using * closer_point_lemma by blast
then show False using y[of "y + u *\<^sub>R (x - y)"] convexD_alt [OF \<open>convex S\<close>]
using \<open>x\<in>S\<close> \<open>y\<in>S\<close> by (auto simp: 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 "(y - z) \<bullet> z + (norm (y - z))\<^sup>2 / 2 < (y - z) \<bullet> x"
by (force simp: field_simps power2_norm_eq_inner inner_commute inner_diff)
qed
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 "(SOME i. i\<in>Basis) \<noteq> (0::'a)"
by (metis Basis_zero SOME_Basis)
then show ?thesis
using True zero_less_one by blast
next
case False
then show ?thesis
using False using separating_hyperplane_closed_point[OF assms]
by (metis all_not_in_conv inner_zero_left inner_zero_right less_eq_real_def not_le)
qed
subsubsection\<^marker>\<open>tag unimportant\<close> \<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> (\<Union>x\<in> S. \<Union>y \<in> T. {x - y}). 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
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: inner_diff)
done
define k where "k = (SUP x\<in>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 ((\<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, auto)
done
qed
subsubsection\<^marker>\<open>tag unimportant\<close> \<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: 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 force
then have *: "\<And>x y. x \<in> T \<Longrightarrow> y \<in> S \<Longrightarrow> inner a y \<le> inner a x"
by (force simp: inner_diff)
then have bdd: "bdd_above (((\<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\<in>S. a \<bullet> x"])
(auto intro!: cSUP_upper bdd cSUP_least \<open>a \<noteq> 0\<close> \<open>S \<noteq> {}\<close> *)
qed
subsection\<^marker>\<open>tag unimportant\<close> \<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
proof clarify
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"
proof (intro exI conjI subsetI)
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: algebra_simps)
done
then show "z \<in> S"
using u by (auto simp: 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\<^marker>\<open>tag unimportant\<close> \<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, 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, auto)
done
qed auto
subsection\<^marker>\<open>tag unimportant\<close> \<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: 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: field_simps)
done
then have "a \<le> u * a + v * b"
unfolding * using as(4)
by (auto simp: 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: field_simps)
then have "a < b"
unfolding * using as(4)
apply (rule_tac mult_left_less_imp_less)
apply (auto simp: field_simps)
done
then have "u * a + v * b \<le> b"
unfolding **
using **(2) as(3)
by (auto simp: 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
text\<open>A non-singleton connected set is perfect (i.e. has no isolated points). \<close>
lemma connected_imp_perfect:
fixes a :: "'a::metric_space"
assumes "connected S" "a \<in> S" and S: "\<And>x. S \<noteq> {x}"
shows "a islimpt S"
proof -
have False if "a \<in> T" "open T" "\<And>y. \<lbrakk>y \<in> S; y \<in> T\<rbrakk> \<Longrightarrow> y = a" for T
proof -
obtain e where "e > 0" and e: "cball a e \<subseteq> T"
using \<open>open T\<close> \<open>a \<in> T\<close> by (auto simp: open_contains_cball)
have "openin (top_of_set S) {a}"
unfolding openin_open using that \<open>a \<in> S\<close> by blast
moreover have "closedin (top_of_set S) {a}"
by (simp add: assms)
ultimately show "False"
using \<open>connected S\<close> connected_clopen S by blast
qed
then show ?thesis
unfolding islimpt_def by blast
qed
lemma connected_imp_perfect_aff_dim:
"\<lbrakk>connected S; aff_dim S \<noteq> 0; a \<in> S\<rbrakk> \<Longrightarrow> a islimpt S"
using aff_dim_sing connected_imp_perfect by blast
subsection\<^marker>\<open>tag unimportant\<close> \<open>On \<open>real\<close>, \<open>is_interval\<close>, \<open>convex\<close> and \<open>connected\<close> are all equivalent\<close>
lemma mem_is_interval_1_I:
fixes a b c::real
assumes "is_interval S"
assumes "a \<in> S" "c \<in> S"
assumes "a \<le> b" "b \<le> c"
shows "b \<in> S"
using assms is_interval_1 by blast
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, rule)
apply (rule open_lessThan, rule)
apply (rule open_greaterThan, 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_compact_interval_1:
"connected S \<and> compact S \<longleftrightarrow> (\<exists>a b. S = {a..b::real})"
by (auto simp: is_interval_connected_1 [symmetric] is_interval_compact)
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[OF subspace_UNIV subspace_UNIV,
where 'a='a and 'b=real]
unfolding Euclidean_Space.dim_UNIV
by (auto simp: 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
lemma [simp]:
fixes r s::real
shows is_interval_io: "is_interval {..<r}"
and is_interval_oi: "is_interval {r<..}"
and is_interval_oo: "is_interval {r<..<s}"
and is_interval_oc: "is_interval {r<..s}"
and is_interval_co: "is_interval {r..<s}"
by (simp_all add: is_interval_convex_1)
subsection\<^marker>\<open>tag unimportant\<close> \<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: 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: 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\<^marker>\<open>tag unimportant\<close> \<open>A bound within an interval\<close>
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_sum_left sum.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_sum_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 add: convex_hull_linear_image)
also have "\<dots> = convex hull (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` {0, 1})"
by (simp only: convex_hull_set_sum)
also have "\<dots> = convex hull {x. \<forall>i\<in>Basis. x\<bullet>i \<in> {0, 1}}"
by (simp only: box_eq_set_sum_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"
proof
show "finite {x::'a. \<forall>i\<in>Basis. x \<bullet> i = 0 \<or> x \<bullet> i = 1}"
proof (rule finite_subset, clarify)
show "finite ((\<lambda>S. \<Sum>i\<in>Basis. (if i \<in> S then 1 else 0) *\<^sub>R i) ` Pow Basis)"
using finite_Basis by blast
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, auto)
done
qed
show "cbox 0 One = convex hull {x. \<forall>i\<in>Basis. x \<bullet> i = 0 \<or> x \<bullet> i = 1}"
using unit_interval_convex_hull by blast
qed
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>R i)::'a"
have *: "cbox (x - ?d) (x + ?d) = (\<lambda>y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 (\<Sum>Basis)"
proof (intro set_eqI iffI)
fix y
assume "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 inner_simps) (simp add: field_simps)
with \<open>0 < d\<close> show "y \<in> (\<lambda>y. x - sum ((*\<^sub>R) d) Basis + (2 * d) *\<^sub>R y) ` cbox 0 One"
by (auto intro: image_eqI[where x= "inverse (2 * d) *\<^sub>R (y - (x - ?d))"])
next
fix y
assume "y \<in> (\<lambda>y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 One"
then obtain z where z: "z \<in> cbox 0 One" "y = x - ?d + (2*d) *\<^sub>R z"
by auto
then show "y \<in> cbox (x - ?d) (x + ?d)"
using z assms by (auto simp: mem_box inner_simps)
qed
obtain S where "finite S" "cbox 0 (\<Sum>Basis::'a) = convex hull S"
using unit_cube_convex_hull by auto
then show ?thesis
by (rule_tac that[of "(\<lambda>y. x - ?d + (2 * d) *\<^sub>R y)` S"]) (auto simp: convex_hull_affinity *)
qed
subsection\<^marker>\<open>tag unimportant\<close>\<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: 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: 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: inner_left_distrib add.commute)
lemma cbox_image_unit_interval:
fixes a :: "'a::euclidean_space"
assumes "cbox a b \<noteq> {}"
shows "cbox a b =
(+) 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 sum_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_sum) (auto simp: algebra_simps linearI)
have "finite ((+) 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\<^marker>\<open>tag unimportant\<close> \<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"
then have B: "0 < B""\<And>x. x\<in>S \<Longrightarrow> \<bar>f x\<bar> \<le> B"
using assms(3) by auto
obtain k where "k > 0" and k: "cball x k \<subseteq> S"
using \<open>x \<in> S\<close> assms(1) open_contains_cball_eq by blast
show "\<exists>d>0. \<forall>x'. norm (x' - x) < d \<longrightarrow> \<bar>f x' - f x\<bar> < e"
proof (intro exI conjI allI impI)
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:field_simps)
have "y \<in> S"
apply (rule k[THEN subsetD])
unfolding mem_cball dist_norm
apply (rule order_trans[of _ "2 * norm (x - y)"])
using as
by (auto simp: field_simps norm_minus_commute)
{
define w where "w = x + t *\<^sub>R (y - x)"
have "w \<in> S"
using \<open>k>0\<close> by (auto simp: dist_norm t_def w_def k[THEN subsetD])
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: algebra_simps)
also have "\<dots> = 0"
using \<open>t > 0\<close> by (auto simp: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: algebra_simps)
have 2: "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:field_simps)
have "f y - f x \<le> (f w - f x) / t"
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:field_simps)
also have "... < e"
using B(2)[OF \<open>w\<in>S\<close>] and B(2)[OF \<open>x\<in>S\<close>] 2 \<open>t > 0\<close> by (auto simp: field_simps)
finally have th1: "f y - f x < e" .
}
moreover
{
define w where "w = x - t *\<^sub>R (y - x)"
have "w \<in> S"
using \<open>k > 0\<close> by (auto simp: dist_norm t_def w_def k[THEN subsetD])
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: algebra_simps)
also have "\<dots> = x"
using \<open>t > 0\<close> by (auto simp: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: 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: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: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:field_simps)
also have "\<dots> = (f w + t * f y) / (1 + t)"
using \<open>t > 0\<close> by (simp add: add_divide_distrib)
also have "\<dots> < e + f y"
using \<open>t > 0\<close> * \<open>e > 0\<close> by (auto simp: 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\<^marker>\<open>tag unimportant\<close> \<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: norm_minus_commute)
have "(1 / 2) *\<^sub>R y + (1 / 2) *\<^sub>R z = x"
unfolding z_def by (auto simp: 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: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\<^marker>\<open>tag unimportant\<close> \<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_sum)
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_sum_Basis
unfolding c_def convex_hull_set_sum
apply (subst convex_hull_linear_image [symmetric])
apply (simp add: linear_iff scaleR_add_left)
apply (rule sum.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
by (clarsimp simp: dist_norm) (metis inner_commute inner_diff_right norm_bound_Basis_le)
have e': "e = (\<Sum>(i::'a)\<in>Basis. d)"
by (simp add: d_def)
show "convex hull c \<subseteq> cball x e"
unfolding 2
apply clarsimp
apply (subst euclidean_dist_l2)
apply (rule order_trans [OF L2_set_le_sum])
apply (rule zero_le_dist)
unfolding e'
apply (rule sum_mono, 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 c1 e(1)])
done
then have k: "\<forall>y\<in>convex hull c. f y \<le> k"
apply (rule_tac convex_on_convex_hull_bound, assumption)
by (simp add: k_def c)
have "e \<le> e * real DIM('a)"
using e(2) real_of_nat_ge_one_iff by auto
then have "d \<le> e"
by (simp add: d_def field_split_simps)
then have dsube: "cball x d \<subseteq> cball x e"
by (rule subset_cball)
have conv: "convex_on (cball x d) f"
using \<open>convex_on (convex hull c) f\<close> c2 convex_on_subset by blast
then have "\<forall>y\<in>cball x d. \<bar>f y\<bar> \<le> k + 2 * \<bar>f x\<bar>"
by (rule convex_bounds_lemma) (use c2 k in blast)
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, 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
end
diff --git a/src/HOL/Analysis/Cross3.thy b/src/HOL/Analysis/Cross3.thy
--- a/src/HOL/Analysis/Cross3.thy
+++ b/src/HOL/Analysis/Cross3.thy
@@ -1,227 +1,227 @@
(* Title: HOL/Analysis/Cross3.thy
Author: L C Paulson, University of Cambridge
Ported from HOL Light
*)
section\<open>Vector Cross Products in 3 Dimensions\<close>
theory "Cross3"
imports Determinants Cartesian_Euclidean_Space
begin
context includes no_Set_Product_syntax
begin \<comment>\<open>locally disable syntax for set product, to avoid warnings\<close>
definition\<^marker>\<open>tag important\<close> cross3 :: "[real^3, real^3] \<Rightarrow> real^3" (infixr "\<times>" 80)
where "a \<times> b \<equiv>
vector [a$2 * b$3 - a$3 * b$2,
a$3 * b$1 - a$1 * b$3,
a$1 * b$2 - a$2 * b$1]"
end
bundle cross3_syntax begin
notation cross3 (infixr "\<times>" 80)
no_notation Product_Type.Times (infixr "\<times>" 80)
end
bundle no_cross3_syntax begin
no_notation cross3 (infixr "\<times>" 80)
notation Product_Type.Times (infixr "\<times>" 80)
end
unbundle cross3_syntax
subsection\<open> Basic lemmas\<close>
lemmas cross3_simps = cross3_def inner_vec_def sum_3 det_3 vec_eq_iff vector_def algebra_simps
lemma dot_cross_self: "x \<bullet> (x \<times> y) = 0" "x \<bullet> (y \<times> x) = 0" "(x \<times> y) \<bullet> y = 0" "(y \<times> x) \<bullet> y = 0"
by (simp_all add: orthogonal_def cross3_simps)
lemma orthogonal_cross: "orthogonal (x \<times> y) x" "orthogonal (x \<times> y) y"
"orthogonal y (x \<times> y)" "orthogonal (x \<times> y) x"
by (simp_all add: orthogonal_def dot_cross_self)
lemma cross_zero_left [simp]: "0 \<times> x = 0" and cross_zero_right [simp]: "x \<times> 0 = 0" for x::"real^3"
by (simp_all add: cross3_simps)
lemma cross_skew: "(x \<times> y) = -(y \<times> x)" for x::"real^3"
by (simp add: cross3_simps)
lemma cross_refl [simp]: "x \<times> x = 0" for x::"real^3"
by (simp add: cross3_simps)
lemma cross_add_left: "(x + y) \<times> z = (x \<times> z) + (y \<times> z)" for x::"real^3"
by (simp add: cross3_simps)
lemma cross_add_right: "x \<times> (y + z) = (x \<times> y) + (x \<times> z)" for x::"real^3"
by (simp add: cross3_simps)
lemma cross_mult_left: "(c *\<^sub>R x) \<times> y = c *\<^sub>R (x \<times> y)" for x::"real^3"
by (simp add: cross3_simps)
lemma cross_mult_right: "x \<times> (c *\<^sub>R y) = c *\<^sub>R (x \<times> y)" for x::"real^3"
by (simp add: cross3_simps)
lemma cross_minus_left [simp]: "(-x) \<times> y = - (x \<times> y)" for x::"real^3"
by (simp add: cross3_simps)
lemma cross_minus_right [simp]: "x \<times> -y = - (x \<times> y)" for x::"real^3"
by (simp add: cross3_simps)
lemma left_diff_distrib: "(x - y) \<times> z = x \<times> z - y \<times> z" for x::"real^3"
by (simp add: cross3_simps)
lemma right_diff_distrib: "x \<times> (y - z) = x \<times> y - x \<times> z" for x::"real^3"
by (simp add: cross3_simps)
hide_fact (open) left_diff_distrib right_diff_distrib
proposition Jacobi: "x \<times> (y \<times> z) + y \<times> (z \<times> x) + z \<times> (x \<times> y) = 0" for x::"real^3"
by (simp add: cross3_simps)
proposition Lagrange: "x \<times> (y \<times> z) = (x \<bullet> z) *\<^sub>R y - (x \<bullet> y) *\<^sub>R z"
by (simp add: cross3_simps) (metis (full_types) exhaust_3)
proposition cross_triple: "(x \<times> y) \<bullet> z = (y \<times> z) \<bullet> x"
by (simp add: cross3_def inner_vec_def sum_3 vec_eq_iff algebra_simps)
lemma cross_components:
"(x \<times> y)$1 = x$2 * y$3 - y$2 * x$3" "(x \<times> y)$2 = x$3 * y$1 - y$3 * x$1" "(x \<times> y)$3 = x$1 * y$2 - y$1 * x$2"
by (simp_all add: cross3_def inner_vec_def sum_3 vec_eq_iff algebra_simps)
lemma cross_basis: "(axis 1 1) \<times> (axis 2 1) = axis 3 1" "(axis 2 1) \<times> (axis 1 1) = -(axis 3 1)"
"(axis 2 1) \<times> (axis 3 1) = axis 1 1" "(axis 3 1) \<times> (axis 2 1) = -(axis 1 1)"
"(axis 3 1) \<times> (axis 1 1) = axis 2 1" "(axis 1 1) \<times> (axis 3 1) = -(axis 2 1)"
using exhaust_3
by (force simp add: axis_def cross3_simps)+
lemma cross_basis_nonzero:
"u \<noteq> 0 \<Longrightarrow> u \<times> axis 1 1 \<noteq> 0 \<or> u \<times> axis 2 1 \<noteq> 0 \<or> u \<times> axis 3 1 \<noteq> 0"
- by (clarsimp simp add: axis_def cross3_simps) (metis vector_3 exhaust_3)
+ by (clarsimp simp add: axis_def cross3_simps) (metis exhaust_3)
lemma cross_dot_cancel:
fixes x::"real^3"
assumes deq: "x \<bullet> y = x \<bullet> z" and veq: "x \<times> y = x \<times> z" and x: "x \<noteq> 0"
shows "y = z"
proof -
have "x \<bullet> x \<noteq> 0"
by (simp add: x)
then have "y - z = 0"
using veq
by (metis (no_types, lifting) Cross3.right_diff_distrib Lagrange deq eq_iff_diff_eq_0 inner_diff_right scale_eq_0_iff)
then show ?thesis
using eq_iff_diff_eq_0 by blast
qed
lemma norm_cross_dot: "(norm (x \<times> y))\<^sup>2 + (x \<bullet> y)\<^sup>2 = (norm x * norm y)\<^sup>2"
unfolding power2_norm_eq_inner power_mult_distrib
by (simp add: cross3_simps power2_eq_square)
lemma dot_cross_det: "x \<bullet> (y \<times> z) = det(vector[x,y,z])"
by (simp add: cross3_simps)
lemma cross_cross_det: "(w \<times> x) \<times> (y \<times> z) = det(vector[w,x,z]) *\<^sub>R y - det(vector[w,x,y]) *\<^sub>R z"
using exhaust_3 by (force simp add: cross3_simps)
proposition dot_cross: "(w \<times> x) \<bullet> (y \<times> z) = (w \<bullet> y) * (x \<bullet> z) - (w \<bullet> z) * (x \<bullet> y)"
by (force simp add: cross3_simps)
proposition norm_cross: "(norm (x \<times> y))\<^sup>2 = (norm x)\<^sup>2 * (norm y)\<^sup>2 - (x \<bullet> y)\<^sup>2"
unfolding power2_norm_eq_inner power_mult_distrib
by (simp add: cross3_simps power2_eq_square)
lemma cross_eq_0: "x \<times> y = 0 \<longleftrightarrow> collinear{0,x,y}"
proof -
have "x \<times> y = 0 \<longleftrightarrow> norm (x \<times> y) = 0"
by simp
also have "... \<longleftrightarrow> (norm x * norm y)\<^sup>2 = (x \<bullet> y)\<^sup>2"
using norm_cross [of x y] by (auto simp: power_mult_distrib)
also have "... \<longleftrightarrow> \<bar>x \<bullet> y\<bar> = norm x * norm y"
using power2_eq_iff
by (metis (mono_tags, hide_lams) abs_minus abs_norm_cancel abs_power2 norm_mult power_abs real_norm_def)
also have "... \<longleftrightarrow> collinear {0, x, y}"
by (rule norm_cauchy_schwarz_equal)
finally show ?thesis .
qed
lemma cross_eq_self: "x \<times> y = x \<longleftrightarrow> x = 0" "x \<times> y = y \<longleftrightarrow> y = 0"
apply (metis cross_zero_left dot_cross_self(1) inner_eq_zero_iff)
by (metis cross_zero_right dot_cross_self(2) inner_eq_zero_iff)
lemma norm_and_cross_eq_0:
"x \<bullet> y = 0 \<and> x \<times> y = 0 \<longleftrightarrow> x = 0 \<or> y = 0" (is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
by (metis cross_dot_cancel cross_zero_right inner_zero_right)
qed auto
lemma bilinear_cross: "bilinear(\<times>)"
apply (auto simp add: bilinear_def linear_def)
apply unfold_locales
apply (simp add: cross_add_right)
apply (simp add: cross_mult_right)
apply (simp add: cross_add_left)
apply (simp add: cross_mult_left)
done
subsection \<open>Preservation by rotation, or other orthogonal transformation up to sign\<close>
lemma cross_matrix_mult: "transpose A *v ((A *v x) \<times> (A *v y)) = det A *\<^sub>R (x \<times> y)"
apply (simp add: vec_eq_iff )
apply (simp add: vector_matrix_mult_def matrix_vector_mult_def forall_3 cross3_simps)
done
lemma cross_orthogonal_matrix:
assumes "orthogonal_matrix A"
shows "(A *v x) \<times> (A *v y) = det A *\<^sub>R (A *v (x \<times> y))"
proof -
have "mat 1 = transpose (A ** transpose A)"
by (metis (no_types) assms orthogonal_matrix_def transpose_mat)
then show ?thesis
by (metis (no_types) vector_matrix_mul_rid vector_transpose_matrix cross_matrix_mult matrix_vector_mul_assoc matrix_vector_mult_scaleR)
qed
lemma cross_rotation_matrix: "rotation_matrix A \<Longrightarrow> (A *v x) \<times> (A *v y) = A *v (x \<times> y)"
by (simp add: rotation_matrix_def cross_orthogonal_matrix)
lemma cross_rotoinversion_matrix: "rotoinversion_matrix A \<Longrightarrow> (A *v x) \<times> (A *v y) = - A *v (x \<times> y)"
by (simp add: rotoinversion_matrix_def cross_orthogonal_matrix scaleR_matrix_vector_assoc)
lemma cross_orthogonal_transformation:
assumes "orthogonal_transformation f"
shows "(f x) \<times> (f y) = det(matrix f) *\<^sub>R f(x \<times> y)"
proof -
have orth: "orthogonal_matrix (matrix f)"
using assms orthogonal_transformation_matrix by blast
have "matrix f *v z = f z" for z
using assms orthogonal_transformation_matrix by force
with cross_orthogonal_matrix [OF orth] show ?thesis
by simp
qed
lemma cross_linear_image:
"\<lbrakk>linear f; \<And>x. norm(f x) = norm x; det(matrix f) = 1\<rbrakk>
\<Longrightarrow> (f x) \<times> (f y) = f(x \<times> y)"
by (simp add: cross_orthogonal_transformation orthogonal_transformation)
subsection \<open>Continuity\<close>
lemma continuous_cross: "\<lbrakk>continuous F f; continuous F g\<rbrakk> \<Longrightarrow> continuous F (\<lambda>x. (f x) \<times> (g x))"
apply (subst continuous_componentwise)
apply (clarsimp simp add: cross3_simps)
apply (intro continuous_intros; simp)
done
lemma continuous_on_cross:
fixes f :: "'a::t2_space \<Rightarrow> real^3"
shows "\<lbrakk>continuous_on S f; continuous_on S g\<rbrakk> \<Longrightarrow> continuous_on S (\<lambda>x. (f x) \<times> (g x))"
by (simp add: continuous_on_eq_continuous_within continuous_cross)
unbundle no_cross3_syntax
end
diff --git a/src/HOL/Analysis/Derivative.thy b/src/HOL/Analysis/Derivative.thy
--- a/src/HOL/Analysis/Derivative.thy
+++ b/src/HOL/Analysis/Derivative.thy
@@ -1,3475 +1,3475 @@
(* Title: HOL/Analysis/Derivative.thy
Author: John Harrison
Author: Robert Himmelmann, TU Muenchen (translation from HOL Light); tidied by LCP
*)
section \<open>Derivative\<close>
theory Derivative
imports
Bounded_Linear_Function
Line_Segment
Convex_Euclidean_Space
begin
declare bounded_linear_inner_left [intro]
declare has_derivative_bounded_linear[dest]
subsection \<open>Derivatives\<close>
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\<^marker>\<open>tag unimportant\<close> \<open>Derivative with composed bilinear function\<close>
text \<open>More explicit epsilon-delta forms.\<close>
proposition 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
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_componentwise_within:
"(f has_derivative f') (at a within S) \<longleftrightarrow>
(\<forall>i \<in> Basis. ((\<lambda>x. f x \<bullet> i) has_derivative (\<lambda>x. f' x \<bullet> i)) (at a within S))"
apply (simp add: has_derivative_within)
apply (subst tendsto_componentwise_iff)
apply (simp add: bounded_linear_componentwise_iff [symmetric] ball_conj_distrib)
apply (simp add: algebra_simps)
done
lemma has_derivative_at_withinI:
"(f has_derivative f') (at x) \<Longrightarrow> (f has_derivative f') (at x within s)"
unfolding has_derivative_within' has_derivative_at'
by blast
lemma has_derivative_right:
fixes f :: "real \<Rightarrow> real"
and y :: "real"
shows "(f has_derivative ((*) 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>
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 has_field_derivative_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 has_field_derivative_iff cong: Lim_cong_within)
qed
subsection \<open>Differentiability\<close>
definition\<^marker>\<open>tag important\<close>
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_withinI
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\<^marker>\<open>tag unimportant\<close> 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_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_const [simp, derivative_intros]: "(\<lambda>z. c) differentiable_on S"
by (simp add: differentiable_on_def)
lemma differentiable_on_mult [simp, derivative_intros]:
fixes f :: "'M::real_normed_vector \<Rightarrow> 'a::real_normed_algebra"
shows "\<lbrakk>f differentiable_on S; g differentiable_on S\<rbrakk> \<Longrightarrow> (\<lambda>z. f z * g z) differentiable_on S"
unfolding differentiable_on_def differentiable_def
using differentiable_def differentiable_mult by blast
lemma differentiable_on_compose:
"\<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 bounded_bilinear.FDERIV [of "(\<bullet>)" id id a _ id id]
by (auto simp: inner_commute dot_square_norm bounded_bilinear_inner)
lemma differentiable_sqnorm_at [derivative_intros, simp]:
fixes a :: "'a :: {real_normed_vector,real_inner}"
shows "(\<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)"
proposition 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)
lemma frechet_derivative_const [simp]: "frechet_derivative (\<lambda>x. c) (at a) = (\<lambda>x. 0)"
using differentiable_const frechet_derivative_works has_derivative_const has_derivative_unique by blast
lemma frechet_derivative_id [simp]: "frechet_derivative id (at a) = id"
using differentiable_def frechet_derivative_works has_derivative_id has_derivative_unique by blast
lemma frechet_derivative_ident [simp]: "frechet_derivative (\<lambda>x. x) (at a) = (\<lambda>x. x)"
by (metis eq_id_iff frechet_derivative_id)
subsection \<open>Differentiability implies continuity\<close>
proposition 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
lemma has_derivative_continuous_on:
"(\<And>x. x \<in> s \<Longrightarrow> (f has_derivative f' x) (at x within s)) \<Longrightarrow> continuous_on s f"
by (auto intro!: differentiable_imp_continuous_on differentiableI simp: differentiable_on_def)
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>
proposition diff_chain_within[derivative_intros]:
assumes "(f has_derivative f') (at x within s)"
and "(g has_derivative g') (at (f x) within (f ` s))"
shows "((g \<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)
lemma has_vector_derivative_within_open:
"a \<in> S \<Longrightarrow> open S \<Longrightarrow>
(f has_vector_derivative f') (at a within S) \<longleftrightarrow> (f has_vector_derivative f') (at a)"
by (simp only: at_within_interior interior_open)
lemma field_vector_diff_chain_within:
assumes Df: "(f has_vector_derivative f') (at x within S)"
and Dg: "(g has_field_derivative g') (at (f x) within f ` S)"
shows "((g \<circ> f) has_vector_derivative (f' * g')) (at x within S)"
using diff_chain_within[OF Df[unfolded has_vector_derivative_def]
Dg [unfolded has_field_derivative_def]]
by (auto simp: o_def mult.commute has_vector_derivative_def)
lemma vector_derivative_diff_chain_within:
assumes Df: "(f has_vector_derivative f') (at x within S)"
and Dg: "(g has_derivative g') (at (f x) within f`S)"
shows "((g \<circ> f) has_vector_derivative (g' f')) (at x within S)"
using diff_chain_within[OF Df[unfolded has_vector_derivative_def] Dg]
linear.scaleR[OF has_derivative_linear[OF Dg]]
unfolding has_vector_derivative_def o_def
by (auto simp: o_def mult.commute has_vector_derivative_def)
subsection\<^marker>\<open>tag unimportant\<close> \<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\<^marker>\<open>tag important\<close> \<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>
proposition frechet_derivative_unique_within:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
assumes 1: "(f has_derivative f') (at x within S)"
and 2: "(f has_derivative f'') (at x within S)"
and S: "\<And>i e. \<lbrakk>i\<in>Basis; e>0\<rbrakk> \<Longrightarrow> \<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 (intro allI impI)
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"
by (rule_tac x="x + d *\<^sub>R (SOME i. i \<in> Basis)" in bexI) (auto simp: dist_norm SOME_Basis nonzero_Basis) qed
then have *: "netlimit (at x within S) = x"
by (simp add: Lim_ident_at trivial_limit_within)
show ?thesis
proof (rule linear_eq_stdbasis)
show "linear f'" "linear f''"
unfolding linear_conv_bounded_linear using as by auto
next
fix i :: 'a
assume i: "i \<in> Basis"
define e where "e = norm (f' i - f'' i)"
show "f' i = f'' i"
proof (rule ccontr)
assume "f' i \<noteq> f'' i"
then have "e > 0"
unfolding e_def by auto
obtain d where d:
"0 < d"
"(\<And>y. y\<in>S \<longrightarrow> 0 < dist y x \<and> dist y x < d \<longrightarrow>
dist ((f y - f x - f' (y - x)) /\<^sub>R norm (y - x) -
(f y - f x - f'' (y - x)) /\<^sub>R norm (y - x)) (0 - 0) < e)"
using tendsto_diff [OF as(1,2)[THEN conjunct2]]
unfolding * Lim_within
using \<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
qed
proposition frechet_derivative_unique_within_closed_interval:
fixes f::"'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
assumes ab: "\<And>i. i\<in>Basis \<Longrightarrow> a\<bullet>i < b\<bullet>i"
and x: "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''"
proof (rule frechet_derivative_unique_within)
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
with ab[of i] \<open>e>0\<close> x i show ?thesis
by (rule_tac x="(min (b\<bullet>i - a\<bullet>i) e) / 2" in exI)
(auto simp add: mem_box field_simps inner_simps inner_Basis)
next
case False
moreover have "a \<bullet> i < x \<bullet> i"
using False i mem_box(2) x by force
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 \<open>a \<bullet> i < x \<bullet> i\<close> 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"
by (simp add: \<open>0 < e\<close> \<open>a \<bullet> i < x \<bullet> i\<close> less_eq_real_def)
then have "x \<bullet> i * 2 \<le> b \<bullet> i * 2 + min (x \<bullet> i - a \<bullet> i) e"
using i mem_box(2) x by force
ultimately show ?thesis
using ab[of i] \<open>e>0\<close> x i
by (rule_tac x="- (min (x\<bullet>i - a\<bullet>i) e) / 2" in exI)
(auto simp add: mem_box field_simps inner_simps inner_Basis)
qed
qed (use assms in auto)
lemma frechet_derivative_unique_within_open_interval:
fixes f::"'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
assumes x: "x \<in> box a b"
and f: "(f has_derivative f' ) (at x within box a b)" "(f has_derivative f'') (at x within box a b)"
shows "f' = f''"
proof -
have "at x within box a b = at x"
by (metis x at_within_interior interior_open open_box)
with f show "f' = f''"
by (simp add: has_derivative_unique)
qed
lemma frechet_derivative_at:
"(f has_derivative f') (at x) \<Longrightarrow> f' = frechet_derivative f (at x)"
using differentiable_def frechet_derivative_works has_derivative_unique by blast
lemma frechet_derivative_compose:
"frechet_derivative (f o g) (at x) = frechet_derivative (f) (at (g x)) o frechet_derivative g (at x)"
if "g differentiable at x" "f differentiable at (g x)"
by (metis diff_chain_at frechet_derivative_at frechet_derivative_works that)
lemma frechet_derivative_within_cbox:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
assumes "\<And>i. i\<in>Basis \<Longrightarrow> 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)
lemma frechet_derivative_transform_within_open:
"frechet_derivative f (at x) = frechet_derivative g (at x)"
if "f differentiable at x" "open X" "x \<in> X" "\<And>x. x \<in> X \<Longrightarrow> f x = g x"
by (meson frechet_derivative_at frechet_derivative_works has_derivative_transform_within_open that)
subsection \<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")
case False
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
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:
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
subsection \<open>One-dimensional mean value theorem\<close>
lemma mvt_simple:
fixes f :: "real \<Rightarrow> real"
assumes "a < b"
and derf: "\<And>x. \<lbrakk>a \<le> x; x \<le> b\<rbrakk> \<Longrightarrow> (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 derf unfolding differentiable_on_def differentiable_def by force
then show "continuous_on {a..b} f"
by (rule differentiable_imp_continuous_on)
show "(f has_derivative f' x) (at x)" if "a < x" "x < b" for x
by (metis at_within_Icc_at derf leI order.asym that)
qed (use assms in auto)
lemma mvt_very_simple:
fixes f :: "real \<Rightarrow> real"
assumes "a \<le> b"
and derf: "\<And>x. \<lbrakk>a \<le> x; x \<le> b\<rbrakk> \<Longrightarrow> (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
by force
next
case False
then show ?thesis
using mvt_simple[OF _ derf]
by (metis \<open>a \<le> b\<close> atLeastAtMost_iff dual_order.order_iff_strict greaterThanLessThan_iff)
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 contf: "continuous_on {a..b} f"
and derf: "\<And>x. \<lbrakk>a < x; x < b\<rbrakk> \<Longrightarrow> (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 [OF \<open>a < b\<close>, where f = "\<lambda>x. (f b - f a) \<bullet> f x"])
apply (intro continuous_intros contf)
using derf apply (auto 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 \<open>a < b\<close> by (rule_tac x="(a + b) /2" in bexI) auto
qed
qed
subsection \<open>More general bound theorems\<close>
proposition 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" .
have "y = b"
proof (cases "a = y")
case True
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
then show "y = b" using \<open>d > 0\<close> \<open>y < b\<close>
by (simp add: d'_def)
next
case False
with \<open>a \<le> y\<close> have "a < y" by simp
show "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 have "\<phi>' y * \<bar>x1 - y\<bar> \<le> \<phi> x1 - \<phi> y + e2 * \<bar>x1 - y\<bar>"
using elim 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'}"
show "?le x1"
proof (cases "x1 < y")
case True
then show ?thesis
using \<open>y \<in> A\<close> local.leI x1 by auto
next
case False
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 show "?le x1"
using False by (auto simp: algebra_simps)
qed
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
qed
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
show ?thesis
proof (rule field_le_epsilon)
fix e::real assume "e > 0"
then show "norm (f b - f a) \<le> \<phi> b - \<phi> a + e"
using *[of "e / (b - a + 1)"] \<open>a < b\<close> by simp
qed
qed
lemma differentiable_bound:
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
assumes "convex S"
and derf: "\<And>x. x\<in>S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
and B: "\<And>x. x \<in> S \<Longrightarrow> 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 *: "x + u *\<^sub>R (y - x) \<in> S" if "u \<in> {0..1}" for u
proof -
have "u *\<^sub>R y = u *\<^sub>R (y - x) + u *\<^sub>R x"
by (simp add: scale_right_diff_distrib)
then show "x + u *\<^sub>R (y - x) \<in> S"
using that \<open>convex S\<close> x y by (simp add: convex_alt)
(metis pth_b(2) pth_c(1) scaleR_collapse)
qed
have "\<And>z. z \<in> (\<lambda>u. x + u *\<^sub>R (y - x)) ` {0..1} \<Longrightarrow>
(f has_derivative f' z) (at z within (\<lambda>u. x + u *\<^sub>R (y - x)) ` {0..1})"
by (auto intro: * has_derivative_within_subset [OF derf])
then have "continuous_on (?p ` {0..1}) f"
unfolding continuous_on_eq_continuous_within
by (meson has_derivative_continuous)
with * have 1: "continuous_on {0 .. 1} (f \<circ> ?p)"
by (intro continuous_intros)+
{
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 derf *)
have "(f \<circ> ?p has_derivative (f' ?u) \<circ> (\<lambda>u. 0 + u *\<^sub>R (y - x))) (at u within box 0 1)"
by (intro derivative_intros has_derivative_within_subset [OF derf]) (use u * in auto)
hence "((f \<circ> ?p) has_vector_derivative f' ?u (y - x)) (at u)"
by (simp add: at_within_open[OF u open_greaterThanLessThan] scaleR has_vector_derivative_def o_def)
} note 2 = this
have 3: "continuous_on {0..1} ?\<phi>"
by (rule continuous_intros)+
have 4: "(?\<phi> has_vector_derivative B * norm (x - y)) (at u)" for u
by (auto simp: has_vector_derivative_def intro!: derivative_eq_intros)
{
fix u::real assume u: "u \<in>{0 <..< 1}"
let ?u = "?p u"
interpret bounded_linear "(f' ?u)"
using u by (auto intro!: has_derivative_bounded_linear derf *)
have "norm (f' ?u (y - x)) \<le> onorm (f' ?u) * norm (y - x)"
by (rule onorm) (rule bounded_linear)
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 field_differentiable_bound:
fixes S :: "'a::real_normed_field 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 (erule df [unfolded has_field_derivative_def])
apply (rule onorm_le, simp_all add: norm_mult mult_right_mono assms)
done
lemma
differentiable_bound_segment:
fixes f::"'a::real_normed_vector \<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: "\<And>x. x \<in> {0..1} \<Longrightarrow> 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 = (+) 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 (force simp: ac_simps)
qed
lemma differentiable_bound_linearization:
fixes f::"'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
assumes S: "\<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: "\<And>x. x \<in> S \<Longrightarrow> 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 "\<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)
with differentiable_bound_segment[OF S g] \<open>x0 \<in> S\<close>
show ?thesis
by (simp add: g_def field_simps linear_diff[OF has_derivative_linear[OF f']])
qed
lemma vector_differentiable_bound_linearization:
fixes f::"real \<Rightarrow> 'b::real_normed_vector"
assumes f': "\<And>x. x \<in> S \<Longrightarrow> (f has_vector_derivative f' x) (at x within S)"
assumes "closed_segment a b \<subseteq> S"
assumes B: "\<And>x. x \<in> S \<Longrightarrow> norm (f' x - f' x0) \<le> B"
assumes "x0 \<in> S"
shows "norm (f b - f a - (b - a) *\<^sub>R f' x0) \<le> norm (b - a) * B"
using assms
by (intro differentiable_bound_linearization[of a b S f "\<lambda>x h. h *\<^sub>R f' x" x0 B])
(force simp: closed_segment_real_eq has_vector_derivative_def
scaleR_diff_right[symmetric] mult.commute[of B]
intro!: onorm_le mult_left_mono)+
text \<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: "(*) 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_vector_derivative_zero_constant:
assumes "convex s"
assumes "\<And>x. x \<in> s \<Longrightarrow> (f has_vector_derivative 0) (at x within s)"
obtains c where "\<And>x. x \<in> s \<Longrightarrow> f x = c"
using has_derivative_zero_constant[of s f] assms
by (auto simp: has_vector_derivative_def)
lemma has_derivative_zero_unique:
fixes f :: "'a::real_normed_vector \<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)
+ (auto simp: at_within_open[OF _ open_ball] f)
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 derf: "(f has_derivative f') (at (g y))"
and ling': "bounded_linear g'"
and "g' \<circ> f' = id"
and contg: "continuous (at y) g"
and "open T"
and "y \<in> T"
and fg: "\<And>z. z \<in> T \<Longrightarrow> 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 (intro allI impI)
fix e :: real
assume "e > 0"
with C(1) have *: "e / C > 0" by auto
obtain d0 where "0 < d0" and d0:
"\<And>u. norm (u - g y) < d0 \<Longrightarrow> norm (f u - f (g y) - f' (u - g y)) \<le> e / C * norm (u - g y)"
using derf * unfolding has_derivative_at_alt by blast
obtain d1 where "0 < d1" and d1: "\<And>x. \<lbrakk>0 < dist x y; dist x y < d1\<rbrakk> \<Longrightarrow> dist (g x) (g y) < d0"
using contg \<open>0 < d0\<close> unfolding continuous_at Lim_at by blast
obtain d2 where "0 < d2" and d2: "\<And>u. dist u y < d2 \<Longrightarrow> u \<in> T"
using \<open>open T\<close> \<open>y \<in> T\<close> unfolding open_dist by blast
obtain d where d: "0 < d" "d < d1" "d < d2"
using field_lbound_gt_zero[OF \<open>0 < d1\<close> \<open>0 < d2\<close>] by blast
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)"
proof (intro exI allI impI conjI)
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] fg[OF \<open>z\<in>T\<close>]
by (simp add: norm_minus_commute)
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"
proof -
have "norm (g z - g y) < d0"
by (metis as cancel_comm_monoid_add_class.diff_cancel d(2) \<open>0 < d0\<close> d1 diff_gt_0_iff_gt diff_strict_mono dist_norm dist_self zero_less_dist_iff)
then show ?thesis
by (metis C(1) \<open>y \<in> T\<close> d0 fg real_mult_le_cancel_iff1)
qed
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 (use d in auto)
qed
have *: "(0::real) < 1 / 2"
by auto
obtain d where "0 < d" and d:
"\<And>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)"
by (rule add_left_mono) (use d z in auto)
also have "\<dots> \<le> norm (z - y) * C + 1 / 2 * norm (g z - g y)"
by (rule add_right_mono) (use C in auto)
finally show "norm (g z - g y) \<le> B * norm (z - y)"
unfolding B_def
by (auto simp add: field_simps)
qed
show ?thesis
unfolding has_derivative_at_alt
proof (intro conjI assms allI impI)
fix e :: real
assume "e > 0"
then have *: "e / B > 0" by (metis \<open>B > 0\<close> divide_pos_pos)
obtain d' where "0 < d'" and d':
"\<And>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 field_lbound_gt_zero[OF \<open>0 < d\<close> \<open>0 < d'\<close>] 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)"
proof (intro exI allI impI conjI)
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] k as \<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 (use k in auto)
qed
qed
text\<^marker>\<open>tag unimportant\<close>\<open>Inverse function theorem for complex derivatives\<close>
lemma has_field_derivative_inverse_basic:
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
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 "\<And>y. y \<in> T \<Longrightarrow> f (g y) = y"
shows "(g has_derivative g') (at (f x))"
by (rule has_derivative_inverse_basic) (use assms in auto)
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 "\<And>x. x \<in> S \<Longrightarrow> 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>
proposition has_derivative_inverse:
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
assumes "compact S"
and "x \<in> S"
and fx: "f x \<in> interior (f ` S)"
and "continuous_on S f"
and gf: "\<And>y. y \<in> S \<Longrightarrow> 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 -
have *: "\<And>y. y \<in> interior (f ` S) \<Longrightarrow> f (g y) = y"
by (metis gf image_iff interior_subset subsetCE)
show ?thesis
apply (rule has_derivative_inverse_basic_x[OF assms(6-8), where T = "interior (f ` S)"])
apply (rule continuous_on_interior[OF _ fx])
apply (rule continuous_on_inv)
apply (simp_all add: assms *)
done
qed
text \<open>Invertible derivative continuous at a point implies local
injectivity. It's only for this we need continuity of the derivative,
except of course if we want the fact that the inverse derivative is
also continuous. So if we know for some other reason that the inverse
function exists, it's OK.\<close>
proposition has_derivative_locally_injective:
fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
assumes "a \<in> S"
and "open S"
and bling: "bounded_linear g'"
and "g' \<circ> f' a = id"
and derf: "\<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)"
using f'g' by auto
then have *: "0 < onorm g'"
unfolding onorm_pos_lt[OF assms(3)]
by fastforce
define k where "k = 1 / onorm g' / 2"
have *: "k > 0"
unfolding k_def using * by auto
obtain d1 where d1:
"0 < d1"
"\<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 < d2" "ball a d2 \<subseteq> S"
using \<open>0 < d2\<close> \<open>ball a d2 \<subseteq> S\<close> by blast
obtain d where d: "0 < d" "d < d1" "d < d2"
using field_lbound_gt_zero[OF d1(1) d2(1)] by blast
show ?thesis
proof
show "0 < d" by (fact d)
show "ball a d \<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 by (simp add: diff f'g')
have "norm (ph x - ph y) \<le> (1 / 2) * norm (x - y)"
proof (rule differentiable_bound[OF convex_ball _ _ as(1-2)])
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
have blin: "bounded_linear (f' a)"
using \<open>a \<in> S\<close> derf by blast
show "(ph has_derivative (\<lambda>v. v - g' (f' u v))) (at u within ball a d)"
unfolding ph' * comp_def
by (rule \<open>u \<in> S\<close> derivative_eq_intros has_derivative_at_withinI [OF derf] bounded_linear.has_derivative [OF blin] bounded_linear.has_derivative [OF bling] |simp)+
have **: "bounded_linear (\<lambda>x. f' u x - f' a x)" "bounded_linear (\<lambda>x. f' a x - f' u x)"
using \<open>u \<in> S\<close> blin bounded_linear_sub derf by auto
then have "onorm (\<lambda>v. v - g' (f' u v)) \<le> onorm g' * onorm (\<lambda>w. f' a w - f' u w)"
by (simp add: "*" bounded_linear_axioms onorm_compose)
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"] d u onorm_pos_le[OF bling] 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)"
by (simp add: as(3) ph_def)
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 derf: "\<And>n x. x \<in> S \<Longrightarrow> ((f n) has_derivative (f' n x)) (at x within S)"
and nle: "\<And>n x h. \<lbrakk>n\<ge>N; x \<in> S\<rbrakk> \<Longrightarrow> 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 clarify
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)"
proof (rule differentiable_bound[where f'="\<lambda>x h. f' m x h - f' n x h", OF \<open>convex S\<close> _ _ as(3-4)])
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 derf \<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"]
by (auto simp add: algebra_simps norm_minus_commute)
also have "\<dots> \<le> e * norm h + e * norm h"
using nle[OF \<open>N \<le> m\<close> \<open>x \<in> S\<close>, of h] nle[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 "\<And>n x. x \<in> S \<Longrightarrow> ((f n) has_derivative (f' n x)) (at x within S)"
and nle: "\<And>e. e > 0 \<Longrightarrow> \<forall>\<^sub>F n in sequentially. \<forall>x\<in>S. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h"
and "e > 0"
shows "\<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 -
have *: "2 * (e/2) = e"
using \<open>e > 0\<close> by auto
obtain N where "\<forall>n\<ge>N. \<forall>x\<in>S. \<forall>h. norm (f' n x h - g' x h) \<le> (e/2) * norm h"
using nle \<open>e > 0\<close>
unfolding eventually_sequentially
by (metis less_divide_eq_numeral1(1) mult_zero_left)
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="e/2", unfolded *])
using assms \<open>e > 0\<close>
apply auto
done
qed
proposition has_derivative_sequence:
fixes f :: "nat \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::banach"
assumes "convex S"
and derf: "\<And>n x. x \<in> S \<Longrightarrow> ((f n) has_derivative (f' n x)) (at x within S)"
and nle: "\<And>e. e > 0 \<Longrightarrow> \<forall>\<^sub>F n in sequentially. \<forall>x\<in>S. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h"
and "x0 \<in> S"
and lim: "((\<lambda>n. f n x0) \<longlongrightarrow> l) sequentially"
shows "\<exists>g. \<forall>x\<in>S. (\<lambda>n. f n x) \<longlonglongrightarrow> g x \<and> (g has_derivative g'(x)) (at x within S)"
proof -
have lem1: "\<And>e. e > 0 \<Longrightarrow> \<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"
proof (intro ballI bchoice)
fix x
assume "x \<in> S"
show "\<exists>y. (\<lambda>n. f n x) \<longlonglongrightarrow> y"
unfolding convergent_eq_Cauchy
proof (cases "x = x0")
case True
then show "Cauchy (\<lambda>n. f n x)"
using LIMSEQ_imp_Cauchy[OF lim] by auto
next
case False
show "Cauchy (\<lambda>n. f n x)"
unfolding Cauchy_def
proof (intro allI impI)
fix e :: real
assume "e > 0"
hence *: "e / 2 > 0" "e / 2 / norm (x - x0) > 0" using False by auto
obtain M where M: "\<forall>m\<ge>M. \<forall>n\<ge>M. dist (f m x0) (f n x0) < e / 2"
using LIMSEQ_imp_Cauchy[OF lim] * unfolding Cauchy_def by blast
obtain N where N:
"\<forall>m\<ge>N. \<forall>n\<ge>N.
\<forall>u\<in>S. \<forall>y\<in>S. norm (f m u - f n u - (f m y - f n y)) \<le>
e / 2 / norm (x - x0) * norm (u - 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"
proof (intro exI allI impI)
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 \<open>x\<in>S\<close> \<open>x0\<in>S\<close> as False by fastforce
also have "\<dots> < e / 2 + e / 2"
by (rule add_strict_right_mono) (use as M in \<open>auto simp: dist_norm\<close>)
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: "\<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)" if "e > 0" for e
proof -
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 \<open>e > 0\<close> 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)"
proof (intro exI ballI allI impI)
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
proof (intro exI allI impI)
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 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 (simp add: tendsto_upperbound)
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 allI impI)
fix x
assume "x \<in> S"
then show "(\<lambda>n. f n x) \<longlonglongrightarrow> g x"
by (simp add: g)
have tog': "(\<lambda>n. f' n x u) \<longlonglongrightarrow> g' x u" for u
unfolding filterlim_def le_nhds_metric_le eventually_filtermap dist_norm
proof (intro allI impI)
fix e :: real
assume "e > 0"
show "eventually (\<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 nle \<open>0 < e\<close> \<open>x \<in> S\<close> by (fast elim: eventually_mono)
then show ?thesis
using \<open>u = 0\<close> \<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 nle \<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 tog'])
unfolding lin[THEN bounded_linear.linear, THEN linear_cmul]
apply (intro tendsto_intros tog')
done
show "g' x (y + z) = g' x y + g' x z"
apply (rule tendsto_unique[OF trivial_limit_sequentially tog'])
unfolding lin[THEN bounded_linear.linear, THEN linear_add]
apply (rule tendsto_add)
apply (rule tog')+
done
obtain N where N: "\<forall>h. norm (f' N x h - g' x h) \<le> 1 * norm h"
using nle \<open>x \<in> S\<close> unfolding eventually_sequentially by (fast intro: zero_less_one)
have "bounded_linear (f' N x)"
using derf \<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 "eventually (\<lambda>y. norm (g y - g x - g' x (y - x)) \<le> e * norm (y - x)) (at x within S)"
if "e > 0" for e
proof -
have *: "e / 3 > 0"
using that by auto
obtain N1 where N1: "\<forall>n\<ge>N1. \<forall>x\<in>S. \<forall>h. norm (f' n x h - g' x h) \<le> e / 3 * norm h"
using nle * unfolding eventually_sequentially by blast
obtain N2 where
N2[rule_format]: "\<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 derf[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[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 der: "\<And>n x. x \<in> S \<Longrightarrow> ((f n) has_derivative (f' n x)) (at x within S)"
and no: "\<And>e. e > 0 \<Longrightarrow> \<forall>\<^sub>F n in sequentially.
\<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 \<open>convex S\<close> _ no, of "\<lambda>n x. f n x + (f 0 a - f n a)"])
apply (metis assms(2) has_derivative_add_const)
using \<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 "\<And>e. e>0 \<Longrightarrow> \<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
*: "\<And>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 * by metis
obtain f' where
f': "\<And>x. \<forall>z\<in>S. (f x has_derivative f' x z) (at z within S) \<and>
(\<forall>h. norm (f' x z h - g' z h) \<le> inverse (real (Suc x)) * norm h)"
using * by metis
show ?thesis
proof (rule has_antiderivative_sequence[OF \<open>convex S\<close>, of f f'])
fix e :: real
assume "e > 0"
obtain N where N: "inverse (real (Suc N)) < e"
using reals_Archimedean[OF \<open>e>0\<close>] ..
show "\<forall>\<^sub>F n in sequentially. \<forall>x\<in>S. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h"
unfolding eventually_sequentially
proof (intro exI allI ballI impI)
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"
by (meson "*" mult_right_mono norm_ge_zero order.trans x f')
qed
qed (use f' in auto)
qed
subsection \<open>Differentiation of a series\<close>
proposition 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 "\<And>e. e>0 \<Longrightarrow> \<forall>\<^sub>F n in sequentially. \<forall>x\<in>S. \<forall>h. norm (sum (\<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_sum)
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>\<^sub>F n in sequentially.
\<forall>x\<in>S. \<forall>h. norm ((\<Sum>i<n. f' i x * h) - g' x * h) \<le> e * norm h" if "e > 0" for e
unfolding eventually_sequentially
proof -
from that 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 sum_distrib_right)
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 (use assms in \<open>auto simp: has_field_derivative_def\<close>)
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 (*) (g' x)) (at x)"
by (simp add: has_field_derivative_def S)
have "((\<lambda>x. \<Sum>n. f n x) has_derivative (*) (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+
subsection \<open>Derivative as a vector\<close>
text \<open>Considering derivative \<^typ>\<open>real \<Rightarrow> 'b::real_normed_vector\<close> 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, simp_all)
show "\<exists>d. d \<noteq> 0 \<and> \<bar>d\<bar> < e \<and> x + d \<in> S" if "0 < e" for e
proof -
from that
obtain x' where "x' \<in> S" "x' \<noteq> x" "\<bar>x' - x\<bar> < e"
using islimpt_approachable_real[of x S] not_bot
by (auto simp add: trivial_limit_within)
then show ?thesis
using eq_iff_diff_eq_0 by fastforce
qed
qed (use f' f'' in \<open>auto simp: has_vector_derivative_def\<close>)
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)
proposition 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)
lemma has_vector_derivative_cong_ev:
assumes *: "eventually (\<lambda>x. x \<in> S \<longrightarrow> f x = g x) (nhds x)" "f x = g x"
shows "(f has_vector_derivative f') (at x within S) = (g has_vector_derivative f') (at x within S)"
unfolding has_vector_derivative_def has_derivative_def
using *
apply (cases "at x within S \<noteq> bot")
apply (intro refl conj_cong filterlim_cong)
apply (auto simp: Lim_ident_at eventually_at_filter elim: eventually_mono)
done
lemma islimpt_closure_open:
fixes s :: "'a::perfect_space set"
assumes "open s" and t: "t = closure s" "x \<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_cbox:
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 vector_derivative_within_closed_interval:
fixes f::"real \<Rightarrow> 'a::euclidean_space"
assumes "a < b" and "x \<in> {a..b}"
assumes "(f has_vector_derivative f') (at x within {a..b})"
shows "vector_derivative f (at x within {a..b}) = f'"
using assms vector_derivative_within_cbox
by fastforce
lemma has_vector_derivative_within_subset:
"(f has_vector_derivative f') (at x within S) \<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_withinI)
lemma has_vector_derivative_weaken:
fixes x D and f g S T
assumes f: "(f has_vector_derivative D) (at x within T)"
and "x \<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 has_vector_derivative_transform:
assumes "x \<in> S" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
assumes f': "(f has_vector_derivative f') (at x within S)"
shows "(g has_vector_derivative f') (at x within S)"
using assms
unfolding has_vector_derivative_def
by (rule has_derivative_transform)
lemma vector_diff_chain_at:
assumes "(f has_vector_derivative f') (at x)"
and "(g has_vector_derivative g') (at (f x))"
shows "((g \<circ> f) has_vector_derivative (f' *\<^sub>R g')) (at x)"
using assms has_vector_derivative_at_within has_vector_derivative_def vector_derivative_diff_chain_within by blast
lemma vector_diff_chain_within:
assumes "(f has_vector_derivative f') (at x within s)"
and "(g has_vector_derivative g') (at (f x) within f ` s)"
shows "((g \<circ> f) has_vector_derivative (f' *\<^sub>R g')) (at x within s)"
using assms has_vector_derivative_def vector_derivative_diff_chain_within by blast
lemma vector_derivative_const_at [simp]: "vector_derivative (\<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_cbox 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_within:
assumes "at x within S \<noteq> bot" "f differentiable (at x within S)"
"(g has_derivative g') (at (f x) within f ` S)"
shows "vector_derivative (g \<circ> f) (at x within S) =
g' (vector_derivative f (at x within S)) "
apply (rule vector_derivative_within [OF \<open>at x within S \<noteq> bot\<close>])
apply (rule vector_derivative_diff_chain_within)
using assms(2-3) vector_derivative_works
by auto
subsection \<open>Field differentiability\<close>
definition\<^marker>\<open>tag important\<close> 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_imp_differentiable:
"f field_differentiable F \<Longrightarrow> f differentiable F"
unfolding field_differentiable_def differentiable_def
using has_field_derivative_imp_has_derivative by auto
lemma field_differentiable_imp_continuous_at:
"f field_differentiable (at x within S) \<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]: "((*) c) field_differentiable F"
unfolding field_differentiable_def has_field_derivative_def mult_commute_abs
by (force intro: has_derivative_mult_right)
lemma field_differentiable_const [simp,derivative_intros]: "(\<lambda>z. c) field_differentiable F"
unfolding field_differentiable_def has_field_derivative_def
using DERIV_const has_field_derivative_imp_has_derivative by blast
lemma field_differentiable_ident [simp,derivative_intros]: "(\<lambda>z. z) field_differentiable F"
unfolding field_differentiable_def has_field_derivative_def
using DERIV_ident has_field_derivative_def by blast
lemma field_differentiable_id [simp,derivative_intros]: "id field_differentiable F"
unfolding id_def by (rule field_differentiable_ident)
lemma field_differentiable_minus [derivative_intros]:
"f field_differentiable F \<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]:
"(+) c field_differentiable F"
by (simp add: field_differentiable_add)
lemma field_differentiable_sum [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)
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))"
apply (auto simp: algebra_split_simps intro!: eventuallyI)
apply (rule mult_left_mono)
apply (auto simp add: field_simps power_abs intro!: divide_right_mono power_le_one)
done
then have "uniform_limit (T \<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 if_distribR
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 tendsto_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 (simp add: field_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)
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)
- (auto simp: algebra_simps field_split_simps exp_add_commuting[symmetric])
+ (auto simp: field_split_simps exp_add_commuting[symmetric])
qed (rule bounded_linear_scaleR_left)
lemma exp_times_scaleR_commute: "exp (t *\<^sub>R A) * A = A * exp (t *\<^sub>R A)"
using exp_times_arg_commute[symmetric, of "t *\<^sub>R A"]
by (auto simp: algebra_simps)
lemma exp_scaleR_has_vector_derivative_left: "((\<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)
lemma field_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 "(\<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(2)[OF x] have g': "(g has_derivative (*) (g' x)) (at x)"
by (simp add: has_field_derivative_def S)
have "((\<lambda>x. \<Sum>n. f n x) has_derivative (*) (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
subsubsection\<^marker>\<open>tag unimportant\<close>\<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>Field derivative\<close>
definition\<^marker>\<open>tag important\<close> 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 deriv_cong_ev:
assumes "eventually (\<lambda>x. f x = g x) (nhds x)" "x = y"
shows "deriv f x = deriv g y"
proof -
have "(\<lambda>D. (f has_field_derivative D) (at x)) = (\<lambda>D. (g has_field_derivative D) (at y))"
by (intro ext DERIV_cong_ev refl assms)
thus ?thesis by (simp add: deriv_def assms)
qed
lemma higher_deriv_cong_ev:
assumes "eventually (\<lambda>x. f x = g x) (nhds x)" "x = y"
shows "(deriv ^^ n) f x = (deriv ^^ n) g y"
proof -
from assms(1) have "eventually (\<lambda>x. (deriv ^^ n) f x = (deriv ^^ n) g x) (nhds x)"
proof (induction n arbitrary: f g)
case (Suc n)
from Suc.prems have "eventually (\<lambda>y. eventually (\<lambda>z. f z = g z) (nhds y)) (nhds x)"
by (simp add: eventually_eventually)
hence "eventually (\<lambda>x. deriv f x = deriv g x) (nhds x)"
by eventually_elim (rule deriv_cong_ev, simp_all)
thus ?case by (auto intro!: deriv_cong_ev Suc simp: funpow_Suc_right simp del: funpow.simps)
qed auto
from eventually_nhds_x_imp_x[OF this] assms(2) show ?thesis by simp
qed
lemma real_derivative_chain:
fixes x :: real
shows "f differentiable at x \<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)
proposition 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 vector_derivative_chain_at_general:
assumes "f differentiable at x" "g field_differentiable 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 vector_derivative_works by (auto simp: field_differentiable_derivI)
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 deriv_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_uminus [simp]: "deriv (\<lambda>w. -w) = (\<lambda>z. -1)"
using deriv_linear[of "-1"] by (simp del: deriv_linear)
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:
"f field_differentiable at z \<Longrightarrow> deriv (\<lambda>w. c * f w) z = c * deriv f z"
by simp
lemma deriv_cmult_right:
"f field_differentiable at z \<Longrightarrow> deriv (\<lambda>w. f w * c) z = deriv f z * c"
by simp
lemma deriv_inverse [simp]:
"\<lbrakk>f field_differentiable at z; f z \<noteq> 0\<rbrakk>
\<Longrightarrow> deriv (\<lambda>w. inverse (f w)) z = - deriv f z / f z ^ 2"
unfolding DERIV_deriv_iff_field_differentiable[symmetric]
by (safe intro!: DERIV_imp_deriv derivative_eq_intros) (auto simp: field_split_simps power2_eq_square)
lemma deriv_divide [simp]:
"\<lbrakk>f field_differentiable at z; g field_differentiable at z; g z \<noteq> 0\<rbrakk>
\<Longrightarrow> deriv (\<lambda>w. f w / g w) z = (deriv f z * g z - f z * deriv g z) / g z ^ 2"
by (simp add: field_class.field_divide_inverse field_differentiable_inverse)
(simp add: field_split_simps power2_eq_square)
lemma deriv_cdivide_right:
"f field_differentiable at z \<Longrightarrow> deriv (\<lambda>w. f w / c) z = deriv f z / c"
by (simp add: field_class.field_divide_inverse)
lemma deriv_compose_linear:
"f field_differentiable at (c * z) \<Longrightarrow> deriv (\<lambda>w. f (c * w)) z = c * deriv f (c * z)"
apply (rule DERIV_imp_deriv)
unfolding DERIV_deriv_iff_field_differentiable [symmetric]
by (metis (full_types) DERIV_chain2 DERIV_cmult_Id mult.commute)
lemma nonzero_deriv_nonconstant:
assumes df: "DERIV f \<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> has_field_derivative_transform_within_open [OF df S] DERIV_const DERIV_unique)
subsection \<open>Relation between convexity and derivative\<close>
(* TODO: Generalise to real vector spaces? *)
proposition convex_on_imp_above_tangent:
assumes convex: "convex_on A f" and connected: "connected A"
assumes c: "c \<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 has_field_derivative_iff using interior_subset[of A] by (blast intro: tendsto_mono at_le)
moreover from eventually_at_right_real[OF xc]
have "eventually (\<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: field_split_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 (simp add: tendsto_upperbound)
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 has_field_derivative_iff using interior_subset[of A] by (blast intro: tendsto_mono at_le)
moreover from eventually_at_left_real[OF xc]
have "eventually (\<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: field_split_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 (simp add: tendsto_lowerbound)
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'" "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
proposition has_derivative_partialsI:
fixes f::"'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector \<Rightarrow> 'c::real_normed_vector"
assumes fx: "((\<lambda>x. f x y) has_derivative fx) (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 fy_cont[unfolded continuous_within]: "continuous (at (x, y) within X \<times> Y) (\<lambda>(x, y). fy x y)"
assumes "y \<in> Y" "convex Y"
shows "((\<lambda>(x, y). f x y) has_derivative (\<lambda>(tx, ty). fx tx + fy x y ty)) (at (x, y) within X \<times> Y)"
proof (safe intro!: has_derivativeI tendstoI, goal_cases)
case (2 e')
interpret fx: bounded_linear "fx" using fx by (rule has_derivative_bounded_linear)
define e where "e = e' / 9"
have "e > 0" using \<open>e' > 0\<close> by (simp add: e_def)
from fy_cont[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>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 "\<And>x. x \<in> ball y d \<inter> Y \<Longrightarrow> 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' - x)) \<le> norm (x' - x) * e"
from fx[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 (simp,
simp add: dist_norm field_split_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[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' - x) + fy x y (y' - y))) /\<^sub>R
norm ((x', y') - (x, y)))
< e'"
apply eventually_elim
proof safe
fix x' y'
have "norm (f x' y' - f x y - (fx (x' - x) + fy x y (y' - y))) \<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' - 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 - (fx) (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' - x) + fy x y (y' - y))) /\<^sub>R norm ((x', y') - (x, y))) < e'"
by (simp add: dist_norm) (auto simp add: field_split_simps)
qed
then show ?case
by eventually_elim (auto simp: dist_norm field_simps)
next
from has_derivative_bounded_linear[OF fx]
obtain fxb where "fx = blinfun_apply fxb"
by (metis bounded_linear_Blinfun_apply)
then show "bounded_linear (\<lambda>(tx, ty). fx tx + blinfun_apply (fy x y) ty)"
by (auto intro!: bounded_linear_intros simp: split_beta')
qed
subsection\<^marker>\<open>tag unimportant\<close> \<open>Differentiable case distinction\<close>
lemma has_derivative_within_If_eq:
"((\<lambda>x. if P x then f x else g x) has_derivative f') (at x within S) =
(bounded_linear f' \<and>
((\<lambda>y.(if P y then (f y - ((if P x then f x else g x) + f' (y - x)))/\<^sub>R norm (y - x)
else (g y - ((if P x then f x else g x) + f' (y - x)))/\<^sub>R norm (y - x)))
\<longlongrightarrow> 0) (at x within S))"
(is "_ = (_ \<and> (?if \<longlongrightarrow> 0) _)")
proof -
have "(\<lambda>y. (1 / norm (y - x)) *\<^sub>R
((if P y then f y else g y) -
((if P x then f x else g x) + f' (y - x)))) = ?if"
by (auto simp: inverse_eq_divide)
thus ?thesis by (auto simp: has_derivative_within)
qed
lemma has_derivative_If_within_closures:
assumes f': "x \<in> S \<union> (closure S \<inter> closure T) \<Longrightarrow>
(f has_derivative f' x) (at x within S \<union> (closure S \<inter> closure T))"
assumes g': "x \<in> T \<union> (closure S \<inter> closure T) \<Longrightarrow>
(g has_derivative g' x) (at x within T \<union> (closure S \<inter> closure T))"
assumes connect: "x \<in> closure S \<Longrightarrow> x \<in> closure T \<Longrightarrow> f x = g x"
assumes connect': "x \<in> closure S \<Longrightarrow> x \<in> closure T \<Longrightarrow> f' x = g' x"
assumes x_in: "x \<in> S \<union> T"
shows "((\<lambda>x. if x \<in> S then f x else g x) has_derivative
(if x \<in> S then f' x else g' x)) (at x within (S \<union> T))"
proof -
from f' x_in interpret f': bounded_linear "if x \<in> S then f' x else (\<lambda>x. 0)"
by (auto simp add: has_derivative_within)
from g' interpret g': bounded_linear "if x \<in> T then g' x else (\<lambda>x. 0)"
by (auto simp add: has_derivative_within)
have bl: "bounded_linear (if x \<in> S then f' x else g' x)"
using f'.scaleR f'.bounded f'.add g'.scaleR g'.bounded g'.add x_in
by (unfold_locales; force)
show ?thesis
using f' g' closure_subset[of T] closure_subset[of S]
unfolding has_derivative_within_If_eq
by (intro conjI bl tendsto_If_within_closures x_in)
(auto simp: has_derivative_within inverse_eq_divide connect connect' subsetD)
qed
lemma has_vector_derivative_If_within_closures:
assumes x_in: "x \<in> S \<union> T"
assumes "u = S \<union> T"
assumes f': "x \<in> S \<union> (closure S \<inter> closure T) \<Longrightarrow>
(f has_vector_derivative f' x) (at x within S \<union> (closure S \<inter> closure T))"
assumes g': "x \<in> T \<union> (closure S \<inter> closure T) \<Longrightarrow>
(g has_vector_derivative g' x) (at x within T \<union> (closure S \<inter> closure T))"
assumes connect: "x \<in> closure S \<Longrightarrow> x \<in> closure T \<Longrightarrow> f x = g x"
assumes connect': "x \<in> closure S \<Longrightarrow> x \<in> closure T \<Longrightarrow> f' x = g' x"
shows "((\<lambda>x. if x \<in> S then f x else g x) has_vector_derivative
(if x \<in> S then f' x else g' x)) (at x within u)"
unfolding has_vector_derivative_def assms
using x_in
apply (intro has_derivative_If_within_closures[where ?f' = "\<lambda>x a. a *\<^sub>R f' x" and ?g' = "\<lambda>x a. a *\<^sub>R g' x",
THEN has_derivative_eq_rhs])
subgoal by (rule f'[unfolded has_vector_derivative_def]; assumption)
subgoal by (rule g'[unfolded has_vector_derivative_def]; assumption)
by (auto simp: assms)
subsection\<^marker>\<open>tag important\<close>\<open>The Inverse Function Theorem\<close>
lemma linear_injective_contraction:
assumes "linear f" "c < 1" and le: "\<And>x. norm (f x - x) \<le> c * norm x"
shows "inj f"
unfolding linear_injective_0[OF \<open>linear f\<close>]
proof safe
fix x
assume "f x = 0"
with le [of x] have "norm x \<le> c * norm x"
by simp
then show "x = 0"
using \<open>c < 1\<close> by (simp add: mult_le_cancel_right1)
qed
text\<open>From an online proof by J. Michael Boardman, Department of Mathematics, Johns Hopkins University\<close>
lemma inverse_function_theorem_scaled:
fixes f::"'a::euclidean_space \<Rightarrow> 'a"
and f'::"'a \<Rightarrow> ('a \<Rightarrow>\<^sub>L 'a)"
assumes "open U"
and derf: "\<And>x. x \<in> U \<Longrightarrow> (f has_derivative blinfun_apply (f' x)) (at x)"
and contf: "continuous_on U f'"
and "0 \<in> U" and [simp]: "f 0 = 0"
and id: "f' 0 = id_blinfun"
obtains U' V g g' where "open U'" "U' \<subseteq> U" "0 \<in> U'" "open V" "0 \<in> V" "homeomorphism U' V f g"
"\<And>y. y \<in> V \<Longrightarrow> (g has_derivative (g' y)) (at y)"
"\<And>y. y \<in> V \<Longrightarrow> g' y = inv (blinfun_apply (f'(g y)))"
"\<And>y. y \<in> V \<Longrightarrow> bij (blinfun_apply (f'(g y)))"
proof -
obtain d1 where "cball 0 d1 \<subseteq> U" "d1 > 0"
using \<open>open U\<close> \<open>0 \<in> U\<close> open_contains_cball by blast
obtain d2 where d2: "\<And>x. \<lbrakk>x \<in> U; dist x 0 \<le> d2\<rbrakk> \<Longrightarrow> dist (f' x) (f' 0) < 1/2" "0 < d2"
using continuous_onE [OF contf, of 0 "1/2"] by (metis \<open>0 \<in> U\<close> half_gt_zero_iff zero_less_one)
obtain \<delta> where le: "\<And>x. norm x \<le> \<delta> \<Longrightarrow> dist (f' x) id_blinfun \<le> 1/2" and "0 < \<delta>"
and subU: "cball 0 \<delta> \<subseteq> U"
proof
show "min d1 d2 > 0"
by (simp add: \<open>0 < d1\<close> \<open>0 < d2\<close>)
show "cball 0 (min d1 d2) \<subseteq> U"
using \<open>cball 0 d1 \<subseteq> U\<close> by auto
show "dist (f' x) id_blinfun \<le> 1/2" if "norm x \<le> min d1 d2" for x
using \<open>cball 0 d1 \<subseteq> U\<close> d2 that id by fastforce
qed
let ?D = "cball 0 \<delta>"
define V:: "'a set" where "V \<equiv> ball 0 (\<delta>/2)"
have 4: "norm (f (x + h) - f x - h) \<le> 1/2 * norm h"
if "x \<in> ?D" "x+h \<in> ?D" for x h
proof -
let ?w = "\<lambda>x. f x - x"
have B: "\<And>x. x \<in> ?D \<Longrightarrow> onorm (blinfun_apply (f' x - id_blinfun)) \<le> 1/2"
by (metis dist_norm le mem_cball_0 norm_blinfun.rep_eq)
have "\<And>x. x \<in> ?D \<Longrightarrow> (?w has_derivative (blinfun_apply (f' x - id_blinfun))) (at x)"
by (rule derivative_eq_intros derf subsetD [OF subU] | force simp: blinfun.diff_left)+
then have Dw: "\<And>x. x \<in> ?D \<Longrightarrow> (?w has_derivative (blinfun_apply (f' x - id_blinfun))) (at x within ?D)"
using has_derivative_at_withinI by blast
have "norm (?w (x+h) - ?w x) \<le> (1/2) * norm h"
using differentiable_bound [OF convex_cball Dw B] that by fastforce
then show ?thesis
by (auto simp: algebra_simps)
qed
have for_g: "\<exists>!x. norm x < \<delta> \<and> f x = y" if y: "norm y < \<delta>/2" for y
proof -
let ?u = "\<lambda>x. x + (y - f x)"
have *: "norm (?u x) < \<delta>" if "x \<in> ?D" for x
proof -
have fxx: "norm (f x - x) \<le> \<delta>/2"
using 4 [of 0 x] \<open>0 < \<delta>\<close> \<open>f 0 = 0\<close> that by auto
have "norm (?u x) \<le> norm y + norm (f x - x)"
by (metis add.commute add_diff_eq norm_minus_commute norm_triangle_ineq)
also have "\<dots> < \<delta>/2 + \<delta>/2"
using fxx y by auto
finally show ?thesis
by simp
qed
have "\<exists>!x \<in> ?D. ?u x = x"
proof (rule banach_fix)
show "cball 0 \<delta> \<noteq> {}"
using \<open>0 < \<delta>\<close> by auto
show "(\<lambda>x. x + (y - f x)) ` cball 0 \<delta> \<subseteq> cball 0 \<delta>"
using * by force
have "dist (x + (y - f x)) (xh + (y - f xh)) * 2 \<le> dist x xh"
if "norm x \<le> \<delta>" and "norm xh \<le> \<delta>" for x xh
using that 4 [of x "xh-x"] by (auto simp: dist_norm norm_minus_commute algebra_simps)
then show "\<forall>x\<in>cball 0 \<delta>. \<forall>ya\<in>cball 0 \<delta>. dist (x + (y - f x)) (ya + (y - f ya)) \<le> (1/2) * dist x ya"
by auto
qed (auto simp: complete_eq_closed)
then show ?thesis
by (metis "*" add_cancel_right_right eq_iff_diff_eq_0 le_less mem_cball_0)
qed
define g where "g \<equiv> \<lambda>y. THE x. norm x < \<delta> \<and> f x = y"
have g: "norm (g y) < \<delta> \<and> f (g y) = y" if "norm y < \<delta>/2" for y
unfolding g_def using that theI' [OF for_g] by meson
then have fg[simp]: "f (g y) = y" if "y \<in> V" for y
using that by (auto simp: V_def)
have 5: "norm (g y' - g y) \<le> 2 * norm (y' - y)" if "y \<in> V" "y' \<in> V" for y y'
proof -
have no: "norm (g y) \<le> \<delta>" "norm (g y') \<le> \<delta>" and [simp]: "f (g y) = y"
using that g unfolding V_def by force+
have "norm (g y' - g y) \<le> norm (g y' - g y - (y' - y)) + norm (y' - y)"
by (simp add: add.commute norm_triangle_sub)
also have "\<dots> \<le> (1/2) * norm (g y' - g y) + norm (y' - y)"
using 4 [of "g y" "g y' - g y"] that no by (simp add: g norm_minus_commute V_def)
finally show ?thesis
by auto
qed
have contg: "continuous_on V g"
proof
fix y::'a and e::real
assume "0 < e" and y: "y \<in> V"
show "\<exists>d>0. \<forall>x'\<in>V. dist x' y < d \<longrightarrow> dist (g x') (g y) \<le> e"
proof (intro exI conjI ballI impI)
show "0 < e/2"
by (simp add: \<open>0 < e\<close>)
qed (use 5 y in \<open>force simp: dist_norm\<close>)
qed
show thesis
proof
define U' where "U' \<equiv> (f -` V) \<inter> ball 0 \<delta>"
have contf: "continuous_on U f"
using derf has_derivative_at_withinI by (fast intro: has_derivative_continuous_on)
then have "continuous_on (ball 0 \<delta>) f"
by (meson ball_subset_cball continuous_on_subset subU)
then show "open U'"
by (simp add: U'_def V_def Int_commute continuous_open_preimage)
show "0 \<in> U'" "U' \<subseteq> U" "open V" "0 \<in> V"
using \<open>0 < \<delta>\<close> subU by (auto simp: U'_def V_def)
show hom: "homeomorphism U' V f g"
proof
show "continuous_on U' f"
using \<open>U' \<subseteq> U\<close> contf continuous_on_subset by blast
show "continuous_on V g"
using contg by blast
show "f ` U' \<subseteq> V"
using U'_def by blast
show "g ` V \<subseteq> U'"
by (simp add: U'_def V_def g image_subset_iff)
show "g (f x) = x" if "x \<in> U'" for x
by (metis that fg Int_iff U'_def V_def for_g g mem_ball_0 vimage_eq)
show "f (g y) = y" if "y \<in> V" for y
using that by (simp add: g V_def)
qed
show bij: "bij (blinfun_apply (f'(g y)))" if "y \<in> V" for y
proof -
have inj: "inj (blinfun_apply (f' (g y)))"
proof (rule linear_injective_contraction)
show "linear (blinfun_apply (f' (g y)))"
using blinfun.bounded_linear_right bounded_linear_def by blast
next
fix x
have "norm (blinfun_apply (f' (g y)) x - x) = norm (blinfun_apply (f' (g y) - id_blinfun) x)"
by (simp add: blinfun.diff_left)
also have "\<dots> \<le> norm (f' (g y) - id_blinfun) * norm x"
by (rule norm_blinfun)
also have "\<dots> \<le> (1/2) * norm x"
proof (rule mult_right_mono)
show "norm (f' (g y) - id_blinfun) \<le> 1/2"
using that g [of y] le by (auto simp: V_def dist_norm)
qed auto
finally show "norm (blinfun_apply (f' (g y)) x - x) \<le> (1/2) * norm x" .
qed auto
moreover
have "surj (blinfun_apply (f' (g y)))"
using blinfun.bounded_linear_right bounded_linear_def
by (blast intro!: linear_inj_imp_surj [OF _ inj])
ultimately show ?thesis
using bijI by blast
qed
define g' where "g' \<equiv> \<lambda>y. inv (blinfun_apply (f'(g y)))"
show "(g has_derivative g' y) (at y)" if "y \<in> V" for y
proof -
have gy: "g y \<in> U"
using g subU that unfolding V_def by fastforce
obtain e where e: "\<And>h. f (g y + h) = y + blinfun_apply (f' (g y)) h + e h"
and e0: "(\<lambda>h. norm (e h) / norm h) \<midarrow>0\<rightarrow> 0"
using iffD1 [OF has_derivative_iff_Ex derf [OF gy]] \<open>y \<in> V\<close> by auto
have [simp]: "e 0 = 0"
using e [of 0] that by simp
let ?INV = "inv (blinfun_apply (f' (g y)))"
have inj: "inj (blinfun_apply (f' (g y)))"
using bij bij_betw_def that by blast
have "(g has_derivative g' y) (at y within V)"
unfolding has_derivative_at_within_iff_Ex [OF \<open>y \<in> V\<close> \<open>open V\<close>]
proof
show blinv: "bounded_linear (g' y)"
unfolding g'_def using derf gy inj inj_linear_imp_inv_bounded_linear by blast
define eg where "eg \<equiv> \<lambda>k. - ?INV (e (g (y+k) - g y))"
have "g (y+k) = g y + g' y k + eg k" if "y + k \<in> V" for k
proof -
have "?INV k = ?INV (blinfun_apply (f' (g y)) (g (y+k) - g y) + e (g (y+k) - g y))"
using e [of "g(y+k) - g y"] that by simp
then have "g (y+k) = g y + ?INV k - ?INV (e (g (y+k) - g y))"
using inj blinv by (simp add: linear_simps g'_def)
then show ?thesis
by (auto simp: eg_def g'_def)
qed
moreover have "(\<lambda>k. norm (eg k) / norm k) \<midarrow>0\<rightarrow> 0"
proof (rule Lim_null_comparison)
let ?g = "\<lambda>k. 2 * onorm ?INV * norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)"
show "\<forall>\<^sub>F k in at 0. norm (norm (eg k) / norm k) \<le> ?g k"
unfolding eventually_at_topological
proof (intro exI conjI ballI impI)
show "open ((+)(-y) ` V)"
using \<open>open V\<close> open_translation by blast
show "0 \<in> (+)(-y) ` V"
by (simp add: that)
show "norm (norm (eg k) / norm k) \<le> 2 * onorm (inv (blinfun_apply (f' (g y)))) * norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)"
if "k \<in> (+)(-y) ` V" "k \<noteq> 0" for k
proof -
have "y+k \<in> V"
using that by auto
have "norm (norm (eg k) / norm k) \<le> onorm ?INV * norm (e (g (y+k) - g y)) / norm k"
using blinv g'_def onorm by (force simp: eg_def divide_simps)
also have "\<dots> = (norm (g (y+k) - g y) / norm k) * (onorm ?INV * (norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)))"
by (simp add: divide_simps)
also have "\<dots> \<le> 2 * (onorm ?INV * (norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)))"
apply (rule mult_right_mono)
using 5 [of y "y+k"] \<open>y \<in> V\<close> \<open>y + k \<in> V\<close> onorm_pos_le [OF blinv]
apply (auto simp: divide_simps zero_le_mult_iff zero_le_divide_iff g'_def)
done
finally show "norm (norm (eg k) / norm k) \<le> 2 * onorm ?INV * norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)"
by simp
qed
qed
have 1: "(\<lambda>h. norm (e h) / norm h) \<midarrow>0\<rightarrow> (norm (e 0) / norm 0)"
using e0 by auto
have 2: "(\<lambda>k. g (y+k) - g y) \<midarrow>0\<rightarrow> 0"
using contg \<open>open V\<close> \<open>y \<in> V\<close> LIM_offset_zero_iff LIM_zero_iff at_within_open continuous_on_def by fastforce
from tendsto_compose [OF 1 2, simplified]
have "(\<lambda>k. norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)) \<midarrow>0\<rightarrow> 0" .
from tendsto_mult_left [OF this] show "?g \<midarrow>0\<rightarrow> 0" by auto
qed
ultimately show "\<exists>e. (\<forall>k. y + k \<in> V \<longrightarrow> g (y+k) = g y + g' y k + e k) \<and> (\<lambda>k. norm (e k) / norm k) \<midarrow>0\<rightarrow> 0"
by blast
qed
then show ?thesis
by (metis \<open>open V\<close> at_within_open that)
qed
show "g' y = inv (blinfun_apply (f' (g y)))"
if "y \<in> V" for y
by (simp add: g'_def)
qed
qed
text\<open>We need all this to justify the scaling and translations.\<close>
theorem inverse_function_theorem:
fixes f::"'a::euclidean_space \<Rightarrow> 'a"
and f'::"'a \<Rightarrow> ('a \<Rightarrow>\<^sub>L 'a)"
assumes "open U"
and derf: "\<And>x. x \<in> U \<Longrightarrow> (f has_derivative (blinfun_apply (f' x))) (at x)"
and contf: "continuous_on U f'"
and "x0 \<in> U"
and invf: "invf o\<^sub>L f' x0 = id_blinfun"
obtains U' V g g' where "open U'" "U' \<subseteq> U" "x0 \<in> U'" "open V" "f x0 \<in> V" "homeomorphism U' V f g"
"\<And>y. y \<in> V \<Longrightarrow> (g has_derivative (g' y)) (at y)"
"\<And>y. y \<in> V \<Longrightarrow> g' y = inv (blinfun_apply (f'(g y)))"
"\<And>y. y \<in> V \<Longrightarrow> bij (blinfun_apply (f'(g y)))"
proof -
have apply1 [simp]: "\<And>i. blinfun_apply invf (blinfun_apply (f' x0) i) = i"
by (metis blinfun_apply_blinfun_compose blinfun_apply_id_blinfun invf)
have apply2 [simp]: "\<And>i. blinfun_apply (f' x0) (blinfun_apply invf i) = i"
by (metis apply1 bij_inv_eq_iff blinfun_bij1 invf)
have [simp]: "(range (blinfun_apply invf)) = UNIV"
using apply1 surjI by blast
let ?f = "invf \<circ> (\<lambda>x. (f \<circ> (+)x0)x - f x0)"
let ?f' = "\<lambda>x. invf o\<^sub>L (f' (x + x0))"
obtain U' V g g' where "open U'" and U': "U' \<subseteq> (+)(-x0) ` U" "0 \<in> U'"
and "open V" "0 \<in> V" and hom: "homeomorphism U' V ?f g"
and derg: "\<And>y. y \<in> V \<Longrightarrow> (g has_derivative (g' y)) (at y)"
and g': "\<And>y. y \<in> V \<Longrightarrow> g' y = inv (?f'(g y))"
and bij: "\<And>y. y \<in> V \<Longrightarrow> bij (?f'(g y))"
proof (rule inverse_function_theorem_scaled [of "(+)(-x0) ` U" ?f "?f'"])
show ope: "open ((+) (- x0) ` U)"
using \<open>open U\<close> open_translation by blast
show "(?f has_derivative blinfun_apply (?f' x)) (at x)"
if "x \<in> (+) (- x0) ` U" for x
using that
apply clarify
apply (rule derf derivative_eq_intros | simp add: blinfun_compose.rep_eq)+
done
have YY: "(\<lambda>x. f' (x + x0)) \<midarrow>u-x0\<rightarrow> f' u"
if "f' \<midarrow>u\<rightarrow> f' u" "u \<in> U" for u
using that LIM_offset [where k = x0] by (auto simp: algebra_simps)
then have "continuous_on ((+) (- x0) ` U) (\<lambda>x. f' (x + x0))"
using contf \<open>open U\<close> Lim_at_imp_Lim_at_within
by (fastforce simp: continuous_on_def at_within_open_NO_MATCH ope)
then show "continuous_on ((+) (- x0) ` U) ?f'"
by (intro continuous_intros) simp
qed (auto simp: invf \<open>x0 \<in> U\<close>)
show thesis
proof
let ?U' = "(+)x0 ` U'"
let ?V = "((+)(f x0) \<circ> f' x0) ` V"
let ?g = "(+)x0 \<circ> g \<circ> invf \<circ> (+)(- f x0)"
let ?g' = "\<lambda>y. inv (blinfun_apply (f' (?g y)))"
show oU': "open ?U'"
by (simp add: \<open>open U'\<close> open_translation)
show subU: "?U' \<subseteq> U"
using ComplI \<open>U' \<subseteq> (+) (- x0) ` U\<close> by auto
show "x0 \<in> ?U'"
by (simp add: \<open>0 \<in> U'\<close>)
show "open ?V"
using blinfun_bij2 [OF invf]
by (metis \<open>open V\<close> bij_is_surj blinfun.bounded_linear_right bounded_linear_def image_comp open_surjective_linear_image open_translation)
show "f x0 \<in> ?V"
using \<open>0 \<in> V\<close> image_iff by fastforce
show "homeomorphism ?U' ?V f ?g"
proof
show "continuous_on ?U' f"
by (meson subU continuous_on_eq_continuous_at derf has_derivative_continuous oU' subsetD)
have "?f ` U' \<subseteq> V"
using hom homeomorphism_image1 by blast
then show "f ` ?U' \<subseteq> ?V"
unfolding image_subset_iff
by (clarsimp simp: image_def) (metis apply2 add.commute diff_add_cancel)
show "?g ` ?V \<subseteq> ?U'"
using hom invf by (auto simp: image_def homeomorphism_def)
show "?g (f x) = x"
if "x \<in> ?U'" for x
using that hom homeomorphism_apply1 by fastforce
have "continuous_on V g"
using hom homeomorphism_def by blast
then show "continuous_on ?V ?g"
by (intro continuous_intros) (auto elim!: continuous_on_subset)
have fg: "?f (g x) = x" if "x \<in> V" for x
using hom homeomorphism_apply2 that by blast
show "f (?g y) = y"
if "y \<in> ?V" for y
using that fg by (simp add: image_iff) (metis apply2 add.commute diff_add_cancel)
qed
show "(?g has_derivative ?g' y) (at y)" "bij (blinfun_apply (f' (?g y)))"
if "y \<in> ?V" for y
proof -
have 1: "bij (blinfun_apply invf)"
using blinfun_bij1 invf by blast
then have 2: "bij (blinfun_apply (f' (x0 + g x)))" if "x \<in> V" for x
by (metis add.commute bij bij_betw_comp_iff2 blinfun_compose.rep_eq that top_greatest)
then show "bij (blinfun_apply (f' (?g y)))"
using that by auto
have "g' x \<circ> blinfun_apply invf = inv (blinfun_apply (f' (x0 + g x)))"
if "x \<in> V" for x
using that
by (simp add: g' o_inv_distrib blinfun_compose.rep_eq 1 2 add.commute bij_is_inj flip: o_assoc)
then show "(?g has_derivative ?g' y) (at y)"
using that invf
by clarsimp (rule derg derivative_eq_intros | simp flip: id_def)+
qed
qed auto
qed
subsection\<^marker>\<open>tag unimportant\<close> \<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 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 \<circ> 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 \<circ> (\<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"
and fd: "\<And>x. x \<in> {a..c} - S \<Longrightarrow> f differentiable at x within {a..c}"
and gd: "\<And>x. x \<in> {c..b} - T \<Longrightarrow> 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
proof (rule differentiable_transform_within [where d = "dist x c"])
have "f differentiable at x"
using x le fd [of x] at_within_interior [of x "{a..c}"] by simp
then show "f differentiable at x within {a..b}"
by (simp add: differentiable_at_withinI)
qed (use x le st dist_real_def in auto)
next
case ge show ?diff_fg
proof (rule differentiable_transform_within [where d = "dist x c"])
have "g differentiable at x"
using x ge gd [of x] at_within_interior [of x "{c..b}"] by simp
then show "g differentiable at x within {a..b}"
by (simp add: differentiable_at_withinI)
qed (use x ge st dist_real_def in auto)
qed
}
then have "\<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)
subsection\<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\ldots{} [namely] continuity plus differentiability except on a
finite set\ldots{} [Our] underlying theory of integration is the Kurzweil-Henstock theory. In contrast to
the Riemann or Lebesgue theory (but in common with a simple notion based on antiderivatives), this
can integrate all derivatives.''
"Formalizing basic complex analysis." From Insight to Proof: Festschrift in Honour of Andrzej Trybulec.
Studies in Logic, Grammar and Rhetoric 10.23 (2007): 151-165.
And indeed he does not assume that his derivatives are continuous, but the penalty is unreasonably
difficult proofs concerning winding numbers. We need a self-contained and straightforward theorem
asserting that all derivatives can be integrated before we can adopt Harrison's choice.\<close>
definition\<^marker>\<open>tag important\<close> 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))"
(is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
unfolding C1_differentiable_on_def
by (metis (no_types, lifting) continuous_on_eq differentiableI_vector vector_derivative_at)
next
assume ?rhs
then show ?lhs
using C1_differentiable_on_def vector_derivative_works by fastforce
qed
lemma C1_differentiable_on_subset:
"f C1_differentiable_on T \<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:
assumes fg: "f C1_differentiable_on S" "g C1_differentiable_on (f ` S)" and fin: "\<And>x. finite (S \<inter> f-`{x})"
shows "(g \<circ> f) C1_differentiable_on S"
proof -
have "\<And>x. x \<in> S \<Longrightarrow> g \<circ> f differentiable at x"
by (meson C1_differentiable_on_eq assms differentiable_chain_at imageI)
moreover have "continuous_on S (\<lambda>x. vector_derivative (g \<circ> f) (at x))"
proof (rule continuous_on_eq [of _ "\<lambda>x. vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x))"])
show "continuous_on S (\<lambda>x. vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x)))"
using fg
apply (clarsimp simp add: C1_differentiable_on_eq)
apply (rule Limits.continuous_on_scaleR, assumption)
by (metis (mono_tags, lifting) continuous_at_imp_continuous_on continuous_on_compose continuous_on_cong differentiable_imp_continuous_within o_def)
show "\<And>x. x \<in> S \<Longrightarrow> vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x)) = vector_derivative (g \<circ> f) (at x)"
by (metis (mono_tags, hide_lams) C1_differentiable_on_eq fg imageI vector_derivative_chain_at)
qed
ultimately show ?thesis
by (simp add: C1_differentiable_on_eq)
qed
lemma C1_diff_imp_diff: "f C1_differentiable_on S \<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)
lemma C1_differentiable_on_const [simp, derivative_intros]: "(\<lambda>z. a) C1_differentiable_on S"
by (auto simp: C1_differentiable_on_eq)
lemma C1_differentiable_on_add [simp, derivative_intros]:
"f C1_differentiable_on S \<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\<^marker>\<open>tag important\<close> 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_withinI)
lemma piecewise_C1_differentiable_compose:
assumes fg: "f piecewise_C1_differentiable_on S" "g piecewise_C1_differentiable_on (f ` S)" and fin: "\<And>x. finite (S \<inter> f-`{x})"
shows "(g \<circ> f) piecewise_C1_differentiable_on S"
proof -
have "continuous_on S (\<lambda>x. g (f x))"
by (metis continuous_on_compose2 fg order_refl piecewise_C1_differentiable_on_def)
moreover have "\<exists>T. finite T \<and> g \<circ> f C1_differentiable_on S - T"
proof -
obtain F where "finite F" and F: "f C1_differentiable_on S - F" and f: "f piecewise_C1_differentiable_on S"
using fg by (auto simp: piecewise_C1_differentiable_on_def)
obtain G where "finite G" and G: "g C1_differentiable_on f ` S - G" and g: "g piecewise_C1_differentiable_on f ` S"
using fg by (auto simp: piecewise_C1_differentiable_on_def)
show ?thesis
proof (intro exI conjI)
show "finite (F \<union> (\<Union>x\<in>G. S \<inter> f-`{x}))"
using fin by (auto simp only: Int_Union \<open>finite F\<close> \<open>finite G\<close> finite_UN finite_imageI)
show "g \<circ> f C1_differentiable_on S - (F \<union> (\<Union>x\<in>G. S \<inter> f -` {x}))"
apply (rule C1_differentiable_compose)
apply (blast intro: C1_differentiable_on_subset [OF F])
apply (blast intro: C1_differentiable_on_subset [OF G])
by (simp add: C1_differentiable_on_subset G Diff_Int_distrib2 fin)
qed
qed
ultimately show ?thesis
by (simp add: piecewise_C1_differentiable_on_def)
qed
lemma piecewise_C1_differentiable_on_subset:
"f piecewise_C1_differentiable_on S \<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 \<circ> (\<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)
next
case False
have *: "\<And>x. finite (S \<inter> {y. m * y + c = x})"
using False not_finite_existsD by fastforce
show ?thesis
apply (rule piecewise_C1_differentiable_compose [OF C1_differentiable_imp_piecewise])
apply (rule * assms derivative_intros | simp add: False vimage_def)+
done
qed
lemma piecewise_C1_differentiable_cases:
fixes c::real
assumes "f piecewise_C1_differentiable_on {a..c}"
"g piecewise_C1_differentiable_on {c..b}"
"a \<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))"
proof -
have "((\<lambda>x. if x \<le> c then f x else g x) has_vector_derivative vector_derivative f (at x)) (at x)"
if "a < x" "x < c" "x \<notin> S" for x
proof -
have f: "f differentiable at x"
by (meson C1_differentiable_on_eq Diff_iff atLeastAtMost_iff less_eq_real_def st(1) that)
show ?thesis
using that
apply (rule_tac f=f and d="dist x c" in has_vector_derivative_transform_within)
apply (auto simp: dist_norm vector_derivative_works [symmetric] f)
done
qed
then show ?thesis
by (metis (no_types, lifting) continuous_on_eq [OF fcon] DiffE greaterThanLessThan_iff vector_derivative_at)
qed
moreover have "continuous_on ({c<..<b} - T) (\<lambda>x. vector_derivative (\<lambda>x. if x \<le> c then f x else g x) (at x))"
proof -
have "((\<lambda>x. if x \<le> c then f x else g x) has_vector_derivative vector_derivative g (at x)) (at x)"
if "c < x" "x < b" "x \<notin> T" for x
proof -
have g: "g differentiable at x"
by (metis C1_differentiable_on_eq DiffD1 DiffI atLeastAtMost_diff_ends greaterThanLessThan_iff st(2) that)
show ?thesis
using that
apply (rule_tac f=g and d="dist x c" in has_vector_derivative_transform_within)
apply (auto simp: dist_norm vector_derivative_works [symmetric] g)
done
qed
then show ?thesis
by (metis (no_types, lifting) continuous_on_eq [OF gcon] DiffE greaterThanLessThan_iff vector_derivative_at)
qed
ultimately have "continuous_on ({a<..<b} - insert c (S \<union> T))
(\<lambda>x. vector_derivative (\<lambda>x. if x \<le> c then f x else g x) (at x))"
by (rule continuous_on_subset [OF continuous_on_open_Un], auto)
} note * = this
have "continuous_on ({a<..<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)
end
diff --git a/src/HOL/Analysis/Elementary_Metric_Spaces.thy b/src/HOL/Analysis/Elementary_Metric_Spaces.thy
--- a/src/HOL/Analysis/Elementary_Metric_Spaces.thy
+++ b/src/HOL/Analysis/Elementary_Metric_Spaces.thy
@@ -1,3311 +1,3311 @@
(* Author: L C Paulson, University of Cambridge
Author: Amine Chaieb, University of Cambridge
Author: Robert Himmelmann, TU Muenchen
Author: Brian Huffman, Portland State University
*)
section \<open>Elementary Metric Spaces\<close>
theory Elementary_Metric_Spaces
imports
Abstract_Topology_2
Metric_Arith
begin
subsection \<open>Open and closed balls\<close>
definition\<^marker>\<open>tag important\<close> ball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set"
where "ball x e = {y. dist x y < e}"
definition\<^marker>\<open>tag important\<close> cball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set"
where "cball x e = {y. dist x y \<le> e}"
definition\<^marker>\<open>tag important\<close> sphere :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set"
where "sphere x e = {y. dist x y = e}"
lemma mem_ball [simp, metric_unfold]: "y \<in> ball x e \<longleftrightarrow> dist x y < e"
by (simp add: ball_def)
lemma mem_cball [simp, metric_unfold]: "y \<in> cball x e \<longleftrightarrow> dist x y \<le> e"
by (simp add: cball_def)
lemma mem_sphere [simp]: "y \<in> sphere x e \<longleftrightarrow> dist x y = e"
by (simp add: sphere_def)
lemma ball_trivial [simp]: "ball x 0 = {}"
by (simp add: ball_def)
lemma cball_trivial [simp]: "cball x 0 = {x}"
by (simp add: cball_def)
lemma sphere_trivial [simp]: "sphere x 0 = {x}"
by (simp add: sphere_def)
lemma disjoint_ballI: "dist x y \<ge> r+s \<Longrightarrow> ball x r \<inter> ball y s = {}"
using dist_triangle_less_add not_le by fastforce
lemma disjoint_cballI: "dist x y > r + s \<Longrightarrow> cball x r \<inter> cball y s = {}"
by (metis add_mono disjoint_iff_not_equal dist_triangle2 dual_order.trans leD mem_cball)
lemma sphere_empty [simp]: "r < 0 \<Longrightarrow> sphere a r = {}"
for a :: "'a::metric_space"
by auto
lemma centre_in_ball [simp]: "x \<in> ball x e \<longleftrightarrow> 0 < e"
by simp
lemma centre_in_cball [simp]: "x \<in> cball x e \<longleftrightarrow> 0 \<le> e"
by simp
lemma ball_subset_cball [simp, intro]: "ball x e \<subseteq> cball x e"
by (simp add: subset_eq)
lemma mem_ball_imp_mem_cball: "x \<in> ball y e \<Longrightarrow> x \<in> cball y e"
by (auto)
lemma sphere_cball [simp,intro]: "sphere z r \<subseteq> cball z r"
by force
lemma cball_diff_sphere: "cball a r - sphere a r = ball a r"
by auto
lemma subset_ball[intro]: "d \<le> e \<Longrightarrow> ball x d \<subseteq> ball x e"
by (simp add: subset_eq)
lemma subset_cball[intro]: "d \<le> e \<Longrightarrow> cball x d \<subseteq> cball x e"
by (simp add: subset_eq)
lemma mem_ball_leI: "x \<in> ball y e \<Longrightarrow> e \<le> f \<Longrightarrow> x \<in> ball y f"
by (auto)
lemma mem_cball_leI: "x \<in> cball y e \<Longrightarrow> e \<le> f \<Longrightarrow> x \<in> cball y f"
by (auto)
lemma cball_trans: "y \<in> cball z b \<Longrightarrow> x \<in> cball y a \<Longrightarrow> x \<in> cball z (b + a)"
by metric
lemma ball_max_Un: "ball a (max r s) = ball a r \<union> ball a s"
by (simp add: set_eq_iff) arith
lemma ball_min_Int: "ball a (min r s) = ball a r \<inter> ball a s"
by (simp add: set_eq_iff)
lemma cball_max_Un: "cball a (max r s) = cball a r \<union> cball a s"
by (simp add: set_eq_iff) arith
lemma cball_min_Int: "cball a (min r s) = cball a r \<inter> cball a s"
by (simp add: set_eq_iff)
lemma cball_diff_eq_sphere: "cball a r - ball a r = sphere a r"
by (auto simp: cball_def ball_def dist_commute)
lemma open_ball [intro, simp]: "open (ball x e)"
proof -
have "open (dist x -` {..<e})"
by (intro open_vimage open_lessThan continuous_intros)
also have "dist x -` {..<e} = ball x e"
by auto
finally show ?thesis .
qed
lemma open_contains_ball: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. ball x e \<subseteq> S)"
- by (simp add: open_dist subset_eq mem_ball Ball_def dist_commute)
+ by (simp add: open_dist subset_eq Ball_def dist_commute)
lemma openI [intro?]: "(\<And>x. x\<in>S \<Longrightarrow> \<exists>e>0. ball x e \<subseteq> S) \<Longrightarrow> open S"
by (auto simp: open_contains_ball)
lemma openE[elim?]:
assumes "open S" "x\<in>S"
obtains e where "e>0" "ball x e \<subseteq> S"
using assms unfolding open_contains_ball by auto
lemma open_contains_ball_eq: "open S \<Longrightarrow> x\<in>S \<longleftrightarrow> (\<exists>e>0. ball x e \<subseteq> S)"
by (metis open_contains_ball subset_eq centre_in_ball)
lemma ball_eq_empty[simp]: "ball x e = {} \<longleftrightarrow> e \<le> 0"
unfolding mem_ball set_eq_iff
by (simp add: not_less) metric
lemma ball_empty: "e \<le> 0 \<Longrightarrow> ball x e = {}" by simp
lemma closed_cball [iff]: "closed (cball x e)"
proof -
have "closed (dist x -` {..e})"
by (intro closed_vimage closed_atMost continuous_intros)
also have "dist x -` {..e} = cball x e"
by auto
finally show ?thesis .
qed
lemma open_contains_cball: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. cball x e \<subseteq> S)"
proof -
{
fix x and e::real
assume "x\<in>S" "e>0" "ball x e \<subseteq> S"
then have "\<exists>d>0. cball x d \<subseteq> S" unfolding subset_eq by (rule_tac x="e/2" in exI, auto)
}
moreover
{
fix x and e::real
assume "x\<in>S" "e>0" "cball x e \<subseteq> S"
then have "\<exists>d>0. ball x d \<subseteq> S"
unfolding subset_eq
apply (rule_tac x="e/2" in exI, auto)
done
}
ultimately show ?thesis
unfolding open_contains_ball by auto
qed
lemma open_contains_cball_eq: "open S \<Longrightarrow> (\<forall>x. x \<in> S \<longleftrightarrow> (\<exists>e>0. cball x e \<subseteq> S))"
by (metis open_contains_cball subset_eq order_less_imp_le centre_in_cball)
lemma eventually_nhds_ball: "d > 0 \<Longrightarrow> eventually (\<lambda>x. x \<in> ball z d) (nhds z)"
by (rule eventually_nhds_in_open) simp_all
lemma eventually_at_ball: "d > 0 \<Longrightarrow> eventually (\<lambda>t. t \<in> ball z d \<and> t \<in> A) (at z within A)"
unfolding eventually_at by (intro exI[of _ d]) (simp_all add: dist_commute)
lemma eventually_at_ball': "d > 0 \<Longrightarrow> eventually (\<lambda>t. t \<in> ball z d \<and> t \<noteq> z \<and> t \<in> A) (at z within A)"
unfolding eventually_at by (intro exI[of _ d]) (simp_all add: dist_commute)
lemma at_within_ball: "e > 0 \<Longrightarrow> dist x y < e \<Longrightarrow> at y within ball x e = at y"
by (subst at_within_open) auto
lemma atLeastAtMost_eq_cball:
fixes a b::real
shows "{a .. b} = cball ((a + b)/2) ((b - a)/2)"
by (auto simp: dist_real_def field_simps)
lemma cball_eq_atLeastAtMost:
fixes a b::real
shows "cball a b = {a - b .. a + b}"
by (auto simp: dist_real_def)
lemma greaterThanLessThan_eq_ball:
fixes a b::real
shows "{a <..< b} = ball ((a + b)/2) ((b - a)/2)"
by (auto simp: dist_real_def field_simps)
lemma ball_eq_greaterThanLessThan:
fixes a b::real
shows "ball a b = {a - b <..< a + b}"
by (auto simp: dist_real_def)
lemma interior_ball [simp]: "interior (ball x e) = ball x e"
by (simp add: interior_open)
lemma cball_eq_empty [simp]: "cball x e = {} \<longleftrightarrow> e < 0"
apply (simp add: set_eq_iff not_le)
apply (metis zero_le_dist dist_self order_less_le_trans)
done
lemma cball_empty [simp]: "e < 0 \<Longrightarrow> cball x e = {}"
by simp
lemma cball_sing:
fixes x :: "'a::metric_space"
shows "e = 0 \<Longrightarrow> cball x e = {x}"
by (auto simp: set_eq_iff)
lemma ball_divide_subset: "d \<ge> 1 \<Longrightarrow> ball x (e/d) \<subseteq> ball x e"
apply (cases "e \<le> 0")
apply (simp add: ball_empty field_split_simps)
apply (rule subset_ball)
apply (simp add: field_split_simps)
done
lemma ball_divide_subset_numeral: "ball x (e / numeral w) \<subseteq> ball x e"
using ball_divide_subset one_le_numeral by blast
lemma cball_divide_subset: "d \<ge> 1 \<Longrightarrow> cball x (e/d) \<subseteq> cball x e"
apply (cases "e < 0")
apply (simp add: field_split_simps)
apply (rule subset_cball)
apply (metis div_by_1 frac_le not_le order_refl zero_less_one)
done
lemma cball_divide_subset_numeral: "cball x (e / numeral w) \<subseteq> cball x e"
using cball_divide_subset one_le_numeral by blast
lemma cball_scale:
assumes "a \<noteq> 0"
shows "(\<lambda>x. a *\<^sub>R x) ` cball c r = cball (a *\<^sub>R c :: 'a :: real_normed_vector) (\<bar>a\<bar> * r)"
proof -
have 1: "(\<lambda>x. a *\<^sub>R x) ` cball c r \<subseteq> cball (a *\<^sub>R c) (\<bar>a\<bar> * r)" if "a \<noteq> 0" for a r and c :: 'a
proof safe
fix x
assume x: "x \<in> cball c r"
have "dist (a *\<^sub>R c) (a *\<^sub>R x) = norm (a *\<^sub>R c - a *\<^sub>R x)"
by (auto simp: dist_norm)
also have "a *\<^sub>R c - a *\<^sub>R x = a *\<^sub>R (c - x)"
by (simp add: algebra_simps)
finally show "a *\<^sub>R x \<in> cball (a *\<^sub>R c) (\<bar>a\<bar> * r)"
using that x by (auto simp: dist_norm)
qed
have "cball (a *\<^sub>R c) (\<bar>a\<bar> * r) = (\<lambda>x. a *\<^sub>R x) ` (\<lambda>x. inverse a *\<^sub>R x) ` cball (a *\<^sub>R c) (\<bar>a\<bar> * r)"
unfolding image_image using assms by simp
also have "\<dots> \<subseteq> (\<lambda>x. a *\<^sub>R x) ` cball (inverse a *\<^sub>R (a *\<^sub>R c)) (\<bar>inverse a\<bar> * (\<bar>a\<bar> * r))"
using assms by (intro image_mono 1) auto
also have "\<dots> = (\<lambda>x. a *\<^sub>R x) ` cball c r"
using assms by (simp add: algebra_simps)
finally have "cball (a *\<^sub>R c) (\<bar>a\<bar> * r) \<subseteq> (\<lambda>x. a *\<^sub>R x) ` cball c r" .
moreover from assms have "(\<lambda>x. a *\<^sub>R x) ` cball c r \<subseteq> cball (a *\<^sub>R c) (\<bar>a\<bar> * r)"
by (intro 1) auto
ultimately show ?thesis by blast
qed
lemma ball_scale:
assumes "a \<noteq> 0"
shows "(\<lambda>x. a *\<^sub>R x) ` ball c r = ball (a *\<^sub>R c :: 'a :: real_normed_vector) (\<bar>a\<bar> * r)"
proof -
have 1: "(\<lambda>x. a *\<^sub>R x) ` ball c r \<subseteq> ball (a *\<^sub>R c) (\<bar>a\<bar> * r)" if "a \<noteq> 0" for a r and c :: 'a
proof safe
fix x
assume x: "x \<in> ball c r"
have "dist (a *\<^sub>R c) (a *\<^sub>R x) = norm (a *\<^sub>R c - a *\<^sub>R x)"
by (auto simp: dist_norm)
also have "a *\<^sub>R c - a *\<^sub>R x = a *\<^sub>R (c - x)"
by (simp add: algebra_simps)
finally show "a *\<^sub>R x \<in> ball (a *\<^sub>R c) (\<bar>a\<bar> * r)"
using that x by (auto simp: dist_norm)
qed
have "ball (a *\<^sub>R c) (\<bar>a\<bar> * r) = (\<lambda>x. a *\<^sub>R x) ` (\<lambda>x. inverse a *\<^sub>R x) ` ball (a *\<^sub>R c) (\<bar>a\<bar> * r)"
unfolding image_image using assms by simp
also have "\<dots> \<subseteq> (\<lambda>x. a *\<^sub>R x) ` ball (inverse a *\<^sub>R (a *\<^sub>R c)) (\<bar>inverse a\<bar> * (\<bar>a\<bar> * r))"
using assms by (intro image_mono 1) auto
also have "\<dots> = (\<lambda>x. a *\<^sub>R x) ` ball c r"
using assms by (simp add: algebra_simps)
finally have "ball (a *\<^sub>R c) (\<bar>a\<bar> * r) \<subseteq> (\<lambda>x. a *\<^sub>R x) ` ball c r" .
moreover from assms have "(\<lambda>x. a *\<^sub>R x) ` ball c r \<subseteq> ball (a *\<^sub>R c) (\<bar>a\<bar> * r)"
by (intro 1) auto
ultimately show ?thesis by blast
qed
subsection \<open>Limit Points\<close>
lemma islimpt_approachable:
fixes x :: "'a::metric_space"
shows "x islimpt S \<longleftrightarrow> (\<forall>e>0. \<exists>x'\<in>S. x' \<noteq> x \<and> dist x' x < e)"
unfolding islimpt_iff_eventually eventually_at by fast
lemma islimpt_approachable_le: "x islimpt S \<longleftrightarrow> (\<forall>e>0. \<exists>x'\<in> S. x' \<noteq> x \<and> dist x' x \<le> e)"
for x :: "'a::metric_space"
unfolding islimpt_approachable
using approachable_lt_le [where f="\<lambda>y. dist y x" and P="\<lambda>y. y \<notin> S \<or> y = x",
THEN arg_cong [where f=Not]]
by (simp add: Bex_def conj_commute conj_left_commute)
lemma limpt_of_limpts: "x islimpt {y. y islimpt S} \<Longrightarrow> x islimpt S"
for x :: "'a::metric_space"
apply (clarsimp simp add: islimpt_approachable)
apply (drule_tac x="e/2" in spec)
apply (auto simp: simp del: less_divide_eq_numeral1)
apply (drule_tac x="dist x' x" in spec)
apply (auto simp del: less_divide_eq_numeral1)
apply metric
done
lemma closed_limpts: "closed {x::'a::metric_space. x islimpt S}"
using closed_limpt limpt_of_limpts by blast
lemma limpt_of_closure: "x islimpt closure S \<longleftrightarrow> x islimpt S"
for x :: "'a::metric_space"
by (auto simp: closure_def islimpt_Un dest: limpt_of_limpts)
lemma islimpt_eq_infinite_ball: "x islimpt S \<longleftrightarrow> (\<forall>e>0. infinite(S \<inter> ball x e))"
apply (simp add: islimpt_eq_acc_point, safe)
apply (metis Int_commute open_ball centre_in_ball)
by (metis open_contains_ball Int_mono finite_subset inf_commute subset_refl)
lemma islimpt_eq_infinite_cball: "x islimpt S \<longleftrightarrow> (\<forall>e>0. infinite(S \<inter> cball x e))"
apply (simp add: islimpt_eq_infinite_ball, safe)
apply (meson Int_mono ball_subset_cball finite_subset order_refl)
by (metis open_ball centre_in_ball finite_Int inf.absorb_iff2 inf_assoc open_contains_cball_eq)
subsection \<open>Perfect Metric Spaces\<close>
lemma perfect_choose_dist: "0 < r \<Longrightarrow> \<exists>a. a \<noteq> x \<and> dist a x < r"
for x :: "'a::{perfect_space,metric_space}"
using islimpt_UNIV [of x] by (simp add: islimpt_approachable)
lemma cball_eq_sing:
fixes x :: "'a::{metric_space,perfect_space}"
shows "cball x e = {x} \<longleftrightarrow> e = 0"
proof (rule linorder_cases)
assume e: "0 < e"
obtain a where "a \<noteq> x" "dist a x < e"
using perfect_choose_dist [OF e] by auto
then have "a \<noteq> x" "dist x a \<le> e"
by (auto simp: dist_commute)
with e show ?thesis by (auto simp: set_eq_iff)
qed auto
subsection \<open>?\<close>
lemma finite_ball_include:
fixes a :: "'a::metric_space"
assumes "finite S"
shows "\<exists>e>0. S \<subseteq> ball a e"
using assms
proof induction
case (insert x S)
then obtain e0 where "e0>0" and e0:"S \<subseteq> ball a e0" by auto
define e where "e = max e0 (2 * dist a x)"
have "e>0" unfolding e_def using \<open>e0>0\<close> by auto
moreover have "insert x S \<subseteq> ball a e"
using e0 \<open>e>0\<close> unfolding e_def by auto
ultimately show ?case by auto
qed (auto intro: zero_less_one)
lemma finite_set_avoid:
fixes a :: "'a::metric_space"
assumes "finite S"
shows "\<exists>d>0. \<forall>x\<in>S. x \<noteq> a \<longrightarrow> d \<le> dist a x"
using assms
proof induction
case (insert x S)
then obtain d where "d > 0" and d: "\<forall>x\<in>S. x \<noteq> a \<longrightarrow> d \<le> dist a x"
by blast
show ?case
proof (cases "x = a")
case True
with \<open>d > 0 \<close>d show ?thesis by auto
next
case False
let ?d = "min d (dist a x)"
from False \<open>d > 0\<close> have dp: "?d > 0"
by auto
from d have d': "\<forall>x\<in>S. x \<noteq> a \<longrightarrow> ?d \<le> dist a x"
by auto
with dp False show ?thesis
by (metis insert_iff le_less min_less_iff_conj not_less)
qed
qed (auto intro: zero_less_one)
lemma discrete_imp_closed:
fixes S :: "'a::metric_space set"
assumes e: "0 < e"
and d: "\<forall>x \<in> S. \<forall>y \<in> S. dist y x < e \<longrightarrow> y = x"
shows "closed S"
proof -
have False if C: "\<And>e. e>0 \<Longrightarrow> \<exists>x'\<in>S. x' \<noteq> x \<and> dist x' x < e" for x
proof -
from e have e2: "e/2 > 0" by arith
from C[rule_format, OF e2] obtain y where y: "y \<in> S" "y \<noteq> x" "dist y x < e/2"
by blast
from e2 y(2) have mp: "min (e/2) (dist x y) > 0"
by simp
from d y C[OF mp] show ?thesis
by metric
qed
then show ?thesis
by (metis islimpt_approachable closed_limpt [where 'a='a])
qed
subsection \<open>Interior\<close>
lemma mem_interior: "x \<in> interior S \<longleftrightarrow> (\<exists>e>0. ball x e \<subseteq> S)"
using open_contains_ball_eq [where S="interior S"]
by (simp add: open_subset_interior)
lemma mem_interior_cball: "x \<in> interior S \<longleftrightarrow> (\<exists>e>0. cball x e \<subseteq> S)"
by (meson ball_subset_cball interior_subset mem_interior open_contains_cball open_interior
subset_trans)
subsection \<open>Frontier\<close>
lemma frontier_straddle:
fixes a :: "'a::metric_space"
shows "a \<in> frontier S \<longleftrightarrow> (\<forall>e>0. (\<exists>x\<in>S. dist a x < e) \<and> (\<exists>x. x \<notin> S \<and> dist a x < e))"
unfolding frontier_def closure_interior
by (auto simp: mem_interior subset_eq ball_def)
subsection \<open>Limits\<close>
proposition Lim: "(f \<longlongrightarrow> l) net \<longleftrightarrow> trivial_limit net \<or> (\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) net)"
by (auto simp: tendsto_iff trivial_limit_eq)
text \<open>Show that they yield usual definitions in the various cases.\<close>
proposition Lim_within_le: "(f \<longlongrightarrow> l)(at a within S) \<longleftrightarrow>
(\<forall>e>0. \<exists>d>0. \<forall>x\<in>S. 0 < dist x a \<and> dist x a \<le> d \<longrightarrow> dist (f x) l < e)"
by (auto simp: tendsto_iff eventually_at_le)
proposition Lim_within: "(f \<longlongrightarrow> l) (at a within S) \<longleftrightarrow>
(\<forall>e >0. \<exists>d>0. \<forall>x \<in> S. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) l < e)"
by (auto simp: tendsto_iff eventually_at)
corollary Lim_withinI [intro?]:
assumes "\<And>e. e > 0 \<Longrightarrow> \<exists>d>0. \<forall>x \<in> S. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) l \<le> e"
shows "(f \<longlongrightarrow> l) (at a within S)"
apply (simp add: Lim_within, clarify)
apply (rule ex_forward [OF assms [OF half_gt_zero]], auto)
done
proposition Lim_at: "(f \<longlongrightarrow> l) (at a) \<longleftrightarrow>
(\<forall>e >0. \<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) l < e)"
by (auto simp: tendsto_iff eventually_at)
lemma Lim_transform_within_set:
fixes a :: "'a::metric_space" and l :: "'b::metric_space"
shows "\<lbrakk>(f \<longlongrightarrow> l) (at a within S); eventually (\<lambda>x. x \<in> S \<longleftrightarrow> x \<in> T) (at a)\<rbrakk>
\<Longrightarrow> (f \<longlongrightarrow> l) (at a within T)"
apply (clarsimp simp: eventually_at Lim_within)
apply (drule_tac x=e in spec, clarify)
apply (rename_tac k)
apply (rule_tac x="min d k" in exI, simp)
done
text \<open>Another limit point characterization.\<close>
lemma limpt_sequential_inj:
fixes x :: "'a::metric_space"
shows "x islimpt S \<longleftrightarrow>
(\<exists>f. (\<forall>n::nat. f n \<in> S - {x}) \<and> inj f \<and> (f \<longlongrightarrow> x) sequentially)"
(is "?lhs = ?rhs")
proof
assume ?lhs
then have "\<forall>e>0. \<exists>x'\<in>S. x' \<noteq> x \<and> dist x' x < e"
by (force simp: islimpt_approachable)
then obtain y where y: "\<And>e. e>0 \<Longrightarrow> y e \<in> S \<and> y e \<noteq> x \<and> dist (y e) x < e"
by metis
define f where "f \<equiv> rec_nat (y 1) (\<lambda>n fn. y (min (inverse(2 ^ (Suc n))) (dist fn x)))"
have [simp]: "f 0 = y 1"
"f(Suc n) = y (min (inverse(2 ^ (Suc n))) (dist (f n) x))" for n
by (simp_all add: f_def)
have f: "f n \<in> S \<and> (f n \<noteq> x) \<and> dist (f n) x < inverse(2 ^ n)" for n
proof (induction n)
case 0 show ?case
by (simp add: y)
next
case (Suc n) then show ?case
apply (auto simp: y)
by (metis half_gt_zero_iff inverse_positive_iff_positive less_divide_eq_numeral1(1) min_less_iff_conj y zero_less_dist_iff zero_less_numeral zero_less_power)
qed
show ?rhs
proof (rule_tac x=f in exI, intro conjI allI)
show "\<And>n. f n \<in> S - {x}"
using f by blast
have "dist (f n) x < dist (f m) x" if "m < n" for m n
using that
proof (induction n)
case 0 then show ?case by simp
next
case (Suc n)
then consider "m < n" | "m = n" using less_Suc_eq by blast
then show ?case
proof cases
assume "m < n"
have "dist (f(Suc n)) x = dist (y (min (inverse(2 ^ (Suc n))) (dist (f n) x))) x"
by simp
also have "\<dots> < dist (f n) x"
by (metis dist_pos_lt f min.strict_order_iff min_less_iff_conj y)
also have "\<dots> < dist (f m) x"
using Suc.IH \<open>m < n\<close> by blast
finally show ?thesis .
next
assume "m = n" then show ?case
by simp (metis dist_pos_lt f half_gt_zero_iff inverse_positive_iff_positive min_less_iff_conj y zero_less_numeral zero_less_power)
qed
qed
then show "inj f"
by (metis less_irrefl linorder_injI)
show "f \<longlonglongrightarrow> x"
apply (rule tendstoI)
apply (rule_tac c="nat (ceiling(1/e))" in eventually_sequentiallyI)
apply (rule less_trans [OF f [THEN conjunct2, THEN conjunct2]])
apply (simp add: field_simps)
by (meson le_less_trans mult_less_cancel_left not_le of_nat_less_two_power)
qed
next
assume ?rhs
then show ?lhs
by (fastforce simp add: islimpt_approachable lim_sequentially)
qed
lemma Lim_dist_ubound:
assumes "\<not>(trivial_limit net)"
and "(f \<longlongrightarrow> l) net"
and "eventually (\<lambda>x. dist a (f x) \<le> e) net"
shows "dist a l \<le> e"
using assms by (fast intro: tendsto_le tendsto_intros)
subsection \<open>Continuity\<close>
text\<open>Derive the epsilon-delta forms, which we often use as "definitions"\<close>
proposition continuous_within_eps_delta:
"continuous (at x within s) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. \<forall>x'\<in> s. dist x' x < d --> dist (f x') (f x) < e)"
unfolding continuous_within and Lim_within by fastforce
corollary continuous_at_eps_delta:
"continuous (at x) f \<longleftrightarrow> (\<forall>e > 0. \<exists>d > 0. \<forall>x'. dist x' x < d \<longrightarrow> dist (f x') (f x) < e)"
using continuous_within_eps_delta [of x UNIV f] by simp
lemma continuous_at_right_real_increasing:
fixes f :: "real \<Rightarrow> real"
assumes nondecF: "\<And>x y. x \<le> y \<Longrightarrow> f x \<le> f y"
shows "continuous (at_right a) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. f (a + d) - f a < e)"
apply (simp add: greaterThan_def dist_real_def continuous_within Lim_within_le)
apply (intro all_cong ex_cong, safe)
apply (erule_tac x="a + d" in allE, simp)
apply (simp add: nondecF field_simps)
apply (drule nondecF, simp)
done
lemma continuous_at_left_real_increasing:
assumes nondecF: "\<And> x y. x \<le> y \<Longrightarrow> f x \<le> ((f y) :: real)"
shows "(continuous (at_left (a :: real)) f) = (\<forall>e > 0. \<exists>delta > 0. f a - f (a - delta) < e)"
apply (simp add: lessThan_def dist_real_def continuous_within Lim_within_le)
apply (intro all_cong ex_cong, safe)
apply (erule_tac x="a - d" in allE, simp)
apply (simp add: nondecF field_simps)
apply (cut_tac x="a - d" and y=x in nondecF, simp_all)
done
text\<open>Versions in terms of open balls.\<close>
lemma continuous_within_ball:
"continuous (at x within s) f \<longleftrightarrow>
(\<forall>e > 0. \<exists>d > 0. f ` (ball x d \<inter> s) \<subseteq> ball (f x) e)"
(is "?lhs = ?rhs")
proof
assume ?lhs
{
fix e :: real
assume "e > 0"
then obtain d where d: "d>0" "\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e"
using \<open>?lhs\<close>[unfolded continuous_within Lim_within] by auto
{
fix y
assume "y \<in> f ` (ball x d \<inter> s)"
then have "y \<in> ball (f x) e"
using d(2)
using \<open>e > 0\<close>
by (auto simp: dist_commute)
}
then have "\<exists>d>0. f ` (ball x d \<inter> s) \<subseteq> ball (f x) e"
using \<open>d > 0\<close>
unfolding subset_eq ball_def by (auto simp: dist_commute)
}
then show ?rhs by auto
next
assume ?rhs
then show ?lhs
unfolding continuous_within Lim_within ball_def subset_eq
apply (auto simp: dist_commute)
apply (erule_tac x=e in allE, auto)
done
qed
lemma continuous_at_ball:
"continuous (at x) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. f ` (ball x d) \<subseteq> ball (f x) e)" (is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball
apply auto
apply (erule_tac x=e in allE, auto)
apply (rule_tac x=d in exI, auto)
apply (erule_tac x=xa in allE)
apply (auto simp: dist_commute)
done
next
assume ?rhs
then show ?lhs
unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball
apply auto
apply (erule_tac x=e in allE, auto)
apply (rule_tac x=d in exI, auto)
apply (erule_tac x="f xa" in allE)
apply (auto simp: dist_commute)
done
qed
text\<open>Define setwise continuity in terms of limits within the set.\<close>
lemma continuous_on_iff:
"continuous_on s f \<longleftrightarrow>
(\<forall>x\<in>s. \<forall>e>0. \<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e)"
unfolding continuous_on_def Lim_within
by (metis dist_pos_lt dist_self)
lemma continuous_within_E:
assumes "continuous (at x within s) f" "e>0"
obtains d where "d>0" "\<And>x'. \<lbrakk>x'\<in> s; dist x' x \<le> d\<rbrakk> \<Longrightarrow> dist (f x') (f x) < e"
using assms apply (simp add: continuous_within_eps_delta)
apply (drule spec [of _ e], clarify)
apply (rule_tac d="d/2" in that, auto)
done
lemma continuous_onI [intro?]:
assumes "\<And>x e. \<lbrakk>e > 0; x \<in> s\<rbrakk> \<Longrightarrow> \<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) \<le> e"
shows "continuous_on s f"
apply (simp add: continuous_on_iff, clarify)
apply (rule ex_forward [OF assms [OF half_gt_zero]], auto)
done
text\<open>Some simple consequential lemmas.\<close>
lemma continuous_onE:
assumes "continuous_on s f" "x\<in>s" "e>0"
obtains d where "d>0" "\<And>x'. \<lbrakk>x' \<in> s; dist x' x \<le> d\<rbrakk> \<Longrightarrow> dist (f x') (f x) < e"
using assms
apply (simp add: continuous_on_iff)
apply (elim ballE allE)
apply (auto intro: that [where d="d/2" for d])
done
text\<open>The usual transformation theorems.\<close>
lemma continuous_transform_within:
fixes f g :: "'a::metric_space \<Rightarrow> 'b::topological_space"
assumes "continuous (at x within s) f"
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 "continuous (at x within s) g"
using assms
unfolding continuous_within
by (force intro: Lim_transform_within)
subsection \<open>Closure and Limit Characterization\<close>
lemma closure_approachable:
fixes S :: "'a::metric_space set"
shows "x \<in> closure S \<longleftrightarrow> (\<forall>e>0. \<exists>y\<in>S. dist y x < e)"
apply (auto simp: closure_def islimpt_approachable)
apply (metis dist_self)
done
lemma closure_approachable_le:
fixes S :: "'a::metric_space set"
shows "x \<in> closure S \<longleftrightarrow> (\<forall>e>0. \<exists>y\<in>S. dist y x \<le> e)"
unfolding closure_approachable
using dense by force
lemma closure_approachableD:
assumes "x \<in> closure S" "e>0"
shows "\<exists>y\<in>S. dist x y < e"
using assms unfolding closure_approachable by (auto simp: dist_commute)
lemma closed_approachable:
fixes S :: "'a::metric_space set"
shows "closed S \<Longrightarrow> (\<forall>e>0. \<exists>y\<in>S. dist y x < e) \<longleftrightarrow> x \<in> S"
by (metis closure_closed closure_approachable)
lemma closure_contains_Inf:
fixes S :: "real set"
assumes "S \<noteq> {}" "bdd_below S"
shows "Inf S \<in> closure S"
proof -
have *: "\<forall>x\<in>S. Inf S \<le> x"
using cInf_lower[of _ S] assms by metis
{
fix e :: real
assume "e > 0"
then have "Inf S < Inf S + e" by simp
with assms obtain x where "x \<in> S" "x < Inf S + e"
by (subst (asm) cInf_less_iff) auto
with * have "\<exists>x\<in>S. dist x (Inf S) < e"
by (intro bexI[of _ x]) (auto simp: dist_real_def)
}
then show ?thesis unfolding closure_approachable by auto
qed
lemma closure_contains_Sup:
fixes S :: "real set"
assumes "S \<noteq> {}" "bdd_above S"
shows "Sup S \<in> closure S"
proof -
have *: "\<forall>x\<in>S. x \<le> Sup S"
using cSup_upper[of _ S] assms by metis
{
fix e :: real
assume "e > 0"
then have "Sup S - e < Sup S" by simp
with assms obtain x where "x \<in> S" "Sup S - e < x"
by (subst (asm) less_cSup_iff) auto
with * have "\<exists>x\<in>S. dist x (Sup S) < e"
by (intro bexI[of _ x]) (auto simp: dist_real_def)
}
then show ?thesis unfolding closure_approachable by auto
qed
lemma not_trivial_limit_within_ball:
"\<not> trivial_limit (at x within S) \<longleftrightarrow> (\<forall>e>0. S \<inter> ball x e - {x} \<noteq> {})"
(is "?lhs \<longleftrightarrow> ?rhs")
proof
show ?rhs if ?lhs
proof -
{
fix e :: real
assume "e > 0"
then obtain y where "y \<in> S - {x}" and "dist y x < e"
using \<open>?lhs\<close> not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"]
by auto
then have "y \<in> S \<inter> ball x e - {x}"
unfolding ball_def by (simp add: dist_commute)
then have "S \<inter> ball x e - {x} \<noteq> {}" by blast
}
then show ?thesis by auto
qed
show ?lhs if ?rhs
proof -
{
fix e :: real
assume "e > 0"
then obtain y where "y \<in> S \<inter> ball x e - {x}"
using \<open>?rhs\<close> by blast
then have "y \<in> S - {x}" and "dist y x < e"
unfolding ball_def by (simp_all add: dist_commute)
then have "\<exists>y \<in> S - {x}. dist y x < e"
by auto
}
then show ?thesis
using not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"]
by auto
qed
qed
subsection \<open>Boundedness\<close>
(* FIXME: This has to be unified with BSEQ!! *)
definition\<^marker>\<open>tag important\<close> (in metric_space) bounded :: "'a set \<Rightarrow> bool"
where "bounded S \<longleftrightarrow> (\<exists>x e. \<forall>y\<in>S. dist x y \<le> e)"
lemma bounded_subset_cball: "bounded S \<longleftrightarrow> (\<exists>e x. S \<subseteq> cball x e \<and> 0 \<le> e)"
unfolding bounded_def subset_eq by auto (meson order_trans zero_le_dist)
lemma bounded_any_center: "bounded S \<longleftrightarrow> (\<exists>e. \<forall>y\<in>S. dist a y \<le> e)"
unfolding bounded_def
by auto (metis add.commute add_le_cancel_right dist_commute dist_triangle_le)
lemma bounded_iff: "bounded S \<longleftrightarrow> (\<exists>a. \<forall>x\<in>S. norm x \<le> a)"
unfolding bounded_any_center [where a=0]
by (simp add: dist_norm)
lemma bdd_above_norm: "bdd_above (norm ` X) \<longleftrightarrow> bounded X"
by (simp add: bounded_iff bdd_above_def)
lemma bounded_norm_comp: "bounded ((\<lambda>x. norm (f x)) ` S) = bounded (f ` S)"
by (simp add: bounded_iff)
lemma boundedI:
assumes "\<And>x. x \<in> S \<Longrightarrow> norm x \<le> B"
shows "bounded S"
using assms bounded_iff by blast
lemma bounded_empty [simp]: "bounded {}"
by (simp add: bounded_def)
lemma bounded_subset: "bounded T \<Longrightarrow> S \<subseteq> T \<Longrightarrow> bounded S"
by (metis bounded_def subset_eq)
lemma bounded_interior[intro]: "bounded S \<Longrightarrow> bounded(interior S)"
by (metis bounded_subset interior_subset)
lemma bounded_closure[intro]:
assumes "bounded S"
shows "bounded (closure S)"
proof -
from assms obtain x and a where a: "\<forall>y\<in>S. dist x y \<le> a"
unfolding bounded_def by auto
{
fix y
assume "y \<in> closure S"
then obtain f where f: "\<forall>n. f n \<in> S" "(f \<longlongrightarrow> y) sequentially"
unfolding closure_sequential by auto
have "\<forall>n. f n \<in> S \<longrightarrow> dist x (f n) \<le> a" using a by simp
then have "eventually (\<lambda>n. dist x (f n) \<le> a) sequentially"
by (simp add: f(1))
have "dist x y \<le> a"
apply (rule Lim_dist_ubound [of sequentially f])
apply (rule trivial_limit_sequentially)
apply (rule f(2))
apply fact
done
}
then show ?thesis
unfolding bounded_def by auto
qed
lemma bounded_closure_image: "bounded (f ` closure S) \<Longrightarrow> bounded (f ` S)"
by (simp add: bounded_subset closure_subset image_mono)
lemma bounded_cball[simp,intro]: "bounded (cball x e)"
apply (simp add: bounded_def)
apply (rule_tac x=x in exI)
apply (rule_tac x=e in exI, auto)
done
lemma bounded_ball[simp,intro]: "bounded (ball x e)"
by (metis ball_subset_cball bounded_cball bounded_subset)
lemma bounded_Un[simp]: "bounded (S \<union> T) \<longleftrightarrow> bounded S \<and> bounded T"
by (auto simp: bounded_def) (metis Un_iff bounded_any_center le_max_iff_disj)
lemma bounded_Union[intro]: "finite F \<Longrightarrow> \<forall>S\<in>F. bounded S \<Longrightarrow> bounded (\<Union>F)"
by (induct rule: finite_induct[of F]) auto
lemma bounded_UN [intro]: "finite A \<Longrightarrow> \<forall>x\<in>A. bounded (B x) \<Longrightarrow> bounded (\<Union>x\<in>A. B x)"
by (induct set: finite) auto
lemma bounded_insert [simp]: "bounded (insert x S) \<longleftrightarrow> bounded S"
proof -
have "\<forall>y\<in>{x}. dist x y \<le> 0"
by simp
then have "bounded {x}"
unfolding bounded_def by fast
then show ?thesis
by (metis insert_is_Un bounded_Un)
qed
lemma bounded_subset_ballI: "S \<subseteq> ball x r \<Longrightarrow> bounded S"
by (meson bounded_ball bounded_subset)
lemma bounded_subset_ballD:
assumes "bounded S" shows "\<exists>r. 0 < r \<and> S \<subseteq> ball x r"
proof -
obtain e::real and y where "S \<subseteq> cball y e" "0 \<le> e"
using assms by (auto simp: bounded_subset_cball)
then show ?thesis
by (intro exI[where x="dist x y + e + 1"]) metric
qed
lemma finite_imp_bounded [intro]: "finite S \<Longrightarrow> bounded S"
by (induct set: finite) simp_all
lemma bounded_Int[intro]: "bounded S \<or> bounded T \<Longrightarrow> bounded (S \<inter> T)"
by (metis Int_lower1 Int_lower2 bounded_subset)
lemma bounded_diff[intro]: "bounded S \<Longrightarrow> bounded (S - T)"
by (metis Diff_subset bounded_subset)
lemma bounded_dist_comp:
assumes "bounded (f ` S)" "bounded (g ` S)"
shows "bounded ((\<lambda>x. dist (f x) (g x)) ` S)"
proof -
from assms obtain M1 M2 where *: "dist (f x) undefined \<le> M1" "dist undefined (g x) \<le> M2" if "x \<in> S" for x
by (auto simp: bounded_any_center[of _ undefined] dist_commute)
have "dist (f x) (g x) \<le> M1 + M2" if "x \<in> S" for x
using *[OF that]
by metric
then show ?thesis
by (auto intro!: boundedI)
qed
lemma bounded_Times:
assumes "bounded s" "bounded t"
shows "bounded (s \<times> t)"
proof -
obtain x y a b where "\<forall>z\<in>s. dist x z \<le> a" "\<forall>z\<in>t. dist y z \<le> b"
using assms [unfolded bounded_def] by auto
then have "\<forall>z\<in>s \<times> t. dist (x, y) z \<le> sqrt (a\<^sup>2 + b\<^sup>2)"
by (auto simp: dist_Pair_Pair real_sqrt_le_mono add_mono power_mono)
then show ?thesis unfolding bounded_any_center [where a="(x, y)"] by auto
qed
subsection \<open>Compactness\<close>
lemma compact_imp_bounded:
assumes "compact U"
shows "bounded U"
proof -
have "compact U" "\<forall>x\<in>U. open (ball x 1)" "U \<subseteq> (\<Union>x\<in>U. ball x 1)"
using assms by auto
then obtain D where D: "D \<subseteq> U" "finite D" "U \<subseteq> (\<Union>x\<in>D. ball x 1)"
by (metis compactE_image)
from \<open>finite D\<close> have "bounded (\<Union>x\<in>D. ball x 1)"
by (simp add: bounded_UN)
then show "bounded U" using \<open>U \<subseteq> (\<Union>x\<in>D. ball x 1)\<close>
by (rule bounded_subset)
qed
lemma closure_Int_ball_not_empty:
assumes "S \<subseteq> closure T" "x \<in> S" "r > 0"
shows "T \<inter> ball x r \<noteq> {}"
using assms centre_in_ball closure_iff_nhds_not_empty by blast
lemma compact_sup_maxdistance:
fixes s :: "'a::metric_space set"
assumes "compact s"
and "s \<noteq> {}"
shows "\<exists>x\<in>s. \<exists>y\<in>s. \<forall>u\<in>s. \<forall>v\<in>s. dist u v \<le> dist x y"
proof -
have "compact (s \<times> s)"
using \<open>compact s\<close> by (intro compact_Times)
moreover have "s \<times> s \<noteq> {}"
using \<open>s \<noteq> {}\<close> by auto
moreover have "continuous_on (s \<times> s) (\<lambda>x. dist (fst x) (snd x))"
by (intro continuous_at_imp_continuous_on ballI continuous_intros)
ultimately show ?thesis
using continuous_attains_sup[of "s \<times> s" "\<lambda>x. dist (fst x) (snd x)"] by auto
qed
subsubsection\<open>Totally bounded\<close>
lemma cauchy_def: "Cauchy s \<longleftrightarrow> (\<forall>e>0. \<exists>N. \<forall>m n. m \<ge> N \<and> n \<ge> N \<longrightarrow> dist (s m) (s n) < e)"
unfolding Cauchy_def by metis
proposition seq_compact_imp_totally_bounded:
assumes "seq_compact s"
shows "\<forall>e>0. \<exists>k. finite k \<and> k \<subseteq> s \<and> s \<subseteq> (\<Union>x\<in>k. ball x e)"
proof -
{ fix e::real assume "e > 0" assume *: "\<And>k. finite k \<Longrightarrow> k \<subseteq> s \<Longrightarrow> \<not> s \<subseteq> (\<Union>x\<in>k. ball x e)"
let ?Q = "\<lambda>x n r. r \<in> s \<and> (\<forall>m < (n::nat). \<not> (dist (x m) r < e))"
have "\<exists>x. \<forall>n::nat. ?Q x n (x n)"
proof (rule dependent_wellorder_choice)
fix n x assume "\<And>y. y < n \<Longrightarrow> ?Q x y (x y)"
then have "\<not> s \<subseteq> (\<Union>x\<in>x ` {0..<n}. ball x e)"
using *[of "x ` {0 ..< n}"] by (auto simp: subset_eq)
then obtain z where z:"z\<in>s" "z \<notin> (\<Union>x\<in>x ` {0..<n}. ball x e)"
unfolding subset_eq by auto
show "\<exists>r. ?Q x n r"
using z by auto
qed simp
then obtain x where "\<forall>n::nat. x n \<in> s" and x:"\<And>n m. m < n \<Longrightarrow> \<not> (dist (x m) (x n) < e)"
by blast
then obtain l r where "l \<in> s" and r:"strict_mono r" and "((x \<circ> r) \<longlongrightarrow> l) sequentially"
using assms by (metis seq_compact_def)
from this(3) have "Cauchy (x \<circ> r)"
using LIMSEQ_imp_Cauchy by auto
then obtain N::nat where "\<And>m n. N \<le> m \<Longrightarrow> N \<le> n \<Longrightarrow> dist ((x \<circ> r) m) ((x \<circ> r) n) < e"
unfolding cauchy_def using \<open>e > 0\<close> by blast
then have False
using x[of "r N" "r (N+1)"] r by (auto simp: strict_mono_def) }
then show ?thesis
by metis
qed
subsubsection\<open>Heine-Borel theorem\<close>
proposition seq_compact_imp_Heine_Borel:
fixes s :: "'a :: metric_space set"
assumes "seq_compact s"
shows "compact s"
proof -
from seq_compact_imp_totally_bounded[OF \<open>seq_compact s\<close>]
obtain f where f: "\<forall>e>0. finite (f e) \<and> f e \<subseteq> s \<and> s \<subseteq> (\<Union>x\<in>f e. ball x e)"
unfolding choice_iff' ..
define K where "K = (\<lambda>(x, r). ball x r) ` ((\<Union>e \<in> \<rat> \<inter> {0 <..}. f e) \<times> \<rat>)"
have "countably_compact s"
using \<open>seq_compact s\<close> by (rule seq_compact_imp_countably_compact)
then show "compact s"
proof (rule countably_compact_imp_compact)
show "countable K"
unfolding K_def using f
by (auto intro: countable_finite countable_subset countable_rat
intro!: countable_image countable_SIGMA countable_UN)
show "\<forall>b\<in>K. open b" by (auto simp: K_def)
next
fix T x
assume T: "open T" "x \<in> T" and x: "x \<in> s"
from openE[OF T] obtain e where "0 < e" "ball x e \<subseteq> T"
by auto
then have "0 < e / 2" "ball x (e / 2) \<subseteq> T"
by auto
from Rats_dense_in_real[OF \<open>0 < e / 2\<close>] obtain r where "r \<in> \<rat>" "0 < r" "r < e / 2"
by auto
from f[rule_format, of r] \<open>0 < r\<close> \<open>x \<in> s\<close> obtain k where "k \<in> f r" "x \<in> ball k r"
by auto
from \<open>r \<in> \<rat>\<close> \<open>0 < r\<close> \<open>k \<in> f r\<close> have "ball k r \<in> K"
by (auto simp: K_def)
then show "\<exists>b\<in>K. x \<in> b \<and> b \<inter> s \<subseteq> T"
proof (rule bexI[rotated], safe)
fix y
assume "y \<in> ball k r"
with \<open>r < e / 2\<close> \<open>x \<in> ball k r\<close> have "dist x y < e"
by (intro dist_triangle_half_r [of k _ e]) (auto simp: dist_commute)
with \<open>ball x e \<subseteq> T\<close> show "y \<in> T"
by auto
next
show "x \<in> ball k r" by fact
qed
qed
qed
proposition compact_eq_seq_compact_metric:
"compact (s :: 'a::metric_space set) \<longleftrightarrow> seq_compact s"
using compact_imp_seq_compact seq_compact_imp_Heine_Borel by blast
proposition compact_def: \<comment> \<open>this is the definition of compactness in HOL Light\<close>
"compact (S :: 'a::metric_space set) \<longleftrightarrow>
(\<forall>f. (\<forall>n. f n \<in> S) \<longrightarrow> (\<exists>l\<in>S. \<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> (f \<circ> r) \<longlonglongrightarrow> l))"
unfolding compact_eq_seq_compact_metric seq_compact_def by auto
subsubsection \<open>Complete the chain of compactness variants\<close>
proposition compact_eq_Bolzano_Weierstrass:
fixes s :: "'a::metric_space set"
shows "compact s \<longleftrightarrow> (\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t))"
(is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
using Heine_Borel_imp_Bolzano_Weierstrass[of s] by auto
next
assume ?rhs
then show ?lhs
unfolding compact_eq_seq_compact_metric by (rule Bolzano_Weierstrass_imp_seq_compact)
qed
proposition Bolzano_Weierstrass_imp_bounded:
"\<forall>t. infinite t \<and> t \<subseteq> s \<longrightarrow> (\<exists>x \<in> s. x islimpt t) \<Longrightarrow> bounded s"
using compact_imp_bounded unfolding compact_eq_Bolzano_Weierstrass .
subsection \<open>Banach fixed point theorem\<close>
theorem banach_fix:\<comment> \<open>TODO: rename to \<open>Banach_fix\<close>\<close>
assumes s: "complete s" "s \<noteq> {}"
and c: "0 \<le> c" "c < 1"
and f: "f ` s \<subseteq> s"
and lipschitz: "\<forall>x\<in>s. \<forall>y\<in>s. dist (f x) (f y) \<le> c * dist x y"
shows "\<exists>!x\<in>s. f x = x"
proof -
from c have "1 - c > 0" by simp
from s(2) obtain z0 where z0: "z0 \<in> s" by blast
define z where "z n = (f ^^ n) z0" for n
with f z0 have z_in_s: "z n \<in> s" for n :: nat
by (induct n) auto
define d where "d = dist (z 0) (z 1)"
have fzn: "f (z n) = z (Suc n)" for n
by (simp add: z_def)
have cf_z: "dist (z n) (z (Suc n)) \<le> (c ^ n) * d" for n :: nat
proof (induct n)
case 0
then show ?case
by (simp add: d_def)
next
case (Suc m)
with \<open>0 \<le> c\<close> have "c * dist (z m) (z (Suc m)) \<le> c ^ Suc m * d"
using mult_left_mono[of "dist (z m) (z (Suc m))" "c ^ m * d" c] by simp
then show ?case
using lipschitz[THEN bspec[where x="z m"], OF z_in_s, THEN bspec[where x="z (Suc m)"], OF z_in_s]
by (simp add: fzn mult_le_cancel_left)
qed
have cf_z2: "(1 - c) * dist (z m) (z (m + n)) \<le> (c ^ m) * d * (1 - c ^ n)" for n m :: nat
proof (induct n)
case 0
show ?case by simp
next
case (Suc k)
from c have "(1 - c) * dist (z m) (z (m + Suc k)) \<le>
(1 - c) * (dist (z m) (z (m + k)) + dist (z (m + k)) (z (Suc (m + k))))"
by (simp add: dist_triangle)
also from c cf_z[of "m + k"] have "\<dots> \<le> (1 - c) * (dist (z m) (z (m + k)) + c ^ (m + k) * d)"
by simp
also from Suc have "\<dots> \<le> c ^ m * d * (1 - c ^ k) + (1 - c) * c ^ (m + k) * d"
by (simp add: field_simps)
also have "\<dots> = (c ^ m) * (d * (1 - c ^ k) + (1 - c) * c ^ k * d)"
by (simp add: power_add field_simps)
also from c have "\<dots> \<le> (c ^ m) * d * (1 - c ^ Suc k)"
by (simp add: field_simps)
finally show ?case by simp
qed
have "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (z m) (z n) < e" if "e > 0" for e
proof (cases "d = 0")
case True
from \<open>1 - c > 0\<close> have "(1 - c) * x \<le> 0 \<longleftrightarrow> x \<le> 0" for x
by (metis mult_zero_left mult.commute real_mult_le_cancel_iff1)
with c cf_z2[of 0] True have "z n = z0" for n
by (simp add: z_def)
with \<open>e > 0\<close> show ?thesis by simp
next
case False
with zero_le_dist[of "z 0" "z 1"] have "d > 0"
by (metis d_def less_le)
with \<open>1 - c > 0\<close> \<open>e > 0\<close> have "0 < e * (1 - c) / d"
by simp
with c obtain N where N: "c ^ N < e * (1 - c) / d"
using real_arch_pow_inv[of "e * (1 - c) / d" c] by auto
have *: "dist (z m) (z n) < e" if "m > n" and as: "m \<ge> N" "n \<ge> N" for m n :: nat
proof -
from c \<open>n \<ge> N\<close> have *: "c ^ n \<le> c ^ N"
using power_decreasing[OF \<open>n\<ge>N\<close>, of c] by simp
from c \<open>m > n\<close> have "1 - c ^ (m - n) > 0"
using power_strict_mono[of c 1 "m - n"] by simp
with \<open>d > 0\<close> \<open>0 < 1 - c\<close> have **: "d * (1 - c ^ (m - n)) / (1 - c) > 0"
by simp
from cf_z2[of n "m - n"] \<open>m > n\<close>
have "dist (z m) (z n) \<le> c ^ n * d * (1 - c ^ (m - n)) / (1 - c)"
by (simp add: pos_le_divide_eq[OF \<open>1 - c > 0\<close>] mult.commute dist_commute)
also have "\<dots> \<le> c ^ N * d * (1 - c ^ (m - n)) / (1 - c)"
using mult_right_mono[OF * order_less_imp_le[OF **]]
by (simp add: mult.assoc)
also have "\<dots> < (e * (1 - c) / d) * d * (1 - c ^ (m - n)) / (1 - c)"
using mult_strict_right_mono[OF N **] by (auto simp: mult.assoc)
also from c \<open>d > 0\<close> \<open>1 - c > 0\<close> have "\<dots> = e * (1 - c ^ (m - n))"
by simp
also from c \<open>1 - c ^ (m - n) > 0\<close> \<open>e > 0\<close> have "\<dots> \<le> e"
using mult_right_le_one_le[of e "1 - c ^ (m - n)"] by auto
finally show ?thesis by simp
qed
have "dist (z n) (z m) < e" if "N \<le> m" "N \<le> n" for m n :: nat
proof (cases "n = m")
case True
with \<open>e > 0\<close> show ?thesis by simp
next
case False
with *[of n m] *[of m n] and that show ?thesis
by (auto simp: dist_commute nat_neq_iff)
qed
then show ?thesis by auto
qed
then have "Cauchy z"
by (simp add: cauchy_def)
then obtain x where "x\<in>s" and x:"(z \<longlongrightarrow> x) sequentially"
using s(1)[unfolded compact_def complete_def, THEN spec[where x=z]] and z_in_s by auto
define e where "e = dist (f x) x"
have "e = 0"
proof (rule ccontr)
assume "e \<noteq> 0"
then have "e > 0"
unfolding e_def using zero_le_dist[of "f x" x]
by (metis dist_eq_0_iff dist_nz e_def)
then obtain N where N:"\<forall>n\<ge>N. dist (z n) x < e / 2"
using x[unfolded lim_sequentially, THEN spec[where x="e/2"]] by auto
then have N':"dist (z N) x < e / 2" by auto
have *: "c * dist (z N) x \<le> dist (z N) x"
unfolding mult_le_cancel_right2
using zero_le_dist[of "z N" x] and c
by (metis dist_eq_0_iff dist_nz order_less_asym less_le)
have "dist (f (z N)) (f x) \<le> c * dist (z N) x"
using lipschitz[THEN bspec[where x="z N"], THEN bspec[where x=x]]
using z_in_s[of N] \<open>x\<in>s\<close>
using c
by auto
also have "\<dots> < e / 2"
using N' and c using * by auto
finally show False
unfolding fzn
using N[THEN spec[where x="Suc N"]] and dist_triangle_half_r[of "z (Suc N)" "f x" e x]
unfolding e_def
by auto
qed
then have "f x = x" by (auto simp: e_def)
moreover have "y = x" if "f y = y" "y \<in> s" for y
proof -
from \<open>x \<in> s\<close> \<open>f x = x\<close> that have "dist x y \<le> c * dist x y"
using lipschitz[THEN bspec[where x=x], THEN bspec[where x=y]] by simp
with c and zero_le_dist[of x y] have "dist x y = 0"
by (simp add: mult_le_cancel_right1)
then show ?thesis by simp
qed
ultimately show ?thesis
using \<open>x\<in>s\<close> by blast
qed
subsection \<open>Edelstein fixed point theorem\<close>
theorem edelstein_fix:\<comment> \<open>TODO: rename to \<open>Edelstein_fix\<close>\<close>
fixes s :: "'a::metric_space set"
assumes s: "compact s" "s \<noteq> {}"
and gs: "(g ` s) \<subseteq> s"
and dist: "\<forall>x\<in>s. \<forall>y\<in>s. x \<noteq> y \<longrightarrow> dist (g x) (g y) < dist x y"
shows "\<exists>!x\<in>s. g x = x"
proof -
let ?D = "(\<lambda>x. (x, x)) ` s"
have D: "compact ?D" "?D \<noteq> {}"
by (rule compact_continuous_image)
(auto intro!: s continuous_Pair continuous_ident simp: continuous_on_eq_continuous_within)
have "\<And>x y e. x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> 0 < e \<Longrightarrow> dist y x < e \<Longrightarrow> dist (g y) (g x) < e"
using dist by fastforce
then have "continuous_on s g"
by (auto simp: continuous_on_iff)
then have cont: "continuous_on ?D (\<lambda>x. dist ((g \<circ> fst) x) (snd x))"
unfolding continuous_on_eq_continuous_within
by (intro continuous_dist ballI continuous_within_compose)
(auto intro!: continuous_fst continuous_snd continuous_ident simp: image_image)
obtain a where "a \<in> s" and le: "\<And>x. x \<in> s \<Longrightarrow> dist (g a) a \<le> dist (g x) x"
using continuous_attains_inf[OF D cont] by auto
have "g a = a"
proof (rule ccontr)
assume "g a \<noteq> a"
with \<open>a \<in> s\<close> gs have "dist (g (g a)) (g a) < dist (g a) a"
by (intro dist[rule_format]) auto
moreover have "dist (g a) a \<le> dist (g (g a)) (g a)"
using \<open>a \<in> s\<close> gs by (intro le) auto
ultimately show False by auto
qed
moreover have "\<And>x. x \<in> s \<Longrightarrow> g x = x \<Longrightarrow> x = a"
using dist[THEN bspec[where x=a]] \<open>g a = a\<close> and \<open>a\<in>s\<close> by auto
ultimately show "\<exists>!x\<in>s. g x = x"
using \<open>a \<in> s\<close> by blast
qed
subsection \<open>The diameter of a set\<close>
definition\<^marker>\<open>tag important\<close> diameter :: "'a::metric_space set \<Rightarrow> real" where
"diameter S = (if S = {} then 0 else SUP (x,y)\<in>S\<times>S. dist x y)"
lemma diameter_empty [simp]: "diameter{} = 0"
by (auto simp: diameter_def)
lemma diameter_singleton [simp]: "diameter{x} = 0"
by (auto simp: diameter_def)
lemma diameter_le:
assumes "S \<noteq> {} \<or> 0 \<le> d"
and no: "\<And>x y. \<lbrakk>x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> norm(x - y) \<le> d"
shows "diameter S \<le> d"
using assms
by (auto simp: dist_norm diameter_def intro: cSUP_least)
lemma diameter_bounded_bound:
fixes s :: "'a :: metric_space set"
assumes s: "bounded s" "x \<in> s" "y \<in> s"
shows "dist x y \<le> diameter s"
proof -
from s obtain z d where z: "\<And>x. x \<in> s \<Longrightarrow> dist z x \<le> d"
unfolding bounded_def by auto
have "bdd_above (case_prod dist ` (s\<times>s))"
proof (intro bdd_aboveI, safe)
fix a b
assume "a \<in> s" "b \<in> s"
with z[of a] z[of b] dist_triangle[of a b z]
show "dist a b \<le> 2 * d"
by (simp add: dist_commute)
qed
moreover have "(x,y) \<in> s\<times>s" using s by auto
ultimately have "dist x y \<le> (SUP (x,y)\<in>s\<times>s. dist x y)"
by (rule cSUP_upper2) simp
with \<open>x \<in> s\<close> show ?thesis
by (auto simp: diameter_def)
qed
lemma diameter_lower_bounded:
fixes s :: "'a :: metric_space set"
assumes s: "bounded s"
and d: "0 < d" "d < diameter s"
shows "\<exists>x\<in>s. \<exists>y\<in>s. d < dist x y"
proof (rule ccontr)
assume contr: "\<not> ?thesis"
moreover have "s \<noteq> {}"
using d by (auto simp: diameter_def)
ultimately have "diameter s \<le> d"
by (auto simp: not_less diameter_def intro!: cSUP_least)
with \<open>d < diameter s\<close> show False by auto
qed
lemma diameter_bounded:
assumes "bounded s"
shows "\<forall>x\<in>s. \<forall>y\<in>s. dist x y \<le> diameter s"
and "\<forall>d>0. d < diameter s \<longrightarrow> (\<exists>x\<in>s. \<exists>y\<in>s. dist x y > d)"
using diameter_bounded_bound[of s] diameter_lower_bounded[of s] assms
by auto
lemma bounded_two_points:
"bounded S \<longleftrightarrow> (\<exists>e. \<forall>x\<in>S. \<forall>y\<in>S. dist x y \<le> e)"
apply (rule iffI)
subgoal using diameter_bounded(1) by auto
subgoal using bounded_any_center[of S] by meson
done
lemma diameter_compact_attained:
assumes "compact s"
and "s \<noteq> {}"
shows "\<exists>x\<in>s. \<exists>y\<in>s. dist x y = diameter s"
proof -
have b: "bounded s" using assms(1)
by (rule compact_imp_bounded)
then obtain x y where xys: "x\<in>s" "y\<in>s"
and xy: "\<forall>u\<in>s. \<forall>v\<in>s. dist u v \<le> dist x y"
using compact_sup_maxdistance[OF assms] by auto
then have "diameter s \<le> dist x y"
unfolding diameter_def
apply clarsimp
apply (rule cSUP_least, fast+)
done
then show ?thesis
by (metis b diameter_bounded_bound order_antisym xys)
qed
lemma diameter_ge_0:
assumes "bounded S" shows "0 \<le> diameter S"
by (metis all_not_in_conv assms diameter_bounded_bound diameter_empty dist_self order_refl)
lemma diameter_subset:
assumes "S \<subseteq> T" "bounded T"
shows "diameter S \<le> diameter T"
proof (cases "S = {} \<or> T = {}")
case True
with assms show ?thesis
by (force simp: diameter_ge_0)
next
case False
then have "bdd_above ((\<lambda>x. case x of (x, xa) \<Rightarrow> dist x xa) ` (T \<times> T))"
using \<open>bounded T\<close> diameter_bounded_bound by (force simp: bdd_above_def)
with False \<open>S \<subseteq> T\<close> show ?thesis
apply (simp add: diameter_def)
apply (rule cSUP_subset_mono, auto)
done
qed
lemma diameter_closure:
assumes "bounded S"
shows "diameter(closure S) = diameter S"
proof (rule order_antisym)
have "False" if "diameter S < diameter (closure S)"
proof -
define d where "d = diameter(closure S) - diameter(S)"
have "d > 0"
using that by (simp add: d_def)
then have "diameter(closure(S)) - d / 2 < diameter(closure(S))"
by simp
have dd: "diameter (closure S) - d / 2 = (diameter(closure(S)) + diameter(S)) / 2"
by (simp add: d_def field_split_simps)
have bocl: "bounded (closure S)"
using assms by blast
moreover have "0 \<le> diameter S"
using assms diameter_ge_0 by blast
ultimately obtain x y where "x \<in> closure S" "y \<in> closure S" and xy: "diameter(closure(S)) - d / 2 < dist x y"
using diameter_bounded(2) [OF bocl, rule_format, of "diameter(closure(S)) - d / 2"] \<open>d > 0\<close> d_def by auto
then obtain x' y' where x'y': "x' \<in> S" "dist x' x < d/4" "y' \<in> S" "dist y' y < d/4"
using closure_approachable
by (metis \<open>0 < d\<close> zero_less_divide_iff zero_less_numeral)
then have "dist x' y' \<le> diameter S"
using assms diameter_bounded_bound by blast
with x'y' have "dist x y \<le> d / 4 + diameter S + d / 4"
by (meson add_mono_thms_linordered_semiring(1) dist_triangle dist_triangle3 less_eq_real_def order_trans)
then show ?thesis
using xy d_def by linarith
qed
then show "diameter (closure S) \<le> diameter S"
by fastforce
next
show "diameter S \<le> diameter (closure S)"
by (simp add: assms bounded_closure closure_subset diameter_subset)
qed
proposition Lebesgue_number_lemma:
assumes "compact S" "\<C> \<noteq> {}" "S \<subseteq> \<Union>\<C>" and ope: "\<And>B. B \<in> \<C> \<Longrightarrow> open B"
obtains \<delta> where "0 < \<delta>" "\<And>T. \<lbrakk>T \<subseteq> S; diameter T < \<delta>\<rbrakk> \<Longrightarrow> \<exists>B \<in> \<C>. T \<subseteq> B"
proof (cases "S = {}")
case True
then show ?thesis
by (metis \<open>\<C> \<noteq> {}\<close> zero_less_one empty_subsetI equals0I subset_trans that)
next
case False
{ fix x assume "x \<in> S"
then obtain C where C: "x \<in> C" "C \<in> \<C>"
using \<open>S \<subseteq> \<Union>\<C>\<close> by blast
then obtain r where r: "r>0" "ball x (2*r) \<subseteq> C"
by (metis mult.commute mult_2_right not_le ope openE field_sum_of_halves zero_le_numeral zero_less_mult_iff)
then have "\<exists>r C. r > 0 \<and> ball x (2*r) \<subseteq> C \<and> C \<in> \<C>"
using C by blast
}
then obtain r where r: "\<And>x. x \<in> S \<Longrightarrow> r x > 0 \<and> (\<exists>C \<in> \<C>. ball x (2*r x) \<subseteq> C)"
by metis
then have "S \<subseteq> (\<Union>x \<in> S. ball x (r x))"
by auto
then obtain \<T> where "finite \<T>" "S \<subseteq> \<Union>\<T>" and \<T>: "\<T> \<subseteq> (\<lambda>x. ball x (r x)) ` S"
by (rule compactE [OF \<open>compact S\<close>]) auto
then obtain S0 where "S0 \<subseteq> S" "finite S0" and S0: "\<T> = (\<lambda>x. ball x (r x)) ` S0"
by (meson finite_subset_image)
then have "S0 \<noteq> {}"
using False \<open>S \<subseteq> \<Union>\<T>\<close> by auto
define \<delta> where "\<delta> = Inf (r ` S0)"
have "\<delta> > 0"
using \<open>finite S0\<close> \<open>S0 \<subseteq> S\<close> \<open>S0 \<noteq> {}\<close> r by (auto simp: \<delta>_def finite_less_Inf_iff)
show ?thesis
proof
show "0 < \<delta>"
by (simp add: \<open>0 < \<delta>\<close>)
show "\<exists>B \<in> \<C>. T \<subseteq> B" if "T \<subseteq> S" and dia: "diameter T < \<delta>" for T
proof (cases "T = {}")
case True
then show ?thesis
using \<open>\<C> \<noteq> {}\<close> by blast
next
case False
then obtain y where "y \<in> T" by blast
then have "y \<in> S"
using \<open>T \<subseteq> S\<close> by auto
then obtain x where "x \<in> S0" and x: "y \<in> ball x (r x)"
using \<open>S \<subseteq> \<Union>\<T>\<close> S0 that by blast
have "ball y \<delta> \<subseteq> ball y (r x)"
by (metis \<delta>_def \<open>S0 \<noteq> {}\<close> \<open>finite S0\<close> \<open>x \<in> S0\<close> empty_is_image finite_imageI finite_less_Inf_iff imageI less_irrefl not_le subset_ball)
also have "... \<subseteq> ball x (2*r x)"
using x by metric
finally obtain C where "C \<in> \<C>" "ball y \<delta> \<subseteq> C"
by (meson r \<open>S0 \<subseteq> S\<close> \<open>x \<in> S0\<close> dual_order.trans subsetCE)
have "bounded T"
using \<open>compact S\<close> bounded_subset compact_imp_bounded \<open>T \<subseteq> S\<close> by blast
then have "T \<subseteq> ball y \<delta>"
using \<open>y \<in> T\<close> dia diameter_bounded_bound by fastforce
then show ?thesis
apply (rule_tac x=C in bexI)
using \<open>ball y \<delta> \<subseteq> C\<close> \<open>C \<in> \<C>\<close> by auto
qed
qed
qed
subsection \<open>Metric spaces with the Heine-Borel property\<close>
text \<open>
A metric space (or topological vector space) is said to have the
Heine-Borel property if every closed and bounded subset is compact.
\<close>
class heine_borel = metric_space +
assumes bounded_imp_convergent_subsequence:
"bounded (range f) \<Longrightarrow> \<exists>l r. strict_mono (r::nat\<Rightarrow>nat) \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially"
proposition bounded_closed_imp_seq_compact:
fixes s::"'a::heine_borel set"
assumes "bounded s"
and "closed s"
shows "seq_compact s"
proof (unfold seq_compact_def, clarify)
fix f :: "nat \<Rightarrow> 'a"
assume f: "\<forall>n. f n \<in> s"
with \<open>bounded s\<close> have "bounded (range f)"
by (auto intro: bounded_subset)
obtain l r where r: "strict_mono (r :: nat \<Rightarrow> nat)" and l: "((f \<circ> r) \<longlongrightarrow> l) sequentially"
using bounded_imp_convergent_subsequence [OF \<open>bounded (range f)\<close>] by auto
from f have fr: "\<forall>n. (f \<circ> r) n \<in> s"
by simp
have "l \<in> s" using \<open>closed s\<close> fr l
by (rule closed_sequentially)
show "\<exists>l\<in>s. \<exists>r. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially"
using \<open>l \<in> s\<close> r l by blast
qed
lemma compact_eq_bounded_closed:
fixes s :: "'a::heine_borel set"
shows "compact s \<longleftrightarrow> bounded s \<and> closed s"
(is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
using compact_imp_closed compact_imp_bounded
by blast
next
assume ?rhs
then show ?lhs
using bounded_closed_imp_seq_compact[of s]
unfolding compact_eq_seq_compact_metric
by auto
qed
lemma compact_Inter:
fixes \<F> :: "'a :: heine_borel set set"
assumes com: "\<And>S. S \<in> \<F> \<Longrightarrow> compact S" and "\<F> \<noteq> {}"
shows "compact(\<Inter> \<F>)"
using assms
by (meson Inf_lower all_not_in_conv bounded_subset closed_Inter compact_eq_bounded_closed)
lemma compact_closure [simp]:
fixes S :: "'a::heine_borel set"
shows "compact(closure S) \<longleftrightarrow> bounded S"
by (meson bounded_closure bounded_subset closed_closure closure_subset compact_eq_bounded_closed)
instance\<^marker>\<open>tag important\<close> real :: heine_borel
proof
fix f :: "nat \<Rightarrow> real"
assume f: "bounded (range f)"
obtain r :: "nat \<Rightarrow> nat" where r: "strict_mono r" "monoseq (f \<circ> r)"
unfolding comp_def by (metis seq_monosub)
then have "Bseq (f \<circ> r)"
unfolding Bseq_eq_bounded using f
by (metis BseqI' bounded_iff comp_apply rangeI)
with r show "\<exists>l r. strict_mono r \<and> (f \<circ> r) \<longlonglongrightarrow> l"
using Bseq_monoseq_convergent[of "f \<circ> r"] by (auto simp: convergent_def)
qed
lemma compact_lemma_general:
fixes f :: "nat \<Rightarrow> 'a"
fixes proj::"'a \<Rightarrow> 'b \<Rightarrow> 'c::heine_borel" (infixl "proj" 60)
fixes unproj:: "('b \<Rightarrow> 'c) \<Rightarrow> 'a"
assumes finite_basis: "finite basis"
assumes bounded_proj: "\<And>k. k \<in> basis \<Longrightarrow> bounded ((\<lambda>x. x proj k) ` range f)"
assumes proj_unproj: "\<And>e k. k \<in> basis \<Longrightarrow> (unproj e) proj k = e k"
assumes unproj_proj: "\<And>x. unproj (\<lambda>k. x proj k) = x"
shows "\<forall>d\<subseteq>basis. \<exists>l::'a. \<exists> r::nat\<Rightarrow>nat.
strict_mono r \<and> (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) proj i) (l proj i) < e) sequentially)"
proof safe
fix d :: "'b set"
assume d: "d \<subseteq> basis"
with finite_basis have "finite d"
by (blast intro: finite_subset)
from this d show "\<exists>l::'a. \<exists>r::nat\<Rightarrow>nat. strict_mono r \<and>
(\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) proj i) (l proj i) < e) sequentially)"
proof (induct d)
case empty
then show ?case
unfolding strict_mono_def by auto
next
case (insert k d)
have k[intro]: "k \<in> basis"
using insert by auto
have s': "bounded ((\<lambda>x. x proj k) ` range f)"
using k
by (rule bounded_proj)
obtain l1::"'a" and r1 where r1: "strict_mono r1"
and lr1: "\<forall>e > 0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) proj i) (l1 proj i) < e) sequentially"
using insert(3) using insert(4) by auto
have f': "\<forall>n. f (r1 n) proj k \<in> (\<lambda>x. x proj k) ` range f"
by simp
have "bounded (range (\<lambda>i. f (r1 i) proj k))"
by (metis (lifting) bounded_subset f' image_subsetI s')
then obtain l2 r2 where r2:"strict_mono r2" and lr2:"((\<lambda>i. f (r1 (r2 i)) proj k) \<longlongrightarrow> l2) sequentially"
using bounded_imp_convergent_subsequence[of "\<lambda>i. f (r1 i) proj k"]
by (auto simp: o_def)
define r where "r = r1 \<circ> r2"
have r:"strict_mono r"
using r1 and r2 unfolding r_def o_def strict_mono_def by auto
moreover
define l where "l = unproj (\<lambda>i. if i = k then l2 else l1 proj i)"
{
fix e::real
assume "e > 0"
from lr1 \<open>e > 0\<close> have N1: "eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) proj i) (l1 proj i) < e) sequentially"
by blast
from lr2 \<open>e > 0\<close> have N2:"eventually (\<lambda>n. dist (f (r1 (r2 n)) proj k) l2 < e) sequentially"
by (rule tendstoD)
from r2 N1 have N1': "eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 (r2 n)) proj i) (l1 proj i) < e) sequentially"
by (rule eventually_subseq)
have "eventually (\<lambda>n. \<forall>i\<in>(insert k d). dist (f (r n) proj i) (l proj i) < e) sequentially"
using N1' N2
by eventually_elim (insert insert.prems, auto simp: l_def r_def o_def proj_unproj)
}
ultimately show ?case by auto
qed
qed
lemma bounded_fst: "bounded s \<Longrightarrow> bounded (fst ` s)"
unfolding bounded_def
by (metis (erased, hide_lams) dist_fst_le image_iff order_trans)
lemma bounded_snd: "bounded s \<Longrightarrow> bounded (snd ` s)"
unfolding bounded_def
by (metis (no_types, hide_lams) dist_snd_le image_iff order.trans)
instance\<^marker>\<open>tag important\<close> prod :: (heine_borel, heine_borel) heine_borel
proof
fix f :: "nat \<Rightarrow> 'a \<times> 'b"
assume f: "bounded (range f)"
then have "bounded (fst ` range f)"
by (rule bounded_fst)
then have s1: "bounded (range (fst \<circ> f))"
by (simp add: image_comp)
obtain l1 r1 where r1: "strict_mono r1" and l1: "(\<lambda>n. fst (f (r1 n))) \<longlonglongrightarrow> l1"
using bounded_imp_convergent_subsequence [OF s1] unfolding o_def by fast
from f have s2: "bounded (range (snd \<circ> f \<circ> r1))"
by (auto simp: image_comp intro: bounded_snd bounded_subset)
obtain l2 r2 where r2: "strict_mono r2" and l2: "((\<lambda>n. snd (f (r1 (r2 n)))) \<longlongrightarrow> l2) sequentially"
using bounded_imp_convergent_subsequence [OF s2]
unfolding o_def by fast
have l1': "((\<lambda>n. fst (f (r1 (r2 n)))) \<longlongrightarrow> l1) sequentially"
using LIMSEQ_subseq_LIMSEQ [OF l1 r2] unfolding o_def .
have l: "((f \<circ> (r1 \<circ> r2)) \<longlongrightarrow> (l1, l2)) sequentially"
using tendsto_Pair [OF l1' l2] unfolding o_def by simp
have r: "strict_mono (r1 \<circ> r2)"
using r1 r2 unfolding strict_mono_def by simp
show "\<exists>l r. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially"
using l r by fast
qed
subsection \<open>Completeness\<close>
proposition (in metric_space) completeI:
assumes "\<And>f. \<forall>n. f n \<in> s \<Longrightarrow> Cauchy f \<Longrightarrow> \<exists>l\<in>s. f \<longlonglongrightarrow> l"
shows "complete s"
using assms unfolding complete_def by fast
proposition (in metric_space) completeE:
assumes "complete s" and "\<forall>n. f n \<in> s" and "Cauchy f"
obtains l where "l \<in> s" and "f \<longlonglongrightarrow> l"
using assms unfolding complete_def by fast
(* TODO: generalize to uniform spaces *)
lemma compact_imp_complete:
fixes s :: "'a::metric_space set"
assumes "compact s"
shows "complete s"
proof -
{
fix f
assume as: "(\<forall>n::nat. f n \<in> s)" "Cauchy f"
from as(1) obtain l r where lr: "l\<in>s" "strict_mono r" "(f \<circ> r) \<longlonglongrightarrow> l"
using assms unfolding compact_def by blast
note lr' = seq_suble [OF lr(2)]
{
fix e :: real
assume "e > 0"
from as(2) obtain N where N:"\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (f m) (f n) < e/2"
unfolding cauchy_def
using \<open>e > 0\<close>
apply (erule_tac x="e/2" in allE, auto)
done
from lr(3)[unfolded lim_sequentially, THEN spec[where x="e/2"]]
obtain M where M:"\<forall>n\<ge>M. dist ((f \<circ> r) n) l < e/2"
using \<open>e > 0\<close> by auto
{
fix n :: nat
assume n: "n \<ge> max N M"
have "dist ((f \<circ> r) n) l < e/2"
using n M by auto
moreover have "r n \<ge> N"
using lr'[of n] n by auto
then have "dist (f n) ((f \<circ> r) n) < e / 2"
using N and n by auto
ultimately have "dist (f n) l < e" using n M
by metric
}
then have "\<exists>N. \<forall>n\<ge>N. dist (f n) l < e" by blast
}
then have "\<exists>l\<in>s. (f \<longlongrightarrow> l) sequentially" using \<open>l\<in>s\<close>
unfolding lim_sequentially by auto
}
then show ?thesis unfolding complete_def by auto
qed
proposition compact_eq_totally_bounded:
"compact s \<longleftrightarrow> complete s \<and> (\<forall>e>0. \<exists>k. finite k \<and> s \<subseteq> (\<Union>x\<in>k. ball x e))"
(is "_ \<longleftrightarrow> ?rhs")
proof
assume assms: "?rhs"
then obtain k where k: "\<And>e. 0 < e \<Longrightarrow> finite (k e)" "\<And>e. 0 < e \<Longrightarrow> s \<subseteq> (\<Union>x\<in>k e. ball x e)"
by (auto simp: choice_iff')
show "compact s"
proof cases
assume "s = {}"
then show "compact s" by (simp add: compact_def)
next
assume "s \<noteq> {}"
show ?thesis
unfolding compact_def
proof safe
fix f :: "nat \<Rightarrow> 'a"
assume f: "\<forall>n. f n \<in> s"
define e where "e n = 1 / (2 * Suc n)" for n
then have [simp]: "\<And>n. 0 < e n" by auto
define B where "B n U = (SOME b. infinite {n. f n \<in> b} \<and> (\<exists>x. b \<subseteq> ball x (e n) \<inter> U))" for n U
{
fix n U
assume "infinite {n. f n \<in> U}"
then have "\<exists>b\<in>k (e n). infinite {i\<in>{n. f n \<in> U}. f i \<in> ball b (e n)}"
using k f by (intro pigeonhole_infinite_rel) (auto simp: subset_eq)
then obtain a where
"a \<in> k (e n)"
"infinite {i \<in> {n. f n \<in> U}. f i \<in> ball a (e n)}" ..
then have "\<exists>b. infinite {i. f i \<in> b} \<and> (\<exists>x. b \<subseteq> ball x (e n) \<inter> U)"
by (intro exI[of _ "ball a (e n) \<inter> U"] exI[of _ a]) (auto simp: ac_simps)
from someI_ex[OF this]
have "infinite {i. f i \<in> B n U}" "\<exists>x. B n U \<subseteq> ball x (e n) \<inter> U"
unfolding B_def by auto
}
note B = this
define F where "F = rec_nat (B 0 UNIV) B"
{
fix n
have "infinite {i. f i \<in> F n}"
by (induct n) (auto simp: F_def B)
}
then have F: "\<And>n. \<exists>x. F (Suc n) \<subseteq> ball x (e n) \<inter> F n"
using B by (simp add: F_def)
then have F_dec: "\<And>m n. m \<le> n \<Longrightarrow> F n \<subseteq> F m"
using decseq_SucI[of F] by (auto simp: decseq_def)
obtain sel where sel: "\<And>k i. i < sel k i" "\<And>k i. f (sel k i) \<in> F k"
proof (atomize_elim, unfold all_conj_distrib[symmetric], intro choice allI)
fix k i
have "infinite ({n. f n \<in> F k} - {.. i})"
using \<open>infinite {n. f n \<in> F k}\<close> by auto
from infinite_imp_nonempty[OF this]
show "\<exists>x>i. f x \<in> F k"
by (simp add: set_eq_iff not_le conj_commute)
qed
define t where "t = rec_nat (sel 0 0) (\<lambda>n i. sel (Suc n) i)"
have "strict_mono t"
unfolding strict_mono_Suc_iff by (simp add: t_def sel)
moreover have "\<forall>i. (f \<circ> t) i \<in> s"
using f by auto
moreover
{
fix n
have "(f \<circ> t) n \<in> F n"
by (cases n) (simp_all add: t_def sel)
}
note t = this
have "Cauchy (f \<circ> t)"
proof (safe intro!: metric_CauchyI exI elim!: nat_approx_posE)
fix r :: real and N n m
assume "1 / Suc N < r" "Suc N \<le> n" "Suc N \<le> m"
then have "(f \<circ> t) n \<in> F (Suc N)" "(f \<circ> t) m \<in> F (Suc N)" "2 * e N < r"
using F_dec t by (auto simp: e_def field_simps)
with F[of N] obtain x where "dist x ((f \<circ> t) n) < e N" "dist x ((f \<circ> t) m) < e N"
by (auto simp: subset_eq)
with \<open>2 * e N < r\<close> show "dist ((f \<circ> t) m) ((f \<circ> t) n) < r"
by metric
qed
ultimately show "\<exists>l\<in>s. \<exists>r. strict_mono r \<and> (f \<circ> r) \<longlonglongrightarrow> l"
using assms unfolding complete_def by blast
qed
qed
qed (metis compact_imp_complete compact_imp_seq_compact seq_compact_imp_totally_bounded)
lemma cauchy_imp_bounded:
assumes "Cauchy s"
shows "bounded (range s)"
proof -
from assms obtain N :: nat where "\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (s m) (s n) < 1"
unfolding cauchy_def by force
then have N:"\<forall>n. N \<le> n \<longrightarrow> dist (s N) (s n) < 1" by auto
moreover
have "bounded (s ` {0..N})"
using finite_imp_bounded[of "s ` {1..N}"] by auto
then obtain a where a:"\<forall>x\<in>s ` {0..N}. dist (s N) x \<le> a"
unfolding bounded_any_center [where a="s N"] by auto
ultimately show "?thesis"
unfolding bounded_any_center [where a="s N"]
apply (rule_tac x="max a 1" in exI, auto)
apply (erule_tac x=y in allE)
apply (erule_tac x=y in ballE, auto)
done
qed
instance heine_borel < complete_space
proof
fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f"
then have "bounded (range f)"
by (rule cauchy_imp_bounded)
then have "compact (closure (range f))"
unfolding compact_eq_bounded_closed by auto
then have "complete (closure (range f))"
by (rule compact_imp_complete)
moreover have "\<forall>n. f n \<in> closure (range f)"
using closure_subset [of "range f"] by auto
ultimately have "\<exists>l\<in>closure (range f). (f \<longlongrightarrow> l) sequentially"
using \<open>Cauchy f\<close> unfolding complete_def by auto
then show "convergent f"
unfolding convergent_def by auto
qed
lemma complete_UNIV: "complete (UNIV :: ('a::complete_space) set)"
proof (rule completeI)
fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f"
then have "convergent f" by (rule Cauchy_convergent)
then show "\<exists>l\<in>UNIV. f \<longlonglongrightarrow> l" unfolding convergent_def by simp
qed
lemma complete_imp_closed:
fixes S :: "'a::metric_space set"
assumes "complete S"
shows "closed S"
proof (unfold closed_sequential_limits, clarify)
fix f x assume "\<forall>n. f n \<in> S" and "f \<longlonglongrightarrow> x"
from \<open>f \<longlonglongrightarrow> x\<close> have "Cauchy f"
by (rule LIMSEQ_imp_Cauchy)
with \<open>complete S\<close> and \<open>\<forall>n. f n \<in> S\<close> obtain l where "l \<in> S" and "f \<longlonglongrightarrow> l"
by (rule completeE)
from \<open>f \<longlonglongrightarrow> x\<close> and \<open>f \<longlonglongrightarrow> l\<close> have "x = l"
by (rule LIMSEQ_unique)
with \<open>l \<in> S\<close> show "x \<in> S"
by simp
qed
lemma complete_Int_closed:
fixes S :: "'a::metric_space set"
assumes "complete S" and "closed t"
shows "complete (S \<inter> t)"
proof (rule completeI)
fix f assume "\<forall>n. f n \<in> S \<inter> t" and "Cauchy f"
then have "\<forall>n. f n \<in> S" and "\<forall>n. f n \<in> t"
by simp_all
from \<open>complete S\<close> obtain l where "l \<in> S" and "f \<longlonglongrightarrow> l"
using \<open>\<forall>n. f n \<in> S\<close> and \<open>Cauchy f\<close> by (rule completeE)
from \<open>closed t\<close> and \<open>\<forall>n. f n \<in> t\<close> and \<open>f \<longlonglongrightarrow> l\<close> have "l \<in> t"
by (rule closed_sequentially)
with \<open>l \<in> S\<close> and \<open>f \<longlonglongrightarrow> l\<close> show "\<exists>l\<in>S \<inter> t. f \<longlonglongrightarrow> l"
by fast
qed
lemma complete_closed_subset:
fixes S :: "'a::metric_space set"
assumes "closed S" and "S \<subseteq> t" and "complete t"
shows "complete S"
using assms complete_Int_closed [of t S] by (simp add: Int_absorb1)
lemma complete_eq_closed:
fixes S :: "('a::complete_space) set"
shows "complete S \<longleftrightarrow> closed S"
proof
assume "closed S" then show "complete S"
using subset_UNIV complete_UNIV by (rule complete_closed_subset)
next
assume "complete S" then show "closed S"
by (rule complete_imp_closed)
qed
lemma convergent_eq_Cauchy:
fixes S :: "nat \<Rightarrow> 'a::complete_space"
shows "(\<exists>l. (S \<longlongrightarrow> l) sequentially) \<longleftrightarrow> Cauchy S"
unfolding Cauchy_convergent_iff convergent_def ..
lemma convergent_imp_bounded:
fixes S :: "nat \<Rightarrow> 'a::metric_space"
shows "(S \<longlongrightarrow> l) sequentially \<Longrightarrow> bounded (range S)"
by (intro cauchy_imp_bounded LIMSEQ_imp_Cauchy)
lemma frontier_subset_compact:
fixes S :: "'a::heine_borel set"
shows "compact S \<Longrightarrow> frontier S \<subseteq> S"
using frontier_subset_closed compact_eq_bounded_closed
by blast
lemma continuous_closed_imp_Cauchy_continuous:
fixes S :: "('a::complete_space) set"
shows "\<lbrakk>continuous_on S f; closed S; Cauchy \<sigma>; \<And>n. (\<sigma> n) \<in> S\<rbrakk> \<Longrightarrow> Cauchy(f \<circ> \<sigma>)"
apply (simp add: complete_eq_closed [symmetric] continuous_on_sequentially)
by (meson LIMSEQ_imp_Cauchy complete_def)
lemma banach_fix_type:
fixes f::"'a::complete_space\<Rightarrow>'a"
assumes c:"0 \<le> c" "c < 1"
and lipschitz:"\<forall>x. \<forall>y. dist (f x) (f y) \<le> c * dist x y"
shows "\<exists>!x. (f x = x)"
using assms banach_fix[OF complete_UNIV UNIV_not_empty assms(1,2) subset_UNIV, of f]
by auto
subsection\<^marker>\<open>tag unimportant\<close>\<open> Finite intersection property\<close>
text\<open>Also developed in HOL's toplogical spaces theory, but the Heine-Borel type class isn't available there.\<close>
lemma closed_imp_fip:
fixes S :: "'a::heine_borel set"
assumes "closed S"
and T: "T \<in> \<F>" "bounded T"
and clof: "\<And>T. T \<in> \<F> \<Longrightarrow> closed T"
and none: "\<And>\<F>'. \<lbrakk>finite \<F>'; \<F>' \<subseteq> \<F>\<rbrakk> \<Longrightarrow> S \<inter> \<Inter>\<F>' \<noteq> {}"
shows "S \<inter> \<Inter>\<F> \<noteq> {}"
proof -
have "compact (S \<inter> T)"
using \<open>closed S\<close> clof compact_eq_bounded_closed T by blast
then have "(S \<inter> T) \<inter> \<Inter>\<F> \<noteq> {}"
apply (rule compact_imp_fip)
apply (simp add: clof)
by (metis Int_assoc complete_lattice_class.Inf_insert finite_insert insert_subset none \<open>T \<in> \<F>\<close>)
then show ?thesis by blast
qed
lemma closed_imp_fip_compact:
fixes S :: "'a::heine_borel set"
shows
"\<lbrakk>closed S; \<And>T. T \<in> \<F> \<Longrightarrow> compact T;
\<And>\<F>'. \<lbrakk>finite \<F>'; \<F>' \<subseteq> \<F>\<rbrakk> \<Longrightarrow> S \<inter> \<Inter>\<F>' \<noteq> {}\<rbrakk>
\<Longrightarrow> S \<inter> \<Inter>\<F> \<noteq> {}"
by (metis Inf_greatest closed_imp_fip compact_eq_bounded_closed empty_subsetI finite.emptyI inf.orderE)
lemma closed_fip_Heine_Borel:
fixes \<F> :: "'a::heine_borel set set"
assumes "closed S" "T \<in> \<F>" "bounded T"
and "\<And>T. T \<in> \<F> \<Longrightarrow> closed T"
and "\<And>\<F>'. \<lbrakk>finite \<F>'; \<F>' \<subseteq> \<F>\<rbrakk> \<Longrightarrow> \<Inter>\<F>' \<noteq> {}"
shows "\<Inter>\<F> \<noteq> {}"
proof -
have "UNIV \<inter> \<Inter>\<F> \<noteq> {}"
using assms closed_imp_fip [OF closed_UNIV] by auto
then show ?thesis by simp
qed
lemma compact_fip_Heine_Borel:
fixes \<F> :: "'a::heine_borel set set"
assumes clof: "\<And>T. T \<in> \<F> \<Longrightarrow> compact T"
and none: "\<And>\<F>'. \<lbrakk>finite \<F>'; \<F>' \<subseteq> \<F>\<rbrakk> \<Longrightarrow> \<Inter>\<F>' \<noteq> {}"
shows "\<Inter>\<F> \<noteq> {}"
by (metis InterI all_not_in_conv clof closed_fip_Heine_Borel compact_eq_bounded_closed none)
lemma compact_sequence_with_limit:
fixes f :: "nat \<Rightarrow> 'a::heine_borel"
shows "(f \<longlongrightarrow> l) sequentially \<Longrightarrow> compact (insert l (range f))"
apply (simp add: compact_eq_bounded_closed, auto)
apply (simp add: convergent_imp_bounded)
by (simp add: closed_limpt islimpt_insert sequence_unique_limpt)
subsection \<open>Properties of Balls and Spheres\<close>
lemma compact_cball[simp]:
fixes x :: "'a::heine_borel"
shows "compact (cball x e)"
using compact_eq_bounded_closed bounded_cball closed_cball
by blast
lemma compact_frontier_bounded[intro]:
fixes S :: "'a::heine_borel set"
shows "bounded S \<Longrightarrow> compact (frontier S)"
unfolding frontier_def
using compact_eq_bounded_closed
by blast
lemma compact_frontier[intro]:
fixes S :: "'a::heine_borel set"
shows "compact S \<Longrightarrow> compact (frontier S)"
using compact_eq_bounded_closed compact_frontier_bounded
by blast
subsection \<open>Distance from a Set\<close>
lemma distance_attains_sup:
assumes "compact s" "s \<noteq> {}"
shows "\<exists>x\<in>s. \<forall>y\<in>s. dist a y \<le> dist a x"
proof (rule continuous_attains_sup [OF assms])
{
fix x
assume "x\<in>s"
have "(dist a \<longlongrightarrow> dist a x) (at x within s)"
by (intro tendsto_dist tendsto_const tendsto_ident_at)
}
then show "continuous_on s (dist a)"
unfolding continuous_on ..
qed
text \<open>For \emph{minimal} distance, we only need closure, not compactness.\<close>
lemma distance_attains_inf:
fixes a :: "'a::heine_borel"
assumes "closed s" and "s \<noteq> {}"
obtains x where "x\<in>s" "\<And>y. y \<in> s \<Longrightarrow> dist a x \<le> dist a y"
proof -
from assms obtain b where "b \<in> s" by auto
let ?B = "s \<inter> cball a (dist b a)"
have "?B \<noteq> {}" using \<open>b \<in> s\<close>
by (auto simp: dist_commute)
moreover have "continuous_on ?B (dist a)"
by (auto intro!: continuous_at_imp_continuous_on continuous_dist continuous_ident continuous_const)
moreover have "compact ?B"
by (intro closed_Int_compact \<open>closed s\<close> compact_cball)
ultimately obtain x where "x \<in> ?B" "\<forall>y\<in>?B. dist a x \<le> dist a y"
by (metis continuous_attains_inf)
with that show ?thesis by fastforce
qed
subsection \<open>Infimum Distance\<close>
definition\<^marker>\<open>tag important\<close> "infdist x A = (if A = {} then 0 else INF a\<in>A. dist x a)"
lemma bdd_below_image_dist[intro, simp]: "bdd_below (dist x ` A)"
by (auto intro!: zero_le_dist)
lemma infdist_notempty: "A \<noteq> {} \<Longrightarrow> infdist x A = (INF a\<in>A. dist x a)"
by (simp add: infdist_def)
lemma infdist_nonneg: "0 \<le> infdist x A"
by (auto simp: infdist_def intro: cINF_greatest)
lemma infdist_le: "a \<in> A \<Longrightarrow> infdist x A \<le> dist x a"
by (auto intro: cINF_lower simp add: infdist_def)
lemma infdist_le2: "a \<in> A \<Longrightarrow> dist x a \<le> d \<Longrightarrow> infdist x A \<le> d"
by (auto intro!: cINF_lower2 simp add: infdist_def)
lemma infdist_zero[simp]: "a \<in> A \<Longrightarrow> infdist a A = 0"
by (auto intro!: antisym infdist_nonneg infdist_le2)
lemma infdist_Un_min:
assumes "A \<noteq> {}" "B \<noteq> {}"
shows "infdist x (A \<union> B) = min (infdist x A) (infdist x B)"
using assms by (simp add: infdist_def cINF_union inf_real_def)
lemma infdist_triangle: "infdist x A \<le> infdist y A + dist x y"
proof (cases "A = {}")
case True
then show ?thesis by (simp add: infdist_def)
next
case False
then obtain a where "a \<in> A" by auto
have "infdist x A \<le> Inf {dist x y + dist y a |a. a \<in> A}"
proof (rule cInf_greatest)
from \<open>A \<noteq> {}\<close> show "{dist x y + dist y a |a. a \<in> A} \<noteq> {}"
by simp
fix d
assume "d \<in> {dist x y + dist y a |a. a \<in> A}"
then obtain a where d: "d = dist x y + dist y a" "a \<in> A"
by auto
show "infdist x A \<le> d"
unfolding infdist_notempty[OF \<open>A \<noteq> {}\<close>]
proof (rule cINF_lower2)
show "a \<in> A" by fact
show "dist x a \<le> d"
unfolding d by (rule dist_triangle)
qed simp
qed
also have "\<dots> = dist x y + infdist y A"
proof (rule cInf_eq, safe)
fix a
assume "a \<in> A"
then show "dist x y + infdist y A \<le> dist x y + dist y a"
by (auto intro: infdist_le)
next
fix i
assume inf: "\<And>d. d \<in> {dist x y + dist y a |a. a \<in> A} \<Longrightarrow> i \<le> d"
then have "i - dist x y \<le> infdist y A"
unfolding infdist_notempty[OF \<open>A \<noteq> {}\<close>] using \<open>a \<in> A\<close>
by (intro cINF_greatest) (auto simp: field_simps)
then show "i \<le> dist x y + infdist y A"
by simp
qed
finally show ?thesis by simp
qed
lemma infdist_triangle_abs: "\<bar>infdist x A - infdist y A\<bar> \<le> dist x y"
by (metis (full_types) abs_diff_le_iff diff_le_eq dist_commute infdist_triangle)
lemma in_closure_iff_infdist_zero:
assumes "A \<noteq> {}"
shows "x \<in> closure A \<longleftrightarrow> infdist x A = 0"
proof
assume "x \<in> closure A"
show "infdist x A = 0"
proof (rule ccontr)
assume "infdist x A \<noteq> 0"
with infdist_nonneg[of x A] have "infdist x A > 0"
by auto
then have "ball x (infdist x A) \<inter> closure A = {}"
apply auto
apply (metis \<open>x \<in> closure A\<close> closure_approachable dist_commute infdist_le not_less)
done
then have "x \<notin> closure A"
by (metis \<open>0 < infdist x A\<close> centre_in_ball disjoint_iff_not_equal)
then show False using \<open>x \<in> closure A\<close> by simp
qed
next
assume x: "infdist x A = 0"
then obtain a where "a \<in> A"
by atomize_elim (metis all_not_in_conv assms)
show "x \<in> closure A"
unfolding closure_approachable
apply safe
proof (rule ccontr)
fix e :: real
assume "e > 0"
assume "\<not> (\<exists>y\<in>A. dist y x < e)"
then have "infdist x A \<ge> e" using \<open>a \<in> A\<close>
unfolding infdist_def
by (force simp: dist_commute intro: cINF_greatest)
with x \<open>e > 0\<close> show False by auto
qed
qed
lemma in_closed_iff_infdist_zero:
assumes "closed A" "A \<noteq> {}"
shows "x \<in> A \<longleftrightarrow> infdist x A = 0"
proof -
have "x \<in> closure A \<longleftrightarrow> infdist x A = 0"
by (rule in_closure_iff_infdist_zero) fact
with assms show ?thesis by simp
qed
lemma infdist_pos_not_in_closed:
assumes "closed S" "S \<noteq> {}" "x \<notin> S"
shows "infdist x S > 0"
using in_closed_iff_infdist_zero[OF assms(1) assms(2), of x] assms(3) infdist_nonneg le_less by fastforce
lemma
infdist_attains_inf:
fixes X::"'a::heine_borel set"
assumes "closed X"
assumes "X \<noteq> {}"
obtains x where "x \<in> X" "infdist y X = dist y x"
proof -
have "bdd_below (dist y ` X)"
by auto
from distance_attains_inf[OF assms, of y]
obtain x where INF: "x \<in> X" "\<And>z. z \<in> X \<Longrightarrow> dist y x \<le> dist y z" by auto
have "infdist y X = dist y x"
by (auto simp: infdist_def assms
intro!: antisym cINF_lower[OF _ \<open>x \<in> X\<close>] cINF_greatest[OF assms(2) INF(2)])
with \<open>x \<in> X\<close> show ?thesis ..
qed
text \<open>Every metric space is a T4 space:\<close>
instance metric_space \<subseteq> t4_space
proof
fix S T::"'a set" assume H: "closed S" "closed T" "S \<inter> T = {}"
consider "S = {}" | "T = {}" | "S \<noteq> {} \<and> T \<noteq> {}" by auto
then show "\<exists>U V. open U \<and> open V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> U \<inter> V = {}"
proof (cases)
case 1
show ?thesis
apply (rule exI[of _ "{}"], rule exI[of _ UNIV]) using 1 by auto
next
case 2
show ?thesis
apply (rule exI[of _ UNIV], rule exI[of _ "{}"]) using 2 by auto
next
case 3
define U where "U = (\<Union>x\<in>S. ball x ((infdist x T)/2))"
have A: "open U" unfolding U_def by auto
have "infdist x T > 0" if "x \<in> S" for x
using H that 3 by (auto intro!: infdist_pos_not_in_closed)
then have B: "S \<subseteq> U" unfolding U_def by auto
define V where "V = (\<Union>x\<in>T. ball x ((infdist x S)/2))"
have C: "open V" unfolding V_def by auto
have "infdist x S > 0" if "x \<in> T" for x
using H that 3 by (auto intro!: infdist_pos_not_in_closed)
then have D: "T \<subseteq> V" unfolding V_def by auto
have "(ball x ((infdist x T)/2)) \<inter> (ball y ((infdist y S)/2)) = {}" if "x \<in> S" "y \<in> T" for x y
proof (auto)
fix z assume H: "dist x z * 2 < infdist x T" "dist y z * 2 < infdist y S"
have "2 * dist x y \<le> 2 * dist x z + 2 * dist y z"
by metric
also have "... < infdist x T + infdist y S"
using H by auto
finally have "dist x y < infdist x T \<or> dist x y < infdist y S"
by auto
then show False
using infdist_le[OF \<open>x \<in> S\<close>, of y] infdist_le[OF \<open>y \<in> T\<close>, of x] by (auto simp add: dist_commute)
qed
then have E: "U \<inter> V = {}"
unfolding U_def V_def by auto
show ?thesis
apply (rule exI[of _ U], rule exI[of _ V]) using A B C D E by auto
qed
qed
lemma tendsto_infdist [tendsto_intros]:
assumes f: "(f \<longlongrightarrow> l) F"
shows "((\<lambda>x. infdist (f x) A) \<longlongrightarrow> infdist l A) F"
proof (rule tendstoI)
fix e ::real
assume "e > 0"
from tendstoD[OF f this]
show "eventually (\<lambda>x. dist (infdist (f x) A) (infdist l A) < e) F"
proof (eventually_elim)
fix x
from infdist_triangle[of l A "f x"] infdist_triangle[of "f x" A l]
have "dist (infdist (f x) A) (infdist l A) \<le> dist (f x) l"
by (simp add: dist_commute dist_real_def)
also assume "dist (f x) l < e"
finally show "dist (infdist (f x) A) (infdist l A) < e" .
qed
qed
lemma continuous_infdist[continuous_intros]:
assumes "continuous F f"
shows "continuous F (\<lambda>x. infdist (f x) A)"
using assms unfolding continuous_def by (rule tendsto_infdist)
lemma continuous_on_infdist [continuous_intros]:
assumes "continuous_on S f"
shows "continuous_on S (\<lambda>x. infdist (f x) A)"
using assms unfolding continuous_on by (auto intro: tendsto_infdist)
lemma compact_infdist_le:
fixes A::"'a::heine_borel set"
assumes "A \<noteq> {}"
assumes "compact A"
assumes "e > 0"
shows "compact {x. infdist x A \<le> e}"
proof -
from continuous_closed_vimage[of "{0..e}" "\<lambda>x. infdist x A"]
continuous_infdist[OF continuous_ident, of _ UNIV A]
have "closed {x. infdist x A \<le> e}" by (auto simp: vimage_def infdist_nonneg)
moreover
from assms obtain x0 b where b: "\<And>x. x \<in> A \<Longrightarrow> dist x0 x \<le> b" "closed A"
by (auto simp: compact_eq_bounded_closed bounded_def)
{
fix y
assume "infdist y A \<le> e"
moreover
from infdist_attains_inf[OF \<open>closed A\<close> \<open>A \<noteq> {}\<close>, of y]
obtain z where "z \<in> A" "infdist y A = dist y z" by blast
ultimately
have "dist x0 y \<le> b + e" using b by metric
} then
have "bounded {x. infdist x A \<le> e}"
by (auto simp: bounded_any_center[where a=x0] intro!: exI[where x="b + e"])
ultimately show "compact {x. infdist x A \<le> e}"
by (simp add: compact_eq_bounded_closed)
qed
subsection \<open>Separation between Points and Sets\<close>
proposition separate_point_closed:
fixes s :: "'a::heine_borel set"
assumes "closed s" and "a \<notin> s"
shows "\<exists>d>0. \<forall>x\<in>s. d \<le> dist a x"
proof (cases "s = {}")
case True
then show ?thesis by(auto intro!: exI[where x=1])
next
case False
from assms obtain x where "x\<in>s" "\<forall>y\<in>s. dist a x \<le> dist a y"
using \<open>s \<noteq> {}\<close> by (blast intro: distance_attains_inf [of s a])
with \<open>x\<in>s\<close> show ?thesis using dist_pos_lt[of a x] and\<open>a \<notin> s\<close>
by blast
qed
proposition separate_compact_closed:
fixes s t :: "'a::heine_borel set"
assumes "compact s"
and t: "closed t" "s \<inter> t = {}"
shows "\<exists>d>0. \<forall>x\<in>s. \<forall>y\<in>t. d \<le> dist x y"
proof cases
assume "s \<noteq> {} \<and> t \<noteq> {}"
then have "s \<noteq> {}" "t \<noteq> {}" by auto
let ?inf = "\<lambda>x. infdist x t"
have "continuous_on s ?inf"
by (auto intro!: continuous_at_imp_continuous_on continuous_infdist continuous_ident)
then obtain x where x: "x \<in> s" "\<forall>y\<in>s. ?inf x \<le> ?inf y"
using continuous_attains_inf[OF \<open>compact s\<close> \<open>s \<noteq> {}\<close>] by auto
then have "0 < ?inf x"
using t \<open>t \<noteq> {}\<close> in_closed_iff_infdist_zero by (auto simp: less_le infdist_nonneg)
moreover have "\<forall>x'\<in>s. \<forall>y\<in>t. ?inf x \<le> dist x' y"
using x by (auto intro: order_trans infdist_le)
ultimately show ?thesis by auto
qed (auto intro!: exI[of _ 1])
proposition separate_closed_compact:
fixes s t :: "'a::heine_borel set"
assumes "closed s"
and "compact t"
and "s \<inter> t = {}"
shows "\<exists>d>0. \<forall>x\<in>s. \<forall>y\<in>t. d \<le> dist x y"
proof -
have *: "t \<inter> s = {}"
using assms(3) by auto
show ?thesis
using separate_compact_closed[OF assms(2,1) *] by (force simp: dist_commute)
qed
proposition compact_in_open_separated:
fixes A::"'a::heine_borel set"
assumes "A \<noteq> {}"
assumes "compact A"
assumes "open B"
assumes "A \<subseteq> B"
obtains e where "e > 0" "{x. infdist x A \<le> e} \<subseteq> B"
proof atomize_elim
have "closed (- B)" "compact A" "- B \<inter> A = {}"
using assms by (auto simp: open_Diff compact_eq_bounded_closed)
from separate_closed_compact[OF this]
obtain d'::real where d': "d'>0" "\<And>x y. x \<notin> B \<Longrightarrow> y \<in> A \<Longrightarrow> d' \<le> dist x y"
by auto
define d where "d = d' / 2"
hence "d>0" "d < d'" using d' by auto
with d' have d: "\<And>x y. x \<notin> B \<Longrightarrow> y \<in> A \<Longrightarrow> d < dist x y"
by force
show "\<exists>e>0. {x. infdist x A \<le> e} \<subseteq> B"
proof (rule ccontr)
assume "\<nexists>e. 0 < e \<and> {x. infdist x A \<le> e} \<subseteq> B"
with \<open>d > 0\<close> obtain x where x: "infdist x A \<le> d" "x \<notin> B"
by auto
from assms have "closed A" "A \<noteq> {}" by (auto simp: compact_eq_bounded_closed)
from infdist_attains_inf[OF this]
obtain y where y: "y \<in> A" "infdist x A = dist x y"
by auto
have "dist x y \<le> d" using x y by simp
also have "\<dots> < dist x y" using y d x by auto
finally show False by simp
qed
qed
subsection \<open>Uniform Continuity\<close>
lemma uniformly_continuous_onE:
assumes "uniformly_continuous_on s f" "0 < e"
obtains d where "d>0" "\<And>x x'. \<lbrakk>x\<in>s; x'\<in>s; dist x' x < d\<rbrakk> \<Longrightarrow> dist (f x') (f x) < e"
using assms
by (auto simp: uniformly_continuous_on_def)
lemma uniformly_continuous_on_sequentially:
"uniformly_continuous_on s f \<longleftrightarrow> (\<forall>x y. (\<forall>n. x n \<in> s) \<and> (\<forall>n. y n \<in> s) \<and>
(\<lambda>n. dist (x n) (y n)) \<longlonglongrightarrow> 0 \<longrightarrow> (\<lambda>n. dist (f(x n)) (f(y n))) \<longlonglongrightarrow> 0)" (is "?lhs = ?rhs")
proof
assume ?lhs
{
fix x y
assume x: "\<forall>n. x n \<in> s"
and y: "\<forall>n. y n \<in> s"
and xy: "((\<lambda>n. dist (x n) (y n)) \<longlongrightarrow> 0) sequentially"
{
fix e :: real
assume "e > 0"
then obtain d where "d > 0" and d: "\<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e"
using \<open>?lhs\<close>[unfolded uniformly_continuous_on_def, THEN spec[where x=e]] by auto
obtain N where N: "\<forall>n\<ge>N. dist (x n) (y n) < d"
using xy[unfolded lim_sequentially dist_norm] and \<open>d>0\<close> by auto
{
fix n
assume "n\<ge>N"
then have "dist (f (x n)) (f (y n)) < e"
using N[THEN spec[where x=n]]
using d[THEN bspec[where x="x n"], THEN bspec[where x="y n"]]
using x and y
by (simp add: dist_commute)
}
then have "\<exists>N. \<forall>n\<ge>N. dist (f (x n)) (f (y n)) < e"
by auto
}
then have "((\<lambda>n. dist (f(x n)) (f(y n))) \<longlongrightarrow> 0) sequentially"
unfolding lim_sequentially and dist_real_def by auto
}
then show ?rhs by auto
next
assume ?rhs
{
assume "\<not> ?lhs"
then obtain e where "e > 0" "\<forall>d>0. \<exists>x\<in>s. \<exists>x'\<in>s. dist x' x < d \<and> \<not> dist (f x') (f x) < e"
unfolding uniformly_continuous_on_def by auto
then obtain fa where fa:
"\<forall>x. 0 < x \<longrightarrow> fst (fa x) \<in> s \<and> snd (fa x) \<in> s \<and> dist (fst (fa x)) (snd (fa x)) < x \<and> \<not> dist (f (fst (fa x))) (f (snd (fa x))) < e"
using choice[of "\<lambda>d x. d>0 \<longrightarrow> fst x \<in> s \<and> snd x \<in> s \<and> dist (snd x) (fst x) < d \<and> \<not> dist (f (snd x)) (f (fst x)) < e"]
unfolding Bex_def
by (auto simp: dist_commute)
define x where "x n = fst (fa (inverse (real n + 1)))" for n
define y where "y n = snd (fa (inverse (real n + 1)))" for n
have xyn: "\<forall>n. x n \<in> s \<and> y n \<in> s"
and xy0: "\<forall>n. dist (x n) (y n) < inverse (real n + 1)"
and fxy:"\<forall>n. \<not> dist (f (x n)) (f (y n)) < e"
unfolding x_def and y_def using fa
by auto
{
fix e :: real
assume "e > 0"
then obtain N :: nat where "N \<noteq> 0" and N: "0 < inverse (real N) \<and> inverse (real N) < e"
unfolding real_arch_inverse[of e] by auto
{
fix n :: nat
assume "n \<ge> N"
then have "inverse (real n + 1) < inverse (real N)"
using of_nat_0_le_iff and \<open>N\<noteq>0\<close> by auto
also have "\<dots> < e" using N by auto
finally have "inverse (real n + 1) < e" by auto
then have "dist (x n) (y n) < e"
using xy0[THEN spec[where x=n]] by auto
}
then have "\<exists>N. \<forall>n\<ge>N. dist (x n) (y n) < e" by auto
}
then have "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (f (x n)) (f (y n)) < e"
using \<open>?rhs\<close>[THEN spec[where x=x], THEN spec[where x=y]] and xyn
unfolding lim_sequentially dist_real_def by auto
then have False using fxy and \<open>e>0\<close> by auto
}
then show ?lhs
unfolding uniformly_continuous_on_def by blast
qed
subsection \<open>Continuity on a Compact Domain Implies Uniform Continuity\<close>
text\<open>From the proof of the Heine-Borel theorem: Lemma 2 in section 3.7, page 69 of
J. C. Burkill and H. Burkill. A Second Course in Mathematical Analysis (CUP, 2002)\<close>
lemma Heine_Borel_lemma:
assumes "compact S" and Ssub: "S \<subseteq> \<Union>\<G>" and opn: "\<And>G. G \<in> \<G> \<Longrightarrow> open G"
obtains e where "0 < e" "\<And>x. x \<in> S \<Longrightarrow> \<exists>G \<in> \<G>. ball x e \<subseteq> G"
proof -
have False if neg: "\<And>e. 0 < e \<Longrightarrow> \<exists>x \<in> S. \<forall>G \<in> \<G>. \<not> ball x e \<subseteq> G"
proof -
have "\<exists>x \<in> S. \<forall>G \<in> \<G>. \<not> ball x (1 / Suc n) \<subseteq> G" for n
using neg by simp
then obtain f where "\<And>n. f n \<in> S" and fG: "\<And>G n. G \<in> \<G> \<Longrightarrow> \<not> ball (f n) (1 / Suc n) \<subseteq> G"
by metis
then obtain l r where "l \<in> S" "strict_mono r" and to_l: "(f \<circ> r) \<longlonglongrightarrow> l"
using \<open>compact S\<close> compact_def that by metis
then obtain G where "l \<in> G" "G \<in> \<G>"
using Ssub by auto
then obtain e where "0 < e" and e: "\<And>z. dist z l < e \<Longrightarrow> z \<in> G"
using opn open_dist by blast
obtain N1 where N1: "\<And>n. n \<ge> N1 \<Longrightarrow> dist (f (r n)) l < e/2"
using to_l apply (simp add: lim_sequentially)
using \<open>0 < e\<close> half_gt_zero that by blast
obtain N2 where N2: "of_nat N2 > 2/e"
using reals_Archimedean2 by blast
obtain x where "x \<in> ball (f (r (max N1 N2))) (1 / real (Suc (r (max N1 N2))))" and "x \<notin> G"
using fG [OF \<open>G \<in> \<G>\<close>, of "r (max N1 N2)"] by blast
then have "dist (f (r (max N1 N2))) x < 1 / real (Suc (r (max N1 N2)))"
by simp
also have "... \<le> 1 / real (Suc (max N1 N2))"
apply (simp add: field_split_simps del: max.bounded_iff)
using \<open>strict_mono r\<close> seq_suble by blast
also have "... \<le> 1 / real (Suc N2)"
by (simp add: field_simps)
also have "... < e/2"
using N2 \<open>0 < e\<close> by (simp add: field_simps)
finally have "dist (f (r (max N1 N2))) x < e / 2" .
moreover have "dist (f (r (max N1 N2))) l < e/2"
using N1 max.cobounded1 by blast
ultimately have "dist x l < e"
by metric
then show ?thesis
using e \<open>x \<notin> G\<close> by blast
qed
then show ?thesis
by (meson that)
qed
lemma compact_uniformly_equicontinuous:
assumes "compact S"
and cont: "\<And>x e. \<lbrakk>x \<in> S; 0 < e\<rbrakk>
\<Longrightarrow> \<exists>d. 0 < d \<and>
(\<forall>f \<in> \<F>. \<forall>x' \<in> S. dist x' x < d \<longrightarrow> dist (f x') (f x) < e)"
and "0 < e"
obtains d where "0 < d"
"\<And>f x x'. \<lbrakk>f \<in> \<F>; x \<in> S; x' \<in> S; dist x' x < d\<rbrakk> \<Longrightarrow> dist (f x') (f x) < e"
proof -
obtain d where d_pos: "\<And>x e. \<lbrakk>x \<in> S; 0 < e\<rbrakk> \<Longrightarrow> 0 < d x e"
and d_dist : "\<And>x x' e f. \<lbrakk>dist x' x < d x e; x \<in> S; x' \<in> S; 0 < e; f \<in> \<F>\<rbrakk> \<Longrightarrow> dist (f x') (f x) < e"
using cont by metis
let ?\<G> = "((\<lambda>x. ball x (d x (e / 2))) ` S)"
have Ssub: "S \<subseteq> \<Union> ?\<G>"
by clarsimp (metis d_pos \<open>0 < e\<close> dist_self half_gt_zero_iff)
then obtain k where "0 < k" and k: "\<And>x. x \<in> S \<Longrightarrow> \<exists>G \<in> ?\<G>. ball x k \<subseteq> G"
by (rule Heine_Borel_lemma [OF \<open>compact S\<close>]) auto
moreover have "dist (f v) (f u) < e" if "f \<in> \<F>" "u \<in> S" "v \<in> S" "dist v u < k" for f u v
proof -
obtain G where "G \<in> ?\<G>" "u \<in> G" "v \<in> G"
using k that
by (metis \<open>dist v u < k\<close> \<open>u \<in> S\<close> \<open>0 < k\<close> centre_in_ball subsetD dist_commute mem_ball)
then obtain w where w: "dist w u < d w (e / 2)" "dist w v < d w (e / 2)" "w \<in> S"
by auto
with that d_dist have "dist (f w) (f v) < e/2"
by (metis \<open>0 < e\<close> dist_commute half_gt_zero)
moreover
have "dist (f w) (f u) < e/2"
using that d_dist w by (metis \<open>0 < e\<close> dist_commute divide_pos_pos zero_less_numeral)
ultimately show ?thesis
using dist_triangle_half_r by blast
qed
ultimately show ?thesis using that by blast
qed
corollary compact_uniformly_continuous:
fixes f :: "'a :: metric_space \<Rightarrow> 'b :: metric_space"
assumes f: "continuous_on S f" and S: "compact S"
shows "uniformly_continuous_on S f"
using f
unfolding continuous_on_iff uniformly_continuous_on_def
by (force intro: compact_uniformly_equicontinuous [OF S, of "{f}"])
subsection\<^marker>\<open>tag unimportant\<close>\<open> Theorems relating continuity and uniform continuity to closures\<close>
lemma continuous_on_closure:
"continuous_on (closure S) f \<longleftrightarrow>
(\<forall>x e. x \<in> closure S \<and> 0 < e
\<longrightarrow> (\<exists>d. 0 < d \<and> (\<forall>y. y \<in> S \<and> dist y x < d \<longrightarrow> dist (f y) (f x) < e)))"
(is "?lhs = ?rhs")
proof
assume ?lhs then show ?rhs
unfolding continuous_on_iff by (metis Un_iff closure_def)
next
assume R [rule_format]: ?rhs
show ?lhs
proof
fix x and e::real
assume "0 < e" and x: "x \<in> closure S"
obtain \<delta>::real where "\<delta> > 0"
and \<delta>: "\<And>y. \<lbrakk>y \<in> S; dist y x < \<delta>\<rbrakk> \<Longrightarrow> dist (f y) (f x) < e/2"
using R [of x "e/2"] \<open>0 < e\<close> x by auto
have "dist (f y) (f x) \<le> e" if y: "y \<in> closure S" and dyx: "dist y x < \<delta>/2" for y
proof -
obtain \<delta>'::real where "\<delta>' > 0"
and \<delta>': "\<And>z. \<lbrakk>z \<in> S; dist z y < \<delta>'\<rbrakk> \<Longrightarrow> dist (f z) (f y) < e/2"
using R [of y "e/2"] \<open>0 < e\<close> y by auto
obtain z where "z \<in> S" and z: "dist z y < min \<delta>' \<delta> / 2"
using closure_approachable y
by (metis \<open>0 < \<delta>'\<close> \<open>0 < \<delta>\<close> divide_pos_pos min_less_iff_conj zero_less_numeral)
have "dist (f z) (f y) < e/2"
using \<delta>' [OF \<open>z \<in> S\<close>] z \<open>0 < \<delta>'\<close> by metric
moreover have "dist (f z) (f x) < e/2"
using \<delta>[OF \<open>z \<in> S\<close>] z dyx by metric
ultimately show ?thesis
by metric
qed
then show "\<exists>d>0. \<forall>x'\<in>closure S. dist x' x < d \<longrightarrow> dist (f x') (f x) \<le> e"
by (rule_tac x="\<delta>/2" in exI) (simp add: \<open>\<delta> > 0\<close>)
qed
qed
lemma continuous_on_closure_sequentially:
fixes f :: "'a::metric_space \<Rightarrow> 'b :: metric_space"
shows
"continuous_on (closure S) f \<longleftrightarrow>
(\<forall>x a. a \<in> closure S \<and> (\<forall>n. x n \<in> S) \<and> x \<longlonglongrightarrow> a \<longrightarrow> (f \<circ> x) \<longlonglongrightarrow> f a)"
(is "?lhs = ?rhs")
proof -
have "continuous_on (closure S) f \<longleftrightarrow>
(\<forall>x \<in> closure S. continuous (at x within S) f)"
by (force simp: continuous_on_closure continuous_within_eps_delta)
also have "... = ?rhs"
by (force simp: continuous_within_sequentially)
finally show ?thesis .
qed
lemma uniformly_continuous_on_closure:
fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space"
assumes ucont: "uniformly_continuous_on S f"
and cont: "continuous_on (closure S) f"
shows "uniformly_continuous_on (closure S) f"
unfolding uniformly_continuous_on_def
proof (intro allI impI)
fix e::real
assume "0 < e"
then obtain d::real
where "d>0"
and d: "\<And>x x'. \<lbrakk>x\<in>S; x'\<in>S; dist x' x < d\<rbrakk> \<Longrightarrow> dist (f x') (f x) < e/3"
using ucont [unfolded uniformly_continuous_on_def, rule_format, of "e/3"] by auto
show "\<exists>d>0. \<forall>x\<in>closure S. \<forall>x'\<in>closure S. dist x' x < d \<longrightarrow> dist (f x') (f x) < e"
proof (rule exI [where x="d/3"], clarsimp simp: \<open>d > 0\<close>)
fix x y
assume x: "x \<in> closure S" and y: "y \<in> closure S" and dyx: "dist y x * 3 < d"
obtain d1::real where "d1 > 0"
and d1: "\<And>w. \<lbrakk>w \<in> closure S; dist w x < d1\<rbrakk> \<Longrightarrow> dist (f w) (f x) < e/3"
using cont [unfolded continuous_on_iff, rule_format, of "x" "e/3"] \<open>0 < e\<close> x by auto
obtain x' where "x' \<in> S" and x': "dist x' x < min d1 (d / 3)"
using closure_approachable [of x S]
by (metis \<open>0 < d1\<close> \<open>0 < d\<close> divide_pos_pos min_less_iff_conj x zero_less_numeral)
obtain d2::real where "d2 > 0"
and d2: "\<forall>w \<in> closure S. dist w y < d2 \<longrightarrow> dist (f w) (f y) < e/3"
using cont [unfolded continuous_on_iff, rule_format, of "y" "e/3"] \<open>0 < e\<close> y by auto
obtain y' where "y' \<in> S" and y': "dist y' y < min d2 (d / 3)"
using closure_approachable [of y S]
by (metis \<open>0 < d2\<close> \<open>0 < d\<close> divide_pos_pos min_less_iff_conj y zero_less_numeral)
have "dist x' x < d/3" using x' by auto
then have "dist x' y' < d"
using dyx y' by metric
then have "dist (f x') (f y') < e/3"
by (rule d [OF \<open>y' \<in> S\<close> \<open>x' \<in> S\<close>])
moreover have "dist (f x') (f x) < e/3" using \<open>x' \<in> S\<close> closure_subset x' d1
by (simp add: closure_def)
moreover have "dist (f y') (f y) < e/3" using \<open>y' \<in> S\<close> closure_subset y' d2
by (simp add: closure_def)
ultimately show "dist (f y) (f x) < e" by metric
qed
qed
lemma uniformly_continuous_on_extension_at_closure:
fixes f::"'a::metric_space \<Rightarrow> 'b::complete_space"
assumes uc: "uniformly_continuous_on X f"
assumes "x \<in> closure X"
obtains l where "(f \<longlongrightarrow> l) (at x within X)"
proof -
from assms obtain xs where xs: "xs \<longlonglongrightarrow> x" "\<And>n. xs n \<in> X"
by (auto simp: closure_sequential)
from uniformly_continuous_on_Cauchy[OF uc LIMSEQ_imp_Cauchy, OF xs]
obtain l where l: "(\<lambda>n. f (xs n)) \<longlonglongrightarrow> l"
by atomize_elim (simp only: convergent_eq_Cauchy)
have "(f \<longlongrightarrow> l) (at x within X)"
proof (safe intro!: Lim_within_LIMSEQ)
fix xs'
assume "\<forall>n. xs' n \<noteq> x \<and> xs' n \<in> X"
and xs': "xs' \<longlonglongrightarrow> x"
then have "xs' n \<noteq> x" "xs' n \<in> X" for n by auto
from uniformly_continuous_on_Cauchy[OF uc LIMSEQ_imp_Cauchy, OF \<open>xs' \<longlonglongrightarrow> x\<close> \<open>xs' _ \<in> X\<close>]
obtain l' where l': "(\<lambda>n. f (xs' n)) \<longlonglongrightarrow> l'"
by atomize_elim (simp only: convergent_eq_Cauchy)
show "(\<lambda>n. f (xs' n)) \<longlonglongrightarrow> l"
proof (rule tendstoI)
fix e::real assume "e > 0"
define e' where "e' \<equiv> e / 2"
have "e' > 0" using \<open>e > 0\<close> by (simp add: e'_def)
have "\<forall>\<^sub>F n in sequentially. dist (f (xs n)) l < e'"
by (simp add: \<open>0 < e'\<close> l tendstoD)
moreover
from uc[unfolded uniformly_continuous_on_def, rule_format, OF \<open>e' > 0\<close>]
obtain d where d: "d > 0" "\<And>x x'. x \<in> X \<Longrightarrow> x' \<in> X \<Longrightarrow> dist x x' < d \<Longrightarrow> dist (f x) (f x') < e'"
by auto
have "\<forall>\<^sub>F n in sequentially. dist (xs n) (xs' n) < d"
by (auto intro!: \<open>0 < d\<close> order_tendstoD tendsto_eq_intros xs xs')
ultimately
show "\<forall>\<^sub>F n in sequentially. dist (f (xs' n)) l < e"
proof eventually_elim
case (elim n)
have "dist (f (xs' n)) l \<le> dist (f (xs n)) (f (xs' n)) + dist (f (xs n)) l"
by metric
also have "dist (f (xs n)) (f (xs' n)) < e'"
by (auto intro!: d xs \<open>xs' _ \<in> _\<close> elim)
also note \<open>dist (f (xs n)) l < e'\<close>
also have "e' + e' = e" by (simp add: e'_def)
finally show ?case by simp
qed
qed
qed
thus ?thesis ..
qed
lemma uniformly_continuous_on_extension_on_closure:
fixes f::"'a::metric_space \<Rightarrow> 'b::complete_space"
assumes uc: "uniformly_continuous_on X f"
obtains g where "uniformly_continuous_on (closure X) g" "\<And>x. x \<in> X \<Longrightarrow> f x = g x"
"\<And>Y h x. X \<subseteq> Y \<Longrightarrow> Y \<subseteq> closure X \<Longrightarrow> continuous_on Y h \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> f x = h x) \<Longrightarrow> x \<in> Y \<Longrightarrow> h x = g x"
proof -
from uc have cont_f: "continuous_on X f"
by (simp add: uniformly_continuous_imp_continuous)
obtain y where y: "(f \<longlongrightarrow> y x) (at x within X)" if "x \<in> closure X" for x
apply atomize_elim
apply (rule choice)
using uniformly_continuous_on_extension_at_closure[OF assms]
by metis
let ?g = "\<lambda>x. if x \<in> X then f x else y x"
have "uniformly_continuous_on (closure X) ?g"
unfolding uniformly_continuous_on_def
proof safe
fix e::real assume "e > 0"
define e' where "e' \<equiv> e / 3"
have "e' > 0" using \<open>e > 0\<close> by (simp add: e'_def)
from uc[unfolded uniformly_continuous_on_def, rule_format, OF \<open>0 < e'\<close>]
obtain d where "d > 0" and d: "\<And>x x'. x \<in> X \<Longrightarrow> x' \<in> X \<Longrightarrow> dist x' x < d \<Longrightarrow> dist (f x') (f x) < e'"
by auto
define d' where "d' = d / 3"
have "d' > 0" using \<open>d > 0\<close> by (simp add: d'_def)
show "\<exists>d>0. \<forall>x\<in>closure X. \<forall>x'\<in>closure X. dist x' x < d \<longrightarrow> dist (?g x') (?g x) < e"
proof (safe intro!: exI[where x=d'] \<open>d' > 0\<close>)
fix x x' assume x: "x \<in> closure X" and x': "x' \<in> closure X" and dist: "dist x' x < d'"
then obtain xs xs' where xs: "xs \<longlonglongrightarrow> x" "\<And>n. xs n \<in> X"
and xs': "xs' \<longlonglongrightarrow> x'" "\<And>n. xs' n \<in> X"
by (auto simp: closure_sequential)
have "\<forall>\<^sub>F n in sequentially. dist (xs' n) x' < d'"
and "\<forall>\<^sub>F n in sequentially. dist (xs n) x < d'"
by (auto intro!: \<open>0 < d'\<close> order_tendstoD tendsto_eq_intros xs xs')
moreover
have "(\<lambda>x. f (xs x)) \<longlonglongrightarrow> y x" if "x \<in> closure X" "x \<notin> X" "xs \<longlonglongrightarrow> x" "\<And>n. xs n \<in> X" for xs x
using that not_eventuallyD
by (force intro!: filterlim_compose[OF y[OF \<open>x \<in> closure X\<close>]] simp: filterlim_at)
then have "(\<lambda>x. f (xs' x)) \<longlonglongrightarrow> ?g x'" "(\<lambda>x. f (xs x)) \<longlonglongrightarrow> ?g x"
using x x'
by (auto intro!: continuous_on_tendsto_compose[OF cont_f] simp: xs' xs)
then have "\<forall>\<^sub>F n in sequentially. dist (f (xs' n)) (?g x') < e'"
"\<forall>\<^sub>F n in sequentially. dist (f (xs n)) (?g x) < e'"
by (auto intro!: \<open>0 < e'\<close> order_tendstoD tendsto_eq_intros)
ultimately
have "\<forall>\<^sub>F n in sequentially. dist (?g x') (?g x) < e"
proof eventually_elim
case (elim n)
have "dist (?g x') (?g x) \<le>
dist (f (xs' n)) (?g x') + dist (f (xs' n)) (f (xs n)) + dist (f (xs n)) (?g x)"
by (metis add.commute add_le_cancel_left dist_commute dist_triangle dist_triangle_le)
also
from \<open>dist (xs' n) x' < d'\<close> \<open>dist x' x < d'\<close> \<open>dist (xs n) x < d'\<close>
have "dist (xs' n) (xs n) < d" unfolding d'_def by metric
with \<open>xs _ \<in> X\<close> \<open>xs' _ \<in> X\<close> have "dist (f (xs' n)) (f (xs n)) < e'"
by (rule d)
also note \<open>dist (f (xs' n)) (?g x') < e'\<close>
also note \<open>dist (f (xs n)) (?g x) < e'\<close>
finally show ?case by (simp add: e'_def)
qed
then show "dist (?g x') (?g x) < e" by simp
qed
qed
moreover have "f x = ?g x" if "x \<in> X" for x using that by simp
moreover
{
fix Y h x
assume Y: "x \<in> Y" "X \<subseteq> Y" "Y \<subseteq> closure X" and cont_h: "continuous_on Y h"
and extension: "(\<And>x. x \<in> X \<Longrightarrow> f x = h x)"
{
assume "x \<notin> X"
have "x \<in> closure X" using Y by auto
then obtain xs where xs: "xs \<longlonglongrightarrow> x" "\<And>n. xs n \<in> X"
by (auto simp: closure_sequential)
from continuous_on_tendsto_compose[OF cont_h xs(1)] xs(2) Y
have hx: "(\<lambda>x. f (xs x)) \<longlonglongrightarrow> h x"
by (auto simp: subsetD extension)
then have "(\<lambda>x. f (xs x)) \<longlonglongrightarrow> y x"
using \<open>x \<notin> X\<close> not_eventuallyD xs(2)
by (force intro!: filterlim_compose[OF y[OF \<open>x \<in> closure X\<close>]] simp: filterlim_at xs)
with hx have "h x = y x" by (rule LIMSEQ_unique)
} then
have "h x = ?g x"
using extension by auto
}
ultimately show ?thesis ..
qed
lemma bounded_uniformly_continuous_image:
fixes f :: "'a :: heine_borel \<Rightarrow> 'b :: heine_borel"
assumes "uniformly_continuous_on S f" "bounded S"
shows "bounded(f ` S)"
by (metis (no_types, lifting) assms bounded_closure_image compact_closure compact_continuous_image compact_eq_bounded_closed image_cong uniformly_continuous_imp_continuous uniformly_continuous_on_extension_on_closure)
subsection \<open>With Abstract Topology (TODO: move and remove dependency?)\<close>
lemma openin_contains_ball:
"openin (top_of_set t) s \<longleftrightarrow>
s \<subseteq> t \<and> (\<forall>x \<in> s. \<exists>e. 0 < e \<and> ball x e \<inter> t \<subseteq> s)"
(is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
apply (simp add: openin_open)
apply (metis Int_commute Int_mono inf.cobounded2 open_contains_ball order_refl subsetCE)
done
next
assume ?rhs
then show ?lhs
apply (simp add: openin_euclidean_subtopology_iff)
by (metis (no_types) Int_iff dist_commute inf.absorb_iff2 mem_ball)
qed
lemma openin_contains_cball:
"openin (top_of_set t) s \<longleftrightarrow>
s \<subseteq> t \<and>
(\<forall>x \<in> s. \<exists>e. 0 < e \<and> cball x e \<inter> t \<subseteq> s)"
apply (simp add: openin_contains_ball)
apply (rule iffI)
apply (auto dest!: bspec)
apply (rule_tac x="e/2" in exI, force+)
done
subsection \<open>Closed Nest\<close>
text \<open>Bounded closed nest property (proof does not use Heine-Borel)\<close>
lemma bounded_closed_nest:
fixes S :: "nat \<Rightarrow> ('a::heine_borel) set"
assumes "\<And>n. closed (S n)"
and "\<And>n. S n \<noteq> {}"
and "\<And>m n. m \<le> n \<Longrightarrow> S n \<subseteq> S m"
and "bounded (S 0)"
obtains a where "\<And>n. a \<in> S n"
proof -
from assms(2) obtain x where x: "\<forall>n. x n \<in> S n"
using choice[of "\<lambda>n x. x \<in> S n"] by auto
from assms(4,1) have "seq_compact (S 0)"
by (simp add: bounded_closed_imp_seq_compact)
then obtain l r where lr: "l \<in> S 0" "strict_mono r" "(x \<circ> r) \<longlonglongrightarrow> l"
using x and assms(3) unfolding seq_compact_def by blast
have "\<forall>n. l \<in> S n"
proof
fix n :: nat
have "closed (S n)"
using assms(1) by simp
moreover have "\<forall>i. (x \<circ> r) i \<in> S i"
using x and assms(3) and lr(2) [THEN seq_suble] by auto
then have "\<forall>i. (x \<circ> r) (i + n) \<in> S n"
using assms(3) by (fast intro!: le_add2)
moreover have "(\<lambda>i. (x \<circ> r) (i + n)) \<longlonglongrightarrow> l"
using lr(3) by (rule LIMSEQ_ignore_initial_segment)
ultimately show "l \<in> S n"
by (rule closed_sequentially)
qed
then show ?thesis
using that by blast
qed
text \<open>Decreasing case does not even need compactness, just completeness.\<close>
lemma decreasing_closed_nest:
fixes S :: "nat \<Rightarrow> ('a::complete_space) set"
assumes "\<And>n. closed (S n)"
"\<And>n. S n \<noteq> {}"
"\<And>m n. m \<le> n \<Longrightarrow> S n \<subseteq> S m"
"\<And>e. e>0 \<Longrightarrow> \<exists>n. \<forall>x\<in>S n. \<forall>y\<in>S n. dist x y < e"
obtains a where "\<And>n. a \<in> S n"
proof -
have "\<forall>n. \<exists>x. x \<in> S n"
using assms(2) by auto
then have "\<exists>t. \<forall>n. t n \<in> S n"
using choice[of "\<lambda>n x. x \<in> S n"] by auto
then obtain t where t: "\<forall>n. t n \<in> S n" by auto
{
fix e :: real
assume "e > 0"
then obtain N where N: "\<forall>x\<in>S N. \<forall>y\<in>S N. dist x y < e"
using assms(4) by blast
{
fix m n :: nat
assume "N \<le> m \<and> N \<le> n"
then have "t m \<in> S N" "t n \<in> S N"
using assms(3) t unfolding subset_eq t by blast+
then have "dist (t m) (t n) < e"
using N by auto
}
then have "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (t m) (t n) < e"
by auto
}
then have "Cauchy t"
unfolding cauchy_def by auto
then obtain l where l:"(t \<longlongrightarrow> l) sequentially"
using complete_UNIV unfolding complete_def by auto
{ fix n :: nat
{ fix e :: real
assume "e > 0"
then obtain N :: nat where N: "\<forall>n\<ge>N. dist (t n) l < e"
using l[unfolded lim_sequentially] by auto
have "t (max n N) \<in> S n"
by (meson assms(3) contra_subsetD max.cobounded1 t)
then have "\<exists>y\<in>S n. dist y l < e"
using N max.cobounded2 by blast
}
then have "l \<in> S n"
using closed_approachable[of "S n" l] assms(1) by auto
}
then show ?thesis
using that by blast
qed
text \<open>Strengthen it to the intersection actually being a singleton.\<close>
lemma decreasing_closed_nest_sing:
fixes S :: "nat \<Rightarrow> 'a::complete_space set"
assumes "\<And>n. closed(S n)"
"\<And>n. S n \<noteq> {}"
"\<And>m n. m \<le> n \<Longrightarrow> S n \<subseteq> S m"
"\<And>e. e>0 \<Longrightarrow> \<exists>n. \<forall>x \<in> (S n). \<forall> y\<in>(S n). dist x y < e"
shows "\<exists>a. \<Inter>(range S) = {a}"
proof -
obtain a where a: "\<forall>n. a \<in> S n"
using decreasing_closed_nest[of S] using assms by auto
{ fix b
assume b: "b \<in> \<Inter>(range S)"
{ fix e :: real
assume "e > 0"
then have "dist a b < e"
using assms(4) and b and a by blast
}
then have "dist a b = 0"
by (metis dist_eq_0_iff dist_nz less_le)
}
with a have "\<Inter>(range S) = {a}"
unfolding image_def by auto
then show ?thesis ..
qed
subsection\<^marker>\<open>tag unimportant\<close> \<open>Making a continuous function avoid some value in a neighbourhood\<close>
lemma continuous_within_avoid:
fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space"
assumes "continuous (at x within s) f"
and "f x \<noteq> a"
shows "\<exists>e>0. \<forall>y \<in> s. dist x y < e --> f y \<noteq> a"
proof -
obtain U where "open U" and "f x \<in> U" and "a \<notin> U"
using t1_space [OF \<open>f x \<noteq> a\<close>] by fast
have "(f \<longlongrightarrow> f x) (at x within s)"
using assms(1) by (simp add: continuous_within)
then have "eventually (\<lambda>y. f y \<in> U) (at x within s)"
using \<open>open U\<close> and \<open>f x \<in> U\<close>
unfolding tendsto_def by fast
then have "eventually (\<lambda>y. f y \<noteq> a) (at x within s)"
using \<open>a \<notin> U\<close> by (fast elim: eventually_mono)
then show ?thesis
using \<open>f x \<noteq> a\<close> by (auto simp: dist_commute eventually_at)
qed
lemma continuous_at_avoid:
fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space"
assumes "continuous (at x) f"
and "f x \<noteq> a"
shows "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> f y \<noteq> a"
using assms continuous_within_avoid[of x UNIV f a] by simp
lemma continuous_on_avoid:
fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space"
assumes "continuous_on s f"
and "x \<in> s"
and "f x \<noteq> a"
shows "\<exists>e>0. \<forall>y \<in> s. dist x y < e \<longrightarrow> f y \<noteq> a"
using assms(1)[unfolded continuous_on_eq_continuous_within, THEN bspec[where x=x],
OF assms(2)] continuous_within_avoid[of x s f a]
using assms(3)
by auto
lemma continuous_on_open_avoid:
fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space"
assumes "continuous_on s f"
and "open s"
and "x \<in> s"
and "f x \<noteq> a"
shows "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> f y \<noteq> a"
using assms(1)[unfolded continuous_on_eq_continuous_at[OF assms(2)], THEN bspec[where x=x], OF assms(3)]
using continuous_at_avoid[of x f a] assms(4)
by auto
subsection \<open>Consequences for Real Numbers\<close>
lemma closed_contains_Inf:
fixes S :: "real set"
shows "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> closed S \<Longrightarrow> Inf S \<in> S"
by (metis closure_contains_Inf closure_closed)
lemma closed_subset_contains_Inf:
fixes A C :: "real set"
shows "closed C \<Longrightarrow> A \<subseteq> C \<Longrightarrow> A \<noteq> {} \<Longrightarrow> bdd_below A \<Longrightarrow> Inf A \<in> C"
by (metis closure_contains_Inf closure_minimal subset_eq)
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 closed_subset_contains_Sup:
fixes A C :: "real set"
shows "closed C \<Longrightarrow> A \<subseteq> C \<Longrightarrow> A \<noteq> {} \<Longrightarrow> bdd_above A \<Longrightarrow> Sup A \<in> C"
by (metis closure_contains_Sup closure_minimal subset_eq)
lemma atLeastAtMost_subset_contains_Inf:
fixes A :: "real set" and a b :: real
shows "A \<noteq> {} \<Longrightarrow> a \<le> b \<Longrightarrow> A \<subseteq> {a..b} \<Longrightarrow> Inf A \<in> {a..b}"
by (rule closed_subset_contains_Inf)
(auto intro: closed_real_atLeastAtMost intro!: bdd_belowI[of A a])
lemma bounded_real: "bounded (S::real set) \<longleftrightarrow> (\<exists>a. \<forall>x\<in>S. \<bar>x\<bar> \<le> a)"
by (simp add: bounded_iff)
lemma bounded_imp_bdd_above: "bounded S \<Longrightarrow> bdd_above (S :: real set)"
by (auto simp: bounded_def bdd_above_def dist_real_def)
(metis abs_le_D1 abs_minus_commute diff_le_eq)
lemma bounded_imp_bdd_below: "bounded S \<Longrightarrow> bdd_below (S :: real set)"
by (auto simp: bounded_def bdd_below_def dist_real_def)
(metis abs_le_D1 add.commute diff_le_eq)
lemma bounded_has_Sup:
fixes S :: "real set"
assumes "bounded S"
and "S \<noteq> {}"
shows "\<forall>x\<in>S. x \<le> Sup S"
and "\<forall>b. (\<forall>x\<in>S. x \<le> b) \<longrightarrow> Sup S \<le> b"
proof
show "\<forall>b. (\<forall>x\<in>S. x \<le> b) \<longrightarrow> Sup S \<le> b"
using assms by (metis cSup_least)
qed (metis cSup_upper assms(1) bounded_imp_bdd_above)
lemma Sup_insert:
fixes S :: "real set"
shows "bounded S \<Longrightarrow> Sup (insert x S) = (if S = {} then x else max x (Sup S))"
by (auto simp: bounded_imp_bdd_above sup_max cSup_insert_If)
lemma bounded_has_Inf:
fixes S :: "real set"
assumes "bounded S"
and "S \<noteq> {}"
shows "\<forall>x\<in>S. x \<ge> Inf S"
and "\<forall>b. (\<forall>x\<in>S. x \<ge> b) \<longrightarrow> Inf S \<ge> b"
proof
show "\<forall>b. (\<forall>x\<in>S. x \<ge> b) \<longrightarrow> Inf S \<ge> b"
using assms by (metis cInf_greatest)
qed (metis cInf_lower assms(1) bounded_imp_bdd_below)
lemma Inf_insert:
fixes S :: "real set"
shows "bounded S \<Longrightarrow> Inf (insert x S) = (if S = {} then x else min x (Inf S))"
by (auto simp: bounded_imp_bdd_below inf_min cInf_insert_If)
lemma open_real:
fixes s :: "real set"
shows "open s \<longleftrightarrow> (\<forall>x \<in> s. \<exists>e>0. \<forall>x'. \<bar>x' - x\<bar> < e --> x' \<in> s)"
unfolding open_dist dist_norm by simp
lemma islimpt_approachable_real:
fixes s :: "real set"
shows "x islimpt s \<longleftrightarrow> (\<forall>e>0. \<exists>x'\<in> s. x' \<noteq> x \<and> \<bar>x' - x\<bar> < e)"
unfolding islimpt_approachable dist_norm by simp
lemma closed_real:
fixes s :: "real set"
shows "closed s \<longleftrightarrow> (\<forall>x. (\<forall>e>0. \<exists>x' \<in> s. x' \<noteq> x \<and> \<bar>x' - x\<bar> < e) \<longrightarrow> x \<in> s)"
unfolding closed_limpt islimpt_approachable dist_norm by simp
lemma continuous_at_real_range:
fixes f :: "'a::real_normed_vector \<Rightarrow> real"
shows "continuous (at x) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. \<forall>x'. norm(x' - x) < d --> \<bar>f x' - f x\<bar> < e)"
unfolding continuous_at
unfolding Lim_at
unfolding dist_norm
apply auto
apply (erule_tac x=e in allE, auto)
apply (rule_tac x=d in exI, auto)
apply (erule_tac x=x' in allE, auto)
apply (erule_tac x=e in allE, auto)
done
lemma continuous_on_real_range:
fixes f :: "'a::real_normed_vector \<Rightarrow> real"
shows "continuous_on s f \<longleftrightarrow>
(\<forall>x \<in> s. \<forall>e>0. \<exists>d>0. (\<forall>x' \<in> s. norm(x' - x) < d \<longrightarrow> \<bar>f x' - f x\<bar> < e))"
unfolding continuous_on_iff dist_norm by simp
lemma continuous_on_closed_Collect_le:
fixes f g :: "'a::topological_space \<Rightarrow> real"
assumes f: "continuous_on s f" and g: "continuous_on s g" and s: "closed s"
shows "closed {x \<in> s. f x \<le> g x}"
proof -
have "closed ((\<lambda>x. g x - f x) -` {0..} \<inter> s)"
using closed_real_atLeast continuous_on_diff [OF g f]
by (simp add: continuous_on_closed_vimage [OF s])
also have "((\<lambda>x. g x - f x) -` {0..} \<inter> s) = {x\<in>s. f x \<le> g x}"
by auto
finally show ?thesis .
qed
lemma continuous_le_on_closure:
fixes a::real
assumes f: "continuous_on (closure s) f"
and x: "x \<in> closure(s)"
and xlo: "\<And>x. x \<in> s ==> f(x) \<le> a"
shows "f(x) \<le> a"
using image_closure_subset [OF f, where T=" {x. x \<le> a}" ] assms
continuous_on_closed_Collect_le[of "UNIV" "\<lambda>x. x" "\<lambda>x. a"]
by auto
lemma continuous_ge_on_closure:
fixes a::real
assumes f: "continuous_on (closure s) f"
and x: "x \<in> closure(s)"
and xlo: "\<And>x. x \<in> s ==> f(x) \<ge> a"
shows "f(x) \<ge> a"
using image_closure_subset [OF f, where T=" {x. a \<le> x}"] assms
continuous_on_closed_Collect_le[of "UNIV" "\<lambda>x. a" "\<lambda>x. x"]
by auto
subsection\<open>The infimum of the distance between two sets\<close>
definition\<^marker>\<open>tag important\<close> 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 intro!: le_setdistI)
done
qed
lemma setdist_closure_2 [simp]: "setdist t (closure s) = setdist t s"
by (metis setdist_closure_1 setdist_sym)
lemma setdist_eq_0I: "\<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_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_le_sing: "x \<in> S ==> setdist S T \<le> setdist {x} T"
using setdist_subset_left by auto
lemma infdist_eq_setdist: "infdist x A = setdist {x} A"
by (simp add: infdist_def setdist_def Setcompr_eq_image)
lemma setdist_eq_infdist: "setdist A B = (if A = {} then 0 else INF a\<in>A. infdist a B)"
proof -
have "Inf {dist x y |x y. x \<in> A \<and> y \<in> B} = (INF x\<in>A. Inf (dist x ` B))"
if "b \<in> B" "a \<in> A" for a b
proof (rule order_antisym)
have "Inf {dist x y |x y. x \<in> A \<and> y \<in> B} \<le> Inf (dist x ` B)"
if "b \<in> B" "a \<in> A" "x \<in> A" for x
proof -
have *: "\<And>b'. b' \<in> B \<Longrightarrow> Inf {dist x y |x y. x \<in> A \<and> y \<in> B} \<le> dist x b'"
by (metis (mono_tags, lifting) ex_in_conv setdist_def setdist_le_dist that(3))
show ?thesis
using that by (subst conditionally_complete_lattice_class.le_cInf_iff) (auto simp: *)+
qed
then show "Inf {dist x y |x y. x \<in> A \<and> y \<in> B} \<le> (INF x\<in>A. Inf (dist x ` B))"
using that
by (subst conditionally_complete_lattice_class.le_cInf_iff) (auto simp: bdd_below_def)
next
have *: "\<And>x y. \<lbrakk>b \<in> B; a \<in> A; x \<in> A; y \<in> B\<rbrakk> \<Longrightarrow> \<exists>a\<in>A. Inf (dist a ` B) \<le> dist x y"
by (meson bdd_below_image_dist cINF_lower)
show "(INF x\<in>A. Inf (dist x ` B)) \<le> Inf {dist x y |x y. x \<in> A \<and> y \<in> B}"
proof (rule conditionally_complete_lattice_class.cInf_mono)
show "bdd_below ((\<lambda>x. Inf (dist x ` B)) ` A)"
by (metis (no_types, lifting) bdd_belowI2 ex_in_conv infdist_def infdist_nonneg that(1))
qed (use that in \<open>auto simp: *\<close>)
qed
then show ?thesis
by (auto simp: setdist_def infdist_def)
qed
lemma infdist_mono:
assumes "A \<subseteq> B" "A \<noteq> {}"
shows "infdist x B \<le> infdist x A"
by (simp add: assms infdist_eq_setdist setdist_subset_right)
lemma infdist_singleton [simp]:
"infdist x {y} = dist x y"
by (simp add: infdist_eq_setdist)
proposition setdist_attains_inf:
assumes "compact B" "B \<noteq> {}"
obtains y where "y \<in> B" "setdist A B = infdist y A"
proof (cases "A = {}")
case True
then show thesis
by (metis assms diameter_compact_attained infdist_def setdist_def that)
next
case False
obtain y where "y \<in> B" and min: "\<And>y'. y' \<in> B \<Longrightarrow> infdist y A \<le> infdist y' A"
by (metis continuous_attains_inf [OF assms continuous_on_infdist] continuous_on_id)
show thesis
proof
have "setdist A B = (INF y\<in>B. infdist y A)"
by (metis \<open>B \<noteq> {}\<close> setdist_eq_infdist setdist_sym)
also have "\<dots> = infdist y A"
proof (rule order_antisym)
show "(INF y\<in>B. infdist y A) \<le> infdist y A"
proof (rule cInf_lower)
show "infdist y A \<in> (\<lambda>y. infdist y A) ` B"
using \<open>y \<in> B\<close> by blast
show "bdd_below ((\<lambda>y. infdist y A) ` B)"
by (meson bdd_belowI2 infdist_nonneg)
qed
next
show "infdist y A \<le> (INF y\<in>B. infdist y A)"
by (simp add: \<open>B \<noteq> {}\<close> cINF_greatest min)
qed
finally show "setdist A B = infdist y A" .
qed (fact \<open>y \<in> B\<close>)
qed
end
diff --git a/src/HOL/Analysis/Elementary_Normed_Spaces.thy b/src/HOL/Analysis/Elementary_Normed_Spaces.thy
--- a/src/HOL/Analysis/Elementary_Normed_Spaces.thy
+++ b/src/HOL/Analysis/Elementary_Normed_Spaces.thy
@@ -1,1794 +1,1794 @@
(* Author: L C Paulson, University of Cambridge
Author: Amine Chaieb, University of Cambridge
Author: Robert Himmelmann, TU Muenchen
Author: Brian Huffman, Portland State University
*)
section \<open>Elementary Normed Vector Spaces\<close>
theory Elementary_Normed_Spaces
imports
"HOL-Library.FuncSet"
Elementary_Metric_Spaces Cartesian_Space
Connected
begin
subsection \<open>Orthogonal Transformation of Balls\<close>
subsection\<^marker>\<open>tag unimportant\<close> \<open>Various Lemmas Combining Imports\<close>
lemma open_sums:
fixes T :: "('b::real_normed_vector) set"
assumes "open S \<or> open T"
shows "open (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})"
using assms
proof
assume S: "open S"
show ?thesis
proof (clarsimp simp: open_dist)
fix x y
assume "x \<in> S" "y \<in> T"
with S obtain e where "e > 0" and e: "\<And>x'. dist x' x < e \<Longrightarrow> x' \<in> S"
by (auto simp: open_dist)
then have "\<And>z. dist z (x + y) < e \<Longrightarrow> \<exists>x\<in>S. \<exists>y\<in>T. z = x + y"
by (metis \<open>y \<in> T\<close> diff_add_cancel dist_add_cancel2)
then show "\<exists>e>0. \<forall>z. dist z (x + y) < e \<longrightarrow> (\<exists>x\<in>S. \<exists>y\<in>T. z = x + y)"
using \<open>0 < e\<close> \<open>x \<in> S\<close> by blast
qed
next
assume T: "open T"
show ?thesis
proof (clarsimp simp: open_dist)
fix x y
assume "x \<in> S" "y \<in> T"
with T obtain e where "e > 0" and e: "\<And>x'. dist x' y < e \<Longrightarrow> x' \<in> T"
by (auto simp: open_dist)
then have "\<And>z. dist z (x + y) < e \<Longrightarrow> \<exists>x\<in>S. \<exists>y\<in>T. z = x + y"
by (metis \<open>x \<in> S\<close> add_diff_cancel_left' add_diff_eq diff_diff_add dist_norm)
then show "\<exists>e>0. \<forall>z. dist z (x + y) < e \<longrightarrow> (\<exists>x\<in>S. \<exists>y\<in>T. z = x + y)"
using \<open>0 < e\<close> \<open>y \<in> T\<close> by blast
qed
qed
lemma image_orthogonal_transformation_ball:
fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
assumes "orthogonal_transformation f"
shows "f ` ball x r = ball (f x) r"
proof (intro equalityI subsetI)
fix y assume "y \<in> f ` ball x r"
with assms show "y \<in> ball (f x) r"
by (auto simp: orthogonal_transformation_isometry)
next
fix y assume y: "y \<in> ball (f x) r"
then obtain z where z: "y = f z"
using assms orthogonal_transformation_surj by blast
with y assms show "y \<in> f ` ball x r"
by (auto simp: orthogonal_transformation_isometry)
qed
lemma image_orthogonal_transformation_cball:
fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
assumes "orthogonal_transformation f"
shows "f ` cball x r = cball (f x) r"
proof (intro equalityI subsetI)
fix y assume "y \<in> f ` cball x r"
with assms show "y \<in> cball (f x) r"
by (auto simp: orthogonal_transformation_isometry)
next
fix y assume y: "y \<in> cball (f x) r"
then obtain z where z: "y = f z"
using assms orthogonal_transformation_surj by blast
with y assms show "y \<in> f ` cball x r"
by (auto simp: orthogonal_transformation_isometry)
qed
subsection \<open>Support\<close>
definition (in monoid_add) support_on :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'b set"
where "support_on s f = {x\<in>s. f x \<noteq> 0}"
lemma in_support_on: "x \<in> support_on s f \<longleftrightarrow> x \<in> s \<and> f x \<noteq> 0"
by (simp add: support_on_def)
lemma support_on_simps[simp]:
"support_on {} f = {}"
"support_on (insert x s) f =
(if f x = 0 then support_on s f else insert x (support_on s f))"
"support_on (s \<union> t) f = support_on s f \<union> support_on t f"
"support_on (s \<inter> t) f = support_on s f \<inter> support_on t f"
"support_on (s - t) f = support_on s f - support_on t f"
"support_on (f ` s) g = f ` (support_on s (g \<circ> f))"
unfolding support_on_def by auto
lemma support_on_cong:
"(\<And>x. x \<in> s \<Longrightarrow> f x = 0 \<longleftrightarrow> g x = 0) \<Longrightarrow> support_on s f = support_on s g"
by (auto simp: support_on_def)
lemma support_on_if: "a \<noteq> 0 \<Longrightarrow> support_on A (\<lambda>x. if P x then a else 0) = {x\<in>A. P x}"
by (auto simp: support_on_def)
lemma support_on_if_subset: "support_on A (\<lambda>x. if P x then a else 0) \<subseteq> {x \<in> A. P x}"
by (auto simp: support_on_def)
lemma finite_support[intro]: "finite S \<Longrightarrow> finite (support_on S f)"
unfolding support_on_def by auto
(* TODO: is supp_sum really needed? TODO: Generalize to Finite_Set.fold *)
definition (in comm_monoid_add) supp_sum :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b set \<Rightarrow> 'a"
where "supp_sum f S = (\<Sum>x\<in>support_on S f. f x)"
lemma supp_sum_empty[simp]: "supp_sum f {} = 0"
unfolding supp_sum_def by auto
lemma supp_sum_insert[simp]:
"finite (support_on S f) \<Longrightarrow>
supp_sum f (insert x S) = (if x \<in> S then supp_sum f S else f x + supp_sum f S)"
by (simp add: supp_sum_def in_support_on insert_absorb)
lemma supp_sum_divide_distrib: "supp_sum f A / (r::'a::field) = supp_sum (\<lambda>n. f n / r) A"
by (cases "r = 0")
(auto simp: supp_sum_def sum_divide_distrib intro!: sum.cong support_on_cong)
subsection \<open>Intervals\<close>
lemma image_affinity_interval:
fixes c :: "'a::ordered_real_vector"
shows "((\<lambda>x. m *\<^sub>R x + c) ` {a..b}) =
(if {a..b}={} then {}
else if 0 \<le> m then {m *\<^sub>R a + c .. m *\<^sub>R b + c}
else {m *\<^sub>R b + c .. m *\<^sub>R a + c})"
(is "?lhs = ?rhs")
proof (cases "m=0")
case True
then show ?thesis
by force
next
case False
show ?thesis
proof
show "?lhs \<subseteq> ?rhs"
by (auto simp: scaleR_left_mono scaleR_left_mono_neg)
show "?rhs \<subseteq> ?lhs"
proof (clarsimp, intro conjI impI subsetI)
show "\<lbrakk>0 \<le> m; a \<le> b; x \<in> {m *\<^sub>R a + c..m *\<^sub>R b + c}\<rbrakk>
\<Longrightarrow> x \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}" for x
apply (rule_tac x="inverse m *\<^sub>R (x-c)" in rev_image_eqI)
using False apply (auto simp: pos_le_divideR_eq pos_divideR_le_eq le_diff_eq diff_le_eq)
done
show "\<lbrakk>\<not> 0 \<le> m; a \<le> b; x \<in> {m *\<^sub>R b + c..m *\<^sub>R a + c}\<rbrakk>
\<Longrightarrow> x \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}" for x
apply (rule_tac x="inverse m *\<^sub>R (x-c)" in rev_image_eqI)
apply (auto simp add: neg_le_divideR_eq neg_divideR_le_eq not_le le_diff_eq diff_le_eq)
done
qed
qed
qed
subsection \<open>Limit Points\<close>
lemma islimpt_ball:
fixes x y :: "'a::{real_normed_vector,perfect_space}"
shows "y islimpt ball x e \<longleftrightarrow> 0 < e \<and> y \<in> cball x e"
(is "?lhs \<longleftrightarrow> ?rhs")
proof
show ?rhs if ?lhs
proof
{
assume "e \<le> 0"
then have *: "ball x e = {}"
using ball_eq_empty[of x e] by auto
have False using \<open>?lhs\<close>
unfolding * using islimpt_EMPTY[of y] by auto
}
then show "e > 0" by (metis not_less)
show "y \<in> cball x e"
using closed_cball[of x e] islimpt_subset[of y "ball x e" "cball x e"]
ball_subset_cball[of x e] \<open>?lhs\<close>
unfolding closed_limpt by auto
qed
show ?lhs if ?rhs
proof -
from that have "e > 0" by auto
{
fix d :: real
assume "d > 0"
have "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
proof (cases "d \<le> dist x y")
case True
then show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
proof (cases "x = y")
case True
then have False
using \<open>d \<le> dist x y\<close> \<open>d>0\<close> by auto
then show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
by auto
next
case False
have "dist x (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) =
norm (x - y + (d / (2 * norm (y - x))) *\<^sub>R (y - x))"
unfolding mem_cball mem_ball dist_norm diff_diff_eq2 diff_add_eq[symmetric]
by auto
also have "\<dots> = \<bar>- 1 + d / (2 * norm (x - y))\<bar> * norm (x - y)"
using scaleR_left_distrib[of "- 1" "d / (2 * norm (y - x))", symmetric, of "y - x"]
unfolding scaleR_minus_left scaleR_one
by (auto simp: norm_minus_commute)
also have "\<dots> = \<bar>- norm (x - y) + d / 2\<bar>"
unfolding abs_mult_pos[of "norm (x - y)", OF norm_ge_zero[of "x - y"]]
unfolding distrib_right using \<open>x\<noteq>y\<close> by auto
also have "\<dots> \<le> e - d/2" using \<open>d \<le> dist x y\<close> and \<open>d>0\<close> and \<open>?rhs\<close>
by (auto simp: dist_norm)
finally have "y - (d / (2 * dist y x)) *\<^sub>R (y - x) \<in> ball x e" using \<open>d>0\<close>
by auto
moreover
have "(d / (2*dist y x)) *\<^sub>R (y - x) \<noteq> 0"
using \<open>x\<noteq>y\<close>[unfolded dist_nz] \<open>d>0\<close> unfolding scaleR_eq_0_iff
by (auto simp: dist_commute)
moreover
have "dist (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) y < d"
unfolding dist_norm
apply simp
unfolding norm_minus_cancel
using \<open>d > 0\<close> \<open>x\<noteq>y\<close>[unfolded dist_nz] dist_commute[of x y]
unfolding dist_norm
apply auto
done
ultimately show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
apply (rule_tac x = "y - (d / (2*dist y x)) *\<^sub>R (y - x)" in bexI)
apply auto
done
qed
next
case False
then have "d > dist x y" by auto
show "\<exists>x' \<in> ball x e. x' \<noteq> y \<and> dist x' y < d"
proof (cases "x = y")
case True
obtain z where **: "z \<noteq> y" "dist z y < min e d"
using perfect_choose_dist[of "min e d" y]
using \<open>d > 0\<close> \<open>e>0\<close> by auto
show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
unfolding \<open>x = y\<close>
using \<open>z \<noteq> y\<close> **
apply (rule_tac x=z in bexI)
apply (auto simp: dist_commute)
done
next
case False
then show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
using \<open>d>0\<close> \<open>d > dist x y\<close> \<open>?rhs\<close>
apply (rule_tac x=x in bexI, auto)
done
qed
qed
}
then show ?thesis
unfolding mem_cball islimpt_approachable mem_ball by auto
qed
qed
lemma closure_ball_lemma:
fixes x y :: "'a::real_normed_vector"
assumes "x \<noteq> y"
shows "y islimpt ball x (dist x y)"
proof (rule islimptI)
fix T
assume "y \<in> T" "open T"
then obtain r where "0 < r" "\<forall>z. dist z y < r \<longrightarrow> z \<in> T"
unfolding open_dist by fast
(* choose point between x and y, within distance r of y. *)
define k where "k = min 1 (r / (2 * dist x y))"
define z where "z = y + scaleR k (x - y)"
have z_def2: "z = x + scaleR (1 - k) (y - x)"
unfolding z_def by (simp add: algebra_simps)
have "dist z y < r"
unfolding z_def k_def using \<open>0 < r\<close>
by (simp add: dist_norm min_def)
then have "z \<in> T"
using \<open>\<forall>z. dist z y < r \<longrightarrow> z \<in> T\<close> by simp
have "dist x z < dist x y"
unfolding z_def2 dist_norm
apply (simp add: norm_minus_commute)
apply (simp only: dist_norm [symmetric])
apply (subgoal_tac "\<bar>1 - k\<bar> * dist x y < 1 * dist x y", simp)
apply (rule mult_strict_right_mono)
apply (simp add: k_def \<open>0 < r\<close> \<open>x \<noteq> y\<close>)
apply (simp add: \<open>x \<noteq> y\<close>)
done
then have "z \<in> ball x (dist x y)"
by simp
have "z \<noteq> y"
unfolding z_def k_def using \<open>x \<noteq> y\<close> \<open>0 < r\<close>
by (simp add: min_def)
show "\<exists>z\<in>ball x (dist x y). z \<in> T \<and> z \<noteq> y"
using \<open>z \<in> ball x (dist x y)\<close> \<open>z \<in> T\<close> \<open>z \<noteq> y\<close>
by fast
qed
subsection \<open>Balls and Spheres in Normed Spaces\<close>
lemma mem_ball_0 [simp]: "x \<in> ball 0 e \<longleftrightarrow> norm x < e"
for x :: "'a::real_normed_vector"
by (simp add: dist_norm)
lemma mem_cball_0 [simp]: "x \<in> cball 0 e \<longleftrightarrow> norm x \<le> e"
for x :: "'a::real_normed_vector"
by (simp add: dist_norm)
lemma closure_ball [simp]:
fixes x :: "'a::real_normed_vector"
shows "0 < e \<Longrightarrow> closure (ball x e) = cball x e"
apply (rule equalityI)
apply (rule closure_minimal)
apply (rule ball_subset_cball)
apply (rule closed_cball)
apply (rule subsetI, rename_tac y)
apply (simp add: le_less [where 'a=real])
apply (erule disjE)
apply (rule subsetD [OF closure_subset], simp)
apply (simp add: closure_def, clarify)
apply (rule closure_ball_lemma)
apply simp
done
lemma mem_sphere_0 [simp]: "x \<in> sphere 0 e \<longleftrightarrow> norm x = e"
for x :: "'a::real_normed_vector"
by (simp add: dist_norm)
(* In a trivial vector space, this fails for e = 0. *)
lemma interior_cball [simp]:
fixes x :: "'a::{real_normed_vector, perfect_space}"
shows "interior (cball x e) = ball x e"
proof (cases "e \<ge> 0")
case False note cs = this
from cs have null: "ball x e = {}"
using ball_empty[of e x] by auto
moreover
have "cball x e = {}"
proof (rule equals0I)
fix y
assume "y \<in> cball x e"
then show False
by (metis ball_eq_empty null cs dist_eq_0_iff dist_le_zero_iff empty_subsetI mem_cball
subset_antisym subset_ball)
qed
then have "interior (cball x e) = {}"
using interior_empty by auto
ultimately show ?thesis by blast
next
case True note cs = this
have "ball x e \<subseteq> cball x e"
using ball_subset_cball by auto
moreover
{
fix S y
assume as: "S \<subseteq> cball x e" "open S" "y\<in>S"
then obtain d where "d>0" and d: "\<forall>x'. dist x' y < d \<longrightarrow> x' \<in> S"
unfolding open_dist by blast
then obtain xa where xa_y: "xa \<noteq> y" and xa: "dist xa y < d"
using perfect_choose_dist [of d] by auto
have "xa \<in> S"
using d[THEN spec[where x = xa]]
using xa by (auto simp: dist_commute)
then have xa_cball: "xa \<in> cball x e"
using as(1) by auto
then have "y \<in> ball x e"
proof (cases "x = y")
case True
then have "e > 0" using cs order.order_iff_strict xa_cball xa_y by fastforce
then show "y \<in> ball x e"
using \<open>x = y \<close> by simp
next
case False
have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) y < d"
unfolding dist_norm
using \<open>d>0\<close> norm_ge_zero[of "y - x"] \<open>x \<noteq> y\<close> by auto
then have *: "y + (d / 2 / dist y x) *\<^sub>R (y - x) \<in> cball x e"
using d as(1)[unfolded subset_eq] by blast
have "y - x \<noteq> 0" using \<open>x \<noteq> y\<close> by auto
hence **:"d / (2 * norm (y - x)) > 0"
unfolding zero_less_norm_iff[symmetric] using \<open>d>0\<close> by auto
have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) x =
norm (y + (d / (2 * norm (y - x))) *\<^sub>R y - (d / (2 * norm (y - x))) *\<^sub>R x - x)"
by (auto simp: dist_norm algebra_simps)
also have "\<dots> = norm ((1 + d / (2 * norm (y - x))) *\<^sub>R (y - x))"
by (auto simp: algebra_simps)
also have "\<dots> = \<bar>1 + d / (2 * norm (y - x))\<bar> * norm (y - x)"
using ** by auto
also have "\<dots> = (dist y x) + d/2"
using ** by (auto simp: distrib_right dist_norm)
finally have "e \<ge> dist x y +d/2"
using *[unfolded mem_cball] by (auto simp: dist_commute)
then show "y \<in> ball x e"
unfolding mem_ball using \<open>d>0\<close> by auto
qed
}
then have "\<forall>S \<subseteq> cball x e. open S \<longrightarrow> S \<subseteq> ball x e"
by auto
ultimately show ?thesis
using interior_unique[of "ball x e" "cball x e"]
using open_ball[of x e]
by auto
qed
lemma frontier_ball [simp]:
fixes a :: "'a::real_normed_vector"
shows "0 < e \<Longrightarrow> frontier (ball a e) = sphere a e"
by (force simp: frontier_def)
lemma frontier_cball [simp]:
fixes a :: "'a::{real_normed_vector, perfect_space}"
shows "frontier (cball a e) = sphere a e"
by (force simp: frontier_def)
corollary compact_sphere [simp]:
fixes a :: "'a::{real_normed_vector,perfect_space,heine_borel}"
shows "compact (sphere a r)"
using compact_frontier [of "cball a r"] by simp
corollary bounded_sphere [simp]:
fixes a :: "'a::{real_normed_vector,perfect_space,heine_borel}"
shows "bounded (sphere a r)"
by (simp add: compact_imp_bounded)
corollary closed_sphere [simp]:
fixes a :: "'a::{real_normed_vector,perfect_space,heine_borel}"
shows "closed (sphere a r)"
by (simp add: compact_imp_closed)
lemma image_add_ball [simp]:
fixes a :: "'a::real_normed_vector"
shows "(+) b ` ball a r = ball (a+b) r"
apply (intro equalityI subsetI)
apply (force simp: dist_norm)
apply (rule_tac x="x-b" in image_eqI)
apply (auto simp: dist_norm algebra_simps)
done
lemma image_add_cball [simp]:
fixes a :: "'a::real_normed_vector"
shows "(+) b ` cball a r = cball (a+b) r"
apply (intro equalityI subsetI)
apply (force simp: dist_norm)
apply (rule_tac x="x-b" in image_eqI)
apply (auto simp: dist_norm algebra_simps)
done
subsection\<^marker>\<open>tag unimportant\<close> \<open>Various Lemmas on Normed Algebras\<close>
lemma closed_of_nat_image: "closed (of_nat ` A :: 'a::real_normed_algebra_1 set)"
by (rule discrete_imp_closed[of 1]) (auto simp: dist_of_nat)
lemma closed_of_int_image: "closed (of_int ` A :: 'a::real_normed_algebra_1 set)"
by (rule discrete_imp_closed[of 1]) (auto simp: dist_of_int)
lemma closed_Nats [simp]: "closed (\<nat> :: 'a :: real_normed_algebra_1 set)"
unfolding Nats_def by (rule closed_of_nat_image)
lemma closed_Ints [simp]: "closed (\<int> :: 'a :: real_normed_algebra_1 set)"
unfolding Ints_def by (rule closed_of_int_image)
lemma closed_subset_Ints:
fixes A :: "'a :: real_normed_algebra_1 set"
assumes "A \<subseteq> \<int>"
shows "closed A"
proof (intro discrete_imp_closed[OF zero_less_one] ballI impI, goal_cases)
case (1 x y)
with assms have "x \<in> \<int>" and "y \<in> \<int>" by auto
with \<open>dist y x < 1\<close> show "y = x"
by (auto elim!: Ints_cases simp: dist_of_int)
qed
subsection \<open>Filters\<close>
definition indirection :: "'a::real_normed_vector \<Rightarrow> 'a \<Rightarrow> 'a filter" (infixr "indirection" 70)
where "a indirection v = at a within {b. \<exists>c\<ge>0. b - a = scaleR c v}"
subsection \<open>Trivial Limits\<close>
lemma trivial_limit_at_infinity:
"\<not> trivial_limit (at_infinity :: ('a::{real_normed_vector,perfect_space}) filter)"
unfolding trivial_limit_def eventually_at_infinity
apply clarsimp
apply (subgoal_tac "\<exists>x::'a. x \<noteq> 0", clarify)
apply (rule_tac x="scaleR (b / norm x) x" in exI, simp)
apply (cut_tac islimpt_UNIV [of "0::'a", unfolded islimpt_def])
apply (drule_tac x=UNIV in spec, simp)
done
lemma at_within_ball_bot_iff:
fixes x y :: "'a::{real_normed_vector,perfect_space}"
shows "at x within ball y r = bot \<longleftrightarrow> (r=0 \<or> x \<notin> cball y r)"
unfolding trivial_limit_within
apply (auto simp add:trivial_limit_within islimpt_ball )
by (metis le_less_trans less_eq_real_def zero_le_dist)
subsection \<open>Limits\<close>
proposition Lim_at_infinity: "(f \<longlongrightarrow> l) at_infinity \<longleftrightarrow> (\<forall>e>0. \<exists>b. \<forall>x. norm x \<ge> b \<longrightarrow> dist (f x) l < e)"
by (auto simp: tendsto_iff eventually_at_infinity)
corollary Lim_at_infinityI [intro?]:
assumes "\<And>e. e > 0 \<Longrightarrow> \<exists>B. \<forall>x. norm x \<ge> B \<longrightarrow> dist (f x) l \<le> e"
shows "(f \<longlongrightarrow> l) at_infinity"
apply (simp add: Lim_at_infinity, clarify)
apply (rule ex_forward [OF assms [OF half_gt_zero]], auto)
done
lemma Lim_transform_within_set_eq:
fixes a :: "'a::metric_space" and l :: "'b::metric_space"
shows "eventually (\<lambda>x. x \<in> S \<longleftrightarrow> x \<in> T) (at a)
\<Longrightarrow> ((f \<longlongrightarrow> l) (at a within S) \<longleftrightarrow> (f \<longlongrightarrow> l) (at a within T))"
by (force intro: Lim_transform_within_set elim: eventually_mono)
lemma Lim_null:
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
shows "(f \<longlongrightarrow> l) net \<longleftrightarrow> ((\<lambda>x. f(x) - l) \<longlongrightarrow> 0) net"
by (simp add: Lim dist_norm)
lemma Lim_null_comparison:
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
assumes "eventually (\<lambda>x. norm (f x) \<le> g x) net" "(g \<longlongrightarrow> 0) net"
shows "(f \<longlongrightarrow> 0) net"
using assms(2)
proof (rule metric_tendsto_imp_tendsto)
show "eventually (\<lambda>x. dist (f x) 0 \<le> dist (g x) 0) net"
using assms(1) by (rule eventually_mono) (simp add: dist_norm)
qed
lemma Lim_transform_bound:
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
and g :: "'a \<Rightarrow> 'c::real_normed_vector"
assumes "eventually (\<lambda>n. norm (f n) \<le> norm (g n)) net"
and "(g \<longlongrightarrow> 0) net"
shows "(f \<longlongrightarrow> 0) net"
using assms(1) tendsto_norm_zero [OF assms(2)]
by (rule Lim_null_comparison)
lemma lim_null_mult_right_bounded:
fixes f :: "'a \<Rightarrow> 'b::real_normed_div_algebra"
assumes f: "(f \<longlongrightarrow> 0) F" and g: "eventually (\<lambda>x. norm(g x) \<le> B) F"
shows "((\<lambda>z. f z * g z) \<longlongrightarrow> 0) F"
proof -
have *: "((\<lambda>x. norm (f x) * B) \<longlongrightarrow> 0) F"
by (simp add: f tendsto_mult_left_zero tendsto_norm_zero)
have "((\<lambda>x. norm (f x) * norm (g x)) \<longlongrightarrow> 0) F"
apply (rule Lim_null_comparison [OF _ *])
apply (simp add: eventually_mono [OF g] mult_left_mono)
done
then show ?thesis
by (subst tendsto_norm_zero_iff [symmetric]) (simp add: norm_mult)
qed
lemma lim_null_mult_left_bounded:
fixes f :: "'a \<Rightarrow> 'b::real_normed_div_algebra"
assumes g: "eventually (\<lambda>x. norm(g x) \<le> B) F" and f: "(f \<longlongrightarrow> 0) F"
shows "((\<lambda>z. g z * f z) \<longlongrightarrow> 0) F"
proof -
have *: "((\<lambda>x. B * norm (f x)) \<longlongrightarrow> 0) F"
by (simp add: f tendsto_mult_right_zero tendsto_norm_zero)
have "((\<lambda>x. norm (g x) * norm (f x)) \<longlongrightarrow> 0) F"
apply (rule Lim_null_comparison [OF _ *])
apply (simp add: eventually_mono [OF g] mult_right_mono)
done
then show ?thesis
by (subst tendsto_norm_zero_iff [symmetric]) (simp add: norm_mult)
qed
lemma lim_null_scaleR_bounded:
assumes f: "(f \<longlongrightarrow> 0) net" and gB: "eventually (\<lambda>a. f a = 0 \<or> norm(g a) \<le> B) net"
shows "((\<lambda>n. f n *\<^sub>R g n) \<longlongrightarrow> 0) net"
proof
fix \<epsilon>::real
assume "0 < \<epsilon>"
then have B: "0 < \<epsilon> / (abs B + 1)" by simp
have *: "\<bar>f x\<bar> * norm (g x) < \<epsilon>" if f: "\<bar>f x\<bar> * (\<bar>B\<bar> + 1) < \<epsilon>" and g: "norm (g x) \<le> B" for x
proof -
have "\<bar>f x\<bar> * norm (g x) \<le> \<bar>f x\<bar> * B"
by (simp add: mult_left_mono g)
also have "\<dots> \<le> \<bar>f x\<bar> * (\<bar>B\<bar> + 1)"
by (simp add: mult_left_mono)
also have "\<dots> < \<epsilon>"
by (rule f)
finally show ?thesis .
qed
show "\<forall>\<^sub>F x in net. dist (f x *\<^sub>R g x) 0 < \<epsilon>"
apply (rule eventually_mono [OF eventually_conj [OF tendstoD [OF f B] gB] ])
apply (auto simp: \<open>0 < \<epsilon>\<close> field_split_simps * split: if_split_asm)
done
qed
lemma Lim_norm_ubound:
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
assumes "\<not>(trivial_limit net)" "(f \<longlongrightarrow> l) net" "eventually (\<lambda>x. norm(f x) \<le> e) net"
shows "norm(l) \<le> e"
using assms by (fast intro: tendsto_le tendsto_intros)
lemma Lim_norm_lbound:
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
assumes "\<not> trivial_limit net"
and "(f \<longlongrightarrow> l) net"
and "eventually (\<lambda>x. e \<le> norm (f x)) net"
shows "e \<le> norm l"
using assms by (fast intro: tendsto_le tendsto_intros)
text\<open>Limit under bilinear function\<close>
lemma Lim_bilinear:
assumes "(f \<longlongrightarrow> l) net"
and "(g \<longlongrightarrow> m) net"
and "bounded_bilinear h"
shows "((\<lambda>x. h (f x) (g x)) \<longlongrightarrow> (h l m)) net"
using \<open>bounded_bilinear h\<close> \<open>(f \<longlongrightarrow> l) net\<close> \<open>(g \<longlongrightarrow> m) net\<close>
by (rule bounded_bilinear.tendsto)
lemma Lim_at_zero:
fixes a :: "'a::real_normed_vector"
and l :: "'b::topological_space"
shows "(f \<longlongrightarrow> l) (at a) \<longleftrightarrow> ((\<lambda>x. f(a + x)) \<longlongrightarrow> l) (at 0)"
using LIM_offset_zero LIM_offset_zero_cancel ..
subsection\<^marker>\<open>tag unimportant\<close> \<open>Limit Point of Filter\<close>
lemma netlimit_at_vector:
fixes a :: "'a::real_normed_vector"
shows "netlimit (at a) = a"
proof (cases "\<exists>x. x \<noteq> a")
case True then obtain x where x: "x \<noteq> a" ..
have "\<not> trivial_limit (at a)"
unfolding trivial_limit_def eventually_at dist_norm
apply clarsimp
apply (rule_tac x="a + scaleR (d / 2) (sgn (x - a))" in exI)
apply (simp add: norm_sgn sgn_zero_iff x)
done
then show ?thesis
by (rule Lim_ident_at [of a UNIV])
qed simp
subsection \<open>Boundedness\<close>
lemma continuous_on_closure_norm_le:
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
assumes "continuous_on (closure s) f"
and "\<forall>y \<in> s. norm(f y) \<le> b"
and "x \<in> (closure s)"
shows "norm (f x) \<le> b"
proof -
have *: "f ` s \<subseteq> cball 0 b"
using assms(2)[unfolded mem_cball_0[symmetric]] by auto
show ?thesis
by (meson "*" assms(1) assms(3) closed_cball image_closure_subset image_subset_iff mem_cball_0)
qed
lemma bounded_pos: "bounded S \<longleftrightarrow> (\<exists>b>0. \<forall>x\<in> S. norm x \<le> b)"
unfolding bounded_iff
by (meson less_imp_le not_le order_trans zero_less_one)
lemma bounded_pos_less: "bounded S \<longleftrightarrow> (\<exists>b>0. \<forall>x\<in> S. norm x < b)"
apply (simp add: bounded_pos)
apply (safe; rule_tac x="b+1" in exI; force)
done
lemma Bseq_eq_bounded:
fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
shows "Bseq f \<longleftrightarrow> bounded (range f)"
unfolding Bseq_def bounded_pos by auto
lemma bounded_linear_image:
assumes "bounded S"
and "bounded_linear f"
shows "bounded (f ` S)"
proof -
from assms(1) obtain b where "b > 0" and b: "\<forall>x\<in>S. norm x \<le> b"
unfolding bounded_pos by auto
from assms(2) obtain B where B: "B > 0" "\<forall>x. norm (f x) \<le> B * norm x"
using bounded_linear.pos_bounded by (auto simp: ac_simps)
show ?thesis
unfolding bounded_pos
proof (intro exI, safe)
show "norm (f x) \<le> B * b" if "x \<in> S" for x
by (meson B b less_imp_le mult_left_mono order_trans that)
qed (use \<open>b > 0\<close> \<open>B > 0\<close> in auto)
qed
lemma bounded_scaling:
fixes S :: "'a::real_normed_vector set"
shows "bounded S \<Longrightarrow> bounded ((\<lambda>x. c *\<^sub>R x) ` S)"
apply (rule bounded_linear_image, assumption)
apply (rule bounded_linear_scaleR_right)
done
lemma bounded_scaleR_comp:
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
assumes "bounded (f ` S)"
shows "bounded ((\<lambda>x. r *\<^sub>R f x) ` S)"
using bounded_scaling[of "f ` S" r] assms
by (auto simp: image_image)
lemma bounded_translation:
fixes S :: "'a::real_normed_vector set"
assumes "bounded S"
shows "bounded ((\<lambda>x. a + x) ` S)"
proof -
from assms obtain b where b: "b > 0" "\<forall>x\<in>S. norm x \<le> b"
unfolding bounded_pos by auto
{
fix x
assume "x \<in> S"
then have "norm (a + x) \<le> b + norm a"
using norm_triangle_ineq[of a x] b by auto
}
then show ?thesis
unfolding bounded_pos
using norm_ge_zero[of a] b(1) and add_strict_increasing[of b 0 "norm a"]
by (auto intro!: exI[of _ "b + norm a"])
qed
lemma bounded_translation_minus:
fixes S :: "'a::real_normed_vector set"
shows "bounded S \<Longrightarrow> bounded ((\<lambda>x. x - a) ` S)"
using bounded_translation [of S "-a"] by simp
lemma bounded_uminus [simp]:
fixes X :: "'a::real_normed_vector set"
shows "bounded (uminus ` X) \<longleftrightarrow> bounded X"
by (auto simp: bounded_def dist_norm; rule_tac x="-x" in exI; force simp: add.commute norm_minus_commute)
lemma uminus_bounded_comp [simp]:
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
shows "bounded ((\<lambda>x. - f x) ` S) \<longleftrightarrow> bounded (f ` S)"
using bounded_uminus[of "f ` S"]
by (auto simp: image_image)
lemma bounded_plus_comp:
fixes f g::"'a \<Rightarrow> 'b::real_normed_vector"
assumes "bounded (f ` S)"
assumes "bounded (g ` S)"
shows "bounded ((\<lambda>x. f x + g x) ` S)"
proof -
{
fix B C
assume "\<And>x. x\<in>S \<Longrightarrow> norm (f x) \<le> B" "\<And>x. x\<in>S \<Longrightarrow> norm (g x) \<le> C"
then have "\<And>x. x \<in> S \<Longrightarrow> norm (f x + g x) \<le> B + C"
by (auto intro!: norm_triangle_le add_mono)
} then show ?thesis
using assms by (fastforce simp: bounded_iff)
qed
lemma bounded_plus:
fixes S ::"'a::real_normed_vector set"
assumes "bounded S" "bounded T"
shows "bounded ((\<lambda>(x,y). x + y) ` (S \<times> T))"
using bounded_plus_comp [of fst "S \<times> T" snd] assms
by (auto simp: split_def split: if_split_asm)
lemma bounded_minus_comp:
"bounded (f ` S) \<Longrightarrow> bounded (g ` S) \<Longrightarrow> bounded ((\<lambda>x. f x - g x) ` S)"
for f g::"'a \<Rightarrow> 'b::real_normed_vector"
using bounded_plus_comp[of "f" S "\<lambda>x. - g x"]
by auto
lemma bounded_minus:
fixes S ::"'a::real_normed_vector set"
assumes "bounded S" "bounded T"
shows "bounded ((\<lambda>(x,y). x - y) ` (S \<times> T))"
using bounded_minus_comp [of fst "S \<times> T" snd] assms
by (auto simp: split_def split: if_split_asm)
lemma not_bounded_UNIV[simp]:
"\<not> bounded (UNIV :: 'a::{real_normed_vector, perfect_space} set)"
proof (auto simp: bounded_pos not_le)
obtain x :: 'a where "x \<noteq> 0"
using perfect_choose_dist [OF zero_less_one] by fast
fix b :: real
assume b: "b >0"
have b1: "b +1 \<ge> 0"
using b by simp
with \<open>x \<noteq> 0\<close> have "b < norm (scaleR (b + 1) (sgn x))"
by (simp add: norm_sgn)
then show "\<exists>x::'a. b < norm x" ..
qed
corollary cobounded_imp_unbounded:
fixes S :: "'a::{real_normed_vector, perfect_space} set"
shows "bounded (- S) \<Longrightarrow> \<not> bounded S"
- using bounded_Un [of S "-S"] by (simp add: sup_compl_top)
+ using bounded_Un [of S "-S"] by (simp)
subsection\<^marker>\<open>tag unimportant\<close>\<open>Relations among convergence and absolute convergence for power series\<close>
lemma summable_imp_bounded:
fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
shows "summable f \<Longrightarrow> bounded (range f)"
by (frule summable_LIMSEQ_zero) (simp add: convergent_imp_bounded)
lemma summable_imp_sums_bounded:
"summable f \<Longrightarrow> bounded (range (\<lambda>n. sum f {..<n}))"
by (auto simp: summable_def sums_def dest: convergent_imp_bounded)
lemma power_series_conv_imp_absconv_weak:
fixes a:: "nat \<Rightarrow> 'a::{real_normed_div_algebra,banach}" and w :: 'a
assumes sum: "summable (\<lambda>n. a n * z ^ n)" and no: "norm w < norm z"
shows "summable (\<lambda>n. of_real(norm(a n)) * w ^ n)"
proof -
obtain M where M: "\<And>x. norm (a x * z ^ x) \<le> M"
using summable_imp_bounded [OF sum] by (force simp: bounded_iff)
then have *: "summable (\<lambda>n. norm (a n) * norm w ^ n)"
by (rule_tac M=M in Abel_lemma) (auto simp: norm_mult norm_power intro: no)
show ?thesis
apply (rule series_comparison_complex [of "(\<lambda>n. of_real(norm(a n) * norm w ^ n))"])
apply (simp only: summable_complex_of_real *)
apply (auto simp: norm_mult norm_power)
done
qed
subsection \<open>Normed spaces with the Heine-Borel property\<close>
lemma not_compact_UNIV[simp]:
fixes s :: "'a::{real_normed_vector,perfect_space,heine_borel} set"
shows "\<not> compact (UNIV::'a set)"
by (simp add: compact_eq_bounded_closed)
lemma not_compact_space_euclideanreal [simp]: "\<not> compact_space euclideanreal"
by (simp add: compact_space_def)
text\<open>Representing sets as the union of a chain of compact sets.\<close>
lemma closed_Union_compact_subsets:
fixes S :: "'a::{heine_borel,real_normed_vector} set"
assumes "closed S"
obtains F where "\<And>n. compact(F n)" "\<And>n. F n \<subseteq> S" "\<And>n. F n \<subseteq> F(Suc n)"
"(\<Union>n. F n) = S" "\<And>K. \<lbrakk>compact K; K \<subseteq> S\<rbrakk> \<Longrightarrow> \<exists>N. \<forall>n \<ge> N. K \<subseteq> F n"
proof
show "compact (S \<inter> cball 0 (of_nat n))" for n
using assms compact_eq_bounded_closed by auto
next
show "(\<Union>n. S \<inter> cball 0 (real n)) = S"
by (auto simp: real_arch_simple)
next
fix K :: "'a set"
assume "compact K" "K \<subseteq> S"
then obtain N where "K \<subseteq> cball 0 N"
by (meson bounded_pos mem_cball_0 compact_imp_bounded subsetI)
then show "\<exists>N. \<forall>n\<ge>N. K \<subseteq> S \<inter> cball 0 (real n)"
by (metis of_nat_le_iff Int_subset_iff \<open>K \<subseteq> S\<close> real_arch_simple subset_cball subset_trans)
qed auto
subsection \<open>Intersecting chains of compact sets and the Baire property\<close>
proposition bounded_closed_chain:
fixes \<F> :: "'a::heine_borel set set"
assumes "B \<in> \<F>" "bounded B" and \<F>: "\<And>S. S \<in> \<F> \<Longrightarrow> closed S" and "{} \<notin> \<F>"
and chain: "\<And>S T. S \<in> \<F> \<and> T \<in> \<F> \<Longrightarrow> S \<subseteq> T \<or> T \<subseteq> S"
shows "\<Inter>\<F> \<noteq> {}"
proof -
have "B \<inter> \<Inter>\<F> \<noteq> {}"
proof (rule compact_imp_fip)
show "compact B" "\<And>T. T \<in> \<F> \<Longrightarrow> closed T"
by (simp_all add: assms compact_eq_bounded_closed)
show "\<lbrakk>finite \<G>; \<G> \<subseteq> \<F>\<rbrakk> \<Longrightarrow> B \<inter> \<Inter>\<G> \<noteq> {}" for \<G>
proof (induction \<G> rule: finite_induct)
case empty
with assms show ?case by force
next
case (insert U \<G>)
then have "U \<in> \<F>" and ne: "B \<inter> \<Inter>\<G> \<noteq> {}" by auto
then consider "B \<subseteq> U" | "U \<subseteq> B"
using \<open>B \<in> \<F>\<close> chain by blast
then show ?case
proof cases
case 1
then show ?thesis
using Int_left_commute ne by auto
next
case 2
have "U \<noteq> {}"
using \<open>U \<in> \<F>\<close> \<open>{} \<notin> \<F>\<close> by blast
moreover
have False if "\<And>x. x \<in> U \<Longrightarrow> \<exists>Y\<in>\<G>. x \<notin> Y"
proof -
have "\<And>x. x \<in> U \<Longrightarrow> \<exists>Y\<in>\<G>. Y \<subseteq> U"
by (metis chain contra_subsetD insert.prems insert_subset that)
then obtain Y where "Y \<in> \<G>" "Y \<subseteq> U"
by (metis all_not_in_conv \<open>U \<noteq> {}\<close>)
moreover obtain x where "x \<in> \<Inter>\<G>"
by (metis Int_emptyI ne)
ultimately show ?thesis
by (metis Inf_lower subset_eq that)
qed
with 2 show ?thesis
by blast
qed
qed
qed
then show ?thesis by blast
qed
corollary compact_chain:
fixes \<F> :: "'a::heine_borel set set"
assumes "\<And>S. S \<in> \<F> \<Longrightarrow> compact S" "{} \<notin> \<F>"
"\<And>S T. S \<in> \<F> \<and> T \<in> \<F> \<Longrightarrow> S \<subseteq> T \<or> T \<subseteq> S"
shows "\<Inter> \<F> \<noteq> {}"
proof (cases "\<F> = {}")
case True
then show ?thesis by auto
next
case False
show ?thesis
by (metis False all_not_in_conv assms compact_imp_bounded compact_imp_closed bounded_closed_chain)
qed
lemma compact_nest:
fixes F :: "'a::linorder \<Rightarrow> 'b::heine_borel set"
assumes F: "\<And>n. compact(F n)" "\<And>n. F n \<noteq> {}" and mono: "\<And>m n. m \<le> n \<Longrightarrow> F n \<subseteq> F m"
shows "\<Inter>(range F) \<noteq> {}"
proof -
have *: "\<And>S T. S \<in> range F \<and> T \<in> range F \<Longrightarrow> S \<subseteq> T \<or> T \<subseteq> S"
by (metis mono image_iff le_cases)
show ?thesis
apply (rule compact_chain [OF _ _ *])
using F apply (blast intro: dest: *)+
done
qed
text\<open>The Baire property of dense sets\<close>
theorem Baire:
fixes S::"'a::{real_normed_vector,heine_borel} set"
assumes "closed S" "countable \<G>"
and ope: "\<And>T. T \<in> \<G> \<Longrightarrow> openin (top_of_set S) T \<and> S \<subseteq> closure T"
shows "S \<subseteq> closure(\<Inter>\<G>)"
proof (cases "\<G> = {}")
case True
then show ?thesis
using closure_subset by auto
next
let ?g = "from_nat_into \<G>"
case False
then have gin: "?g n \<in> \<G>" for n
by (simp add: from_nat_into)
show ?thesis
proof (clarsimp simp: closure_approachable)
fix x and e::real
assume "x \<in> S" "0 < e"
obtain TF where opeF: "\<And>n. openin (top_of_set S) (TF n)"
and ne: "\<And>n. TF n \<noteq> {}"
and subg: "\<And>n. S \<inter> closure(TF n) \<subseteq> ?g n"
and subball: "\<And>n. closure(TF n) \<subseteq> ball x e"
and decr: "\<And>n. TF(Suc n) \<subseteq> TF n"
proof -
have *: "\<exists>Y. (openin (top_of_set S) Y \<and> Y \<noteq> {} \<and>
S \<inter> closure Y \<subseteq> ?g n \<and> closure Y \<subseteq> ball x e) \<and> Y \<subseteq> U"
if opeU: "openin (top_of_set S) U" and "U \<noteq> {}" and cloU: "closure U \<subseteq> ball x e" for U n
proof -
obtain T where T: "open T" "U = T \<inter> S"
using \<open>openin (top_of_set S) U\<close> by (auto simp: openin_subtopology)
with \<open>U \<noteq> {}\<close> have "T \<inter> closure (?g n) \<noteq> {}"
using gin ope by fastforce
then have "T \<inter> ?g n \<noteq> {}"
using \<open>open T\<close> open_Int_closure_eq_empty by blast
then obtain y where "y \<in> U" "y \<in> ?g n"
using T ope [of "?g n", OF gin] by (blast dest: openin_imp_subset)
moreover have "openin (top_of_set S) (U \<inter> ?g n)"
using gin ope opeU by blast
ultimately obtain d where U: "U \<inter> ?g n \<subseteq> S" and "d > 0" and d: "ball y d \<inter> S \<subseteq> U \<inter> ?g n"
by (force simp: openin_contains_ball)
show ?thesis
proof (intro exI conjI)
show "openin (top_of_set S) (S \<inter> ball y (d/2))"
by (simp add: openin_open_Int)
show "S \<inter> ball y (d/2) \<noteq> {}"
using \<open>0 < d\<close> \<open>y \<in> U\<close> opeU openin_imp_subset by fastforce
have "S \<inter> closure (S \<inter> ball y (d/2)) \<subseteq> S \<inter> closure (ball y (d/2))"
using closure_mono by blast
also have "... \<subseteq> ?g n"
using \<open>d > 0\<close> d by force
finally show "S \<inter> closure (S \<inter> ball y (d/2)) \<subseteq> ?g n" .
have "closure (S \<inter> ball y (d/2)) \<subseteq> S \<inter> ball y d"
proof -
have "closure (ball y (d/2)) \<subseteq> ball y d"
using \<open>d > 0\<close> by auto
then have "closure (S \<inter> ball y (d/2)) \<subseteq> ball y d"
by (meson closure_mono inf.cobounded2 subset_trans)
then show ?thesis
by (simp add: \<open>closed S\<close> closure_minimal)
qed
also have "... \<subseteq> ball x e"
using cloU closure_subset d by blast
finally show "closure (S \<inter> ball y (d/2)) \<subseteq> ball x e" .
show "S \<inter> ball y (d/2) \<subseteq> U"
using ball_divide_subset_numeral d by blast
qed
qed
let ?\<Phi> = "\<lambda>n X. openin (top_of_set S) X \<and> X \<noteq> {} \<and>
S \<inter> closure X \<subseteq> ?g n \<and> closure X \<subseteq> ball x e"
have "closure (S \<inter> ball x (e / 2)) \<subseteq> closure(ball x (e/2))"
by (simp add: closure_mono)
also have "... \<subseteq> ball x e"
using \<open>e > 0\<close> by auto
finally have "closure (S \<inter> ball x (e / 2)) \<subseteq> ball x e" .
moreover have"openin (top_of_set S) (S \<inter> ball x (e / 2))" "S \<inter> ball x (e / 2) \<noteq> {}"
using \<open>0 < e\<close> \<open>x \<in> S\<close> by auto
ultimately obtain Y where Y: "?\<Phi> 0 Y \<and> Y \<subseteq> S \<inter> ball x (e / 2)"
using * [of "S \<inter> ball x (e/2)" 0] by metis
show thesis
proof (rule exE [OF dependent_nat_choice [of ?\<Phi> "\<lambda>n X Y. Y \<subseteq> X"]])
show "\<exists>x. ?\<Phi> 0 x"
using Y by auto
show "\<exists>Y. ?\<Phi> (Suc n) Y \<and> Y \<subseteq> X" if "?\<Phi> n X" for X n
using that by (blast intro: *)
qed (use that in metis)
qed
have "(\<Inter>n. S \<inter> closure (TF n)) \<noteq> {}"
proof (rule compact_nest)
show "\<And>n. compact (S \<inter> closure (TF n))"
by (metis closed_closure subball bounded_subset_ballI compact_eq_bounded_closed closed_Int_compact [OF \<open>closed S\<close>])
show "\<And>n. S \<inter> closure (TF n) \<noteq> {}"
by (metis Int_absorb1 opeF \<open>closed S\<close> closure_eq_empty closure_minimal ne openin_imp_subset)
show "\<And>m n. m \<le> n \<Longrightarrow> S \<inter> closure (TF n) \<subseteq> S \<inter> closure (TF m)"
by (meson closure_mono decr dual_order.refl inf_mono lift_Suc_antimono_le)
qed
moreover have "(\<Inter>n. S \<inter> closure (TF n)) \<subseteq> {y \<in> \<Inter>\<G>. dist y x < e}"
proof (clarsimp, intro conjI)
fix y
assume "y \<in> S" and y: "\<forall>n. y \<in> closure (TF n)"
then show "\<forall>T\<in>\<G>. y \<in> T"
by (metis Int_iff from_nat_into_surj [OF \<open>countable \<G>\<close>] subsetD subg)
show "dist y x < e"
by (metis y dist_commute mem_ball subball subsetCE)
qed
ultimately show "\<exists>y \<in> \<Inter>\<G>. dist y x < e"
by auto
qed
qed
subsection \<open>Continuity\<close>
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Structural rules for uniform continuity\<close>
lemma (in bounded_linear) uniformly_continuous_on[continuous_intros]:
fixes g :: "_::metric_space \<Rightarrow> _"
assumes "uniformly_continuous_on s g"
shows "uniformly_continuous_on s (\<lambda>x. f (g x))"
using assms unfolding uniformly_continuous_on_sequentially
unfolding dist_norm tendsto_norm_zero_iff diff[symmetric]
by (auto intro: tendsto_zero)
lemma uniformly_continuous_on_dist[continuous_intros]:
fixes f g :: "'a::metric_space \<Rightarrow> 'b::metric_space"
assumes "uniformly_continuous_on s f"
and "uniformly_continuous_on s g"
shows "uniformly_continuous_on s (\<lambda>x. dist (f x) (g x))"
proof -
{
fix a b c d :: 'b
have "\<bar>dist a b - dist c d\<bar> \<le> dist a c + dist b d"
using dist_triangle2 [of a b c] dist_triangle2 [of b c d]
using dist_triangle3 [of c d a] dist_triangle [of a d b]
by arith
} note le = this
{
fix x y
assume f: "(\<lambda>n. dist (f (x n)) (f (y n))) \<longlonglongrightarrow> 0"
assume g: "(\<lambda>n. dist (g (x n)) (g (y n))) \<longlonglongrightarrow> 0"
have "(\<lambda>n. \<bar>dist (f (x n)) (g (x n)) - dist (f (y n)) (g (y n))\<bar>) \<longlonglongrightarrow> 0"
by (rule Lim_transform_bound [OF _ tendsto_add_zero [OF f g]],
simp add: le)
}
then show ?thesis
using assms unfolding uniformly_continuous_on_sequentially
unfolding dist_real_def by simp
qed
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 uniformly_continuous_on_norm[continuous_intros]:
fixes f :: "'a :: metric_space \<Rightarrow> 'b :: real_normed_vector"
assumes "uniformly_continuous_on s f"
shows "uniformly_continuous_on s (\<lambda>x. norm (f x))"
unfolding norm_conv_dist using assms
by (intro uniformly_continuous_on_dist uniformly_continuous_on_const)
lemma uniformly_continuous_on_cmul[continuous_intros]:
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
assumes "uniformly_continuous_on s f"
shows "uniformly_continuous_on s (\<lambda>x. c *\<^sub>R f(x))"
using bounded_linear_scaleR_right assms
by (rule bounded_linear.uniformly_continuous_on)
lemma dist_minus:
fixes x y :: "'a::real_normed_vector"
shows "dist (- x) (- y) = dist x y"
unfolding dist_norm minus_diff_minus norm_minus_cancel ..
lemma uniformly_continuous_on_minus[continuous_intros]:
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
shows "uniformly_continuous_on s f \<Longrightarrow> uniformly_continuous_on s (\<lambda>x. - f x)"
unfolding uniformly_continuous_on_def dist_minus .
lemma uniformly_continuous_on_add[continuous_intros]:
fixes f g :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
assumes "uniformly_continuous_on s f"
and "uniformly_continuous_on s g"
shows "uniformly_continuous_on s (\<lambda>x. f x + g x)"
using assms
unfolding uniformly_continuous_on_sequentially
unfolding dist_norm tendsto_norm_zero_iff add_diff_add
by (auto intro: tendsto_add_zero)
lemma uniformly_continuous_on_diff[continuous_intros]:
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
assumes "uniformly_continuous_on s f"
and "uniformly_continuous_on s g"
shows "uniformly_continuous_on s (\<lambda>x. f x - g x)"
using assms uniformly_continuous_on_add [of s f "- g"]
by (simp add: fun_Compl_def uniformly_continuous_on_minus)
subsection\<^marker>\<open>tag unimportant\<close> \<open>Arithmetic Preserves Topological Properties\<close>
lemma open_scaling[intro]:
fixes s :: "'a::real_normed_vector set"
assumes "c \<noteq> 0"
and "open s"
shows "open((\<lambda>x. c *\<^sub>R x) ` s)"
proof -
{
fix x
assume "x \<in> s"
then obtain e where "e>0"
and e:"\<forall>x'. dist x' x < e \<longrightarrow> x' \<in> s" using assms(2)[unfolded open_dist, THEN bspec[where x=x]]
by auto
have "e * \<bar>c\<bar> > 0"
using assms(1)[unfolded zero_less_abs_iff[symmetric]] \<open>e>0\<close> by auto
moreover
{
fix y
assume "dist y (c *\<^sub>R x) < e * \<bar>c\<bar>"
then have "norm ((1 / c) *\<^sub>R y - x) < e"
unfolding dist_norm
using norm_scaleR[of c "(1 / c) *\<^sub>R y - x", unfolded scaleR_right_diff_distrib, unfolded scaleR_scaleR] assms(1)
assms(1)[unfolded zero_less_abs_iff[symmetric]] by (simp del:zero_less_abs_iff)
then have "y \<in> (*\<^sub>R) c ` s"
using rev_image_eqI[of "(1 / c) *\<^sub>R y" s y "(*\<^sub>R) c"]
using e[THEN spec[where x="(1 / c) *\<^sub>R y"]]
using assms(1)
unfolding dist_norm scaleR_scaleR
by auto
}
ultimately have "\<exists>e>0. \<forall>x'. dist x' (c *\<^sub>R x) < e \<longrightarrow> x' \<in> (*\<^sub>R) c ` s"
apply (rule_tac x="e * \<bar>c\<bar>" in exI, auto)
done
}
then show ?thesis unfolding open_dist by auto
qed
lemma minus_image_eq_vimage:
fixes A :: "'a::ab_group_add set"
shows "(\<lambda>x. - x) ` A = (\<lambda>x. - x) -` A"
by (auto intro!: image_eqI [where f="\<lambda>x. - x"])
lemma open_negations:
fixes S :: "'a::real_normed_vector set"
shows "open S \<Longrightarrow> open ((\<lambda>x. - x) ` S)"
using open_scaling [of "- 1" S] by simp
lemma open_translation:
fixes S :: "'a::real_normed_vector set"
assumes "open S"
shows "open((\<lambda>x. a + x) ` S)"
proof -
{
fix x
have "continuous (at x) (\<lambda>x. x - a)"
by (intro continuous_diff continuous_ident continuous_const)
}
moreover have "{x. x - a \<in> S} = (+) a ` S"
by force
ultimately show ?thesis
by (metis assms continuous_open_vimage vimage_def)
qed
lemma open_translation_subtract:
fixes S :: "'a::real_normed_vector set"
assumes "open S"
shows "open ((\<lambda>x. x - a) ` S)"
using assms open_translation [of S "- a"] by (simp cong: image_cong_simp)
lemma open_neg_translation:
fixes S :: "'a::real_normed_vector set"
assumes "open S"
shows "open((\<lambda>x. a - x) ` S)"
using open_translation[OF open_negations[OF assms], of a]
by (auto simp: image_image)
lemma open_affinity:
fixes S :: "'a::real_normed_vector set"
assumes "open S" "c \<noteq> 0"
shows "open ((\<lambda>x. a + c *\<^sub>R x) ` S)"
proof -
have *: "(\<lambda>x. a + c *\<^sub>R x) = (\<lambda>x. a + x) \<circ> (\<lambda>x. c *\<^sub>R x)"
unfolding o_def ..
have "(+) a ` (*\<^sub>R) c ` S = ((+) a \<circ> (*\<^sub>R) c) ` S"
by auto
then show ?thesis
using assms open_translation[of "(*\<^sub>R) c ` S" a]
unfolding *
by auto
qed
lemma interior_translation:
"interior ((+) a ` S) = (+) a ` (interior S)" for S :: "'a::real_normed_vector set"
proof (rule set_eqI, rule)
fix x
assume "x \<in> interior ((+) a ` S)"
then obtain e where "e > 0" and e: "ball x e \<subseteq> (+) a ` S"
unfolding mem_interior by auto
then have "ball (x - a) e \<subseteq> S"
unfolding subset_eq Ball_def mem_ball dist_norm
by (auto simp: diff_diff_eq)
then show "x \<in> (+) a ` interior S"
unfolding image_iff
apply (rule_tac x="x - a" in bexI)
unfolding mem_interior
using \<open>e > 0\<close>
apply auto
done
next
fix x
assume "x \<in> (+) a ` interior S"
then obtain y e where "e > 0" and e: "ball y e \<subseteq> S" and y: "x = a + y"
unfolding image_iff Bex_def mem_interior by auto
{
fix z
have *: "a + y - z = y + a - z" by auto
assume "z \<in> ball x e"
then have "z - a \<in> S"
using e[unfolded subset_eq, THEN bspec[where x="z - a"]]
unfolding mem_ball dist_norm y group_add_class.diff_diff_eq2 *
by auto
then have "z \<in> (+) a ` S"
unfolding image_iff by (auto intro!: bexI[where x="z - a"])
}
then have "ball x e \<subseteq> (+) a ` S"
unfolding subset_eq by auto
then show "x \<in> interior ((+) a ` S)"
unfolding mem_interior using \<open>e > 0\<close> by auto
qed
lemma interior_translation_subtract:
"interior ((\<lambda>x. x - a) ` S) = (\<lambda>x. x - a) ` interior S" for S :: "'a::real_normed_vector set"
using interior_translation [of "- a"] by (simp cong: image_cong_simp)
lemma compact_scaling:
fixes s :: "'a::real_normed_vector set"
assumes "compact s"
shows "compact ((\<lambda>x. c *\<^sub>R x) ` s)"
proof -
let ?f = "\<lambda>x. scaleR c x"
have *: "bounded_linear ?f" by (rule bounded_linear_scaleR_right)
show ?thesis
using compact_continuous_image[of s ?f] continuous_at_imp_continuous_on[of s ?f]
using linear_continuous_at[OF *] assms
by auto
qed
lemma compact_negations:
fixes s :: "'a::real_normed_vector set"
assumes "compact s"
shows "compact ((\<lambda>x. - x) ` s)"
using compact_scaling [OF assms, of "- 1"] by auto
lemma compact_sums:
fixes s t :: "'a::real_normed_vector set"
assumes "compact s"
and "compact t"
shows "compact {x + y | x y. x \<in> s \<and> y \<in> t}"
proof -
have *: "{x + y | x y. x \<in> s \<and> y \<in> t} = (\<lambda>z. fst z + snd z) ` (s \<times> t)"
apply auto
unfolding image_iff
apply (rule_tac x="(xa, y)" in bexI)
apply auto
done
have "continuous_on (s \<times> t) (\<lambda>z. fst z + snd z)"
unfolding continuous_on by (rule ballI) (intro tendsto_intros)
then show ?thesis
unfolding * using compact_continuous_image compact_Times [OF assms] by auto
qed
lemma compact_differences:
fixes s t :: "'a::real_normed_vector set"
assumes "compact s"
and "compact t"
shows "compact {x - y | x y. x \<in> s \<and> y \<in> t}"
proof-
have "{x - y | x y. x\<in>s \<and> y \<in> t} = {x + y | x y. x \<in> s \<and> y \<in> (uminus ` t)}"
apply auto
apply (rule_tac x= xa in exI, auto)
done
then show ?thesis
using compact_sums[OF assms(1) compact_negations[OF assms(2)]] by auto
qed
lemma compact_translation:
"compact ((+) a ` s)" if "compact s" for s :: "'a::real_normed_vector set"
proof -
have "{x + y |x y. x \<in> s \<and> y \<in> {a}} = (\<lambda>x. a + x) ` s"
by auto
then show ?thesis
using compact_sums [OF that compact_sing [of a]] by auto
qed
lemma compact_translation_subtract:
"compact ((\<lambda>x. x - a) ` s)" if "compact s" for s :: "'a::real_normed_vector set"
using that compact_translation [of s "- a"] by (simp cong: image_cong_simp)
lemma compact_affinity:
fixes s :: "'a::real_normed_vector set"
assumes "compact s"
shows "compact ((\<lambda>x. a + c *\<^sub>R x) ` s)"
proof -
have "(+) a ` (*\<^sub>R) c ` s = (\<lambda>x. a + c *\<^sub>R x) ` s"
by auto
then show ?thesis
using compact_translation[OF compact_scaling[OF assms], of a c] by auto
qed
lemma closed_scaling:
fixes S :: "'a::real_normed_vector set"
assumes "closed S"
shows "closed ((\<lambda>x. c *\<^sub>R x) ` S)"
proof (cases "c = 0")
case True then show ?thesis
by (auto simp: image_constant_conv)
next
case False
from assms have "closed ((\<lambda>x. inverse c *\<^sub>R x) -` S)"
by (simp add: continuous_closed_vimage)
also have "(\<lambda>x. inverse c *\<^sub>R x) -` S = (\<lambda>x. c *\<^sub>R x) ` S"
using \<open>c \<noteq> 0\<close> by (auto elim: image_eqI [rotated])
finally show ?thesis .
qed
lemma closed_negations:
fixes S :: "'a::real_normed_vector set"
assumes "closed S"
shows "closed ((\<lambda>x. -x) ` S)"
using closed_scaling[OF assms, of "- 1"] by simp
lemma compact_closed_sums:
fixes S :: "'a::real_normed_vector set"
assumes "compact S" and "closed T"
shows "closed (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})"
proof -
let ?S = "{x + y |x y. x \<in> S \<and> y \<in> T}"
{
fix x l
assume as: "\<forall>n. x n \<in> ?S" "(x \<longlongrightarrow> l) sequentially"
from as(1) obtain f where f: "\<forall>n. x n = fst (f n) + snd (f n)" "\<forall>n. fst (f n) \<in> S" "\<forall>n. snd (f n) \<in> T"
using choice[of "\<lambda>n y. x n = (fst y) + (snd y) \<and> fst y \<in> S \<and> snd y \<in> T"] by auto
obtain l' r where "l'\<in>S" and r: "strict_mono r" and lr: "(((\<lambda>n. fst (f n)) \<circ> r) \<longlongrightarrow> l') sequentially"
using assms(1)[unfolded compact_def, THEN spec[where x="\<lambda> n. fst (f n)"]] using f(2) by auto
have "((\<lambda>n. snd (f (r n))) \<longlongrightarrow> l - l') sequentially"
using tendsto_diff[OF LIMSEQ_subseq_LIMSEQ[OF as(2) r] lr] and f(1)
unfolding o_def
by auto
then have "l - l' \<in> T"
using assms(2)[unfolded closed_sequential_limits,
THEN spec[where x="\<lambda> n. snd (f (r n))"],
THEN spec[where x="l - l'"]]
using f(3)
by auto
then have "l \<in> ?S"
using \<open>l' \<in> S\<close>
apply auto
apply (rule_tac x=l' in exI)
apply (rule_tac x="l - l'" in exI, auto)
done
}
moreover have "?S = (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})"
by force
ultimately show ?thesis
unfolding closed_sequential_limits
by (metis (no_types, lifting))
qed
lemma closed_compact_sums:
fixes S T :: "'a::real_normed_vector set"
assumes "closed S" "compact T"
shows "closed (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})"
proof -
have "(\<Union>x\<in> T. \<Union>y \<in> S. {x + y}) = (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})"
by auto
then show ?thesis
using compact_closed_sums[OF assms(2,1)] by simp
qed
lemma compact_closed_differences:
fixes S T :: "'a::real_normed_vector set"
assumes "compact S" "closed T"
shows "closed (\<Union>x\<in> S. \<Union>y \<in> T. {x - y})"
proof -
have "(\<Union>x\<in> S. \<Union>y \<in> uminus ` T. {x + y}) = (\<Union>x\<in> S. \<Union>y \<in> T. {x - y})"
by force
then show ?thesis
using compact_closed_sums[OF assms(1) closed_negations[OF assms(2)]] by auto
qed
lemma closed_compact_differences:
fixes S T :: "'a::real_normed_vector set"
assumes "closed S" "compact T"
shows "closed (\<Union>x\<in> S. \<Union>y \<in> T. {x - y})"
proof -
have "(\<Union>x\<in> S. \<Union>y \<in> uminus ` T. {x + y}) = {x - y |x y. x \<in> S \<and> y \<in> T}"
by auto
then show ?thesis
using closed_compact_sums[OF assms(1) compact_negations[OF assms(2)]] by simp
qed
lemma closed_translation:
"closed ((+) a ` S)" if "closed S" for a :: "'a::real_normed_vector"
proof -
have "(\<Union>x\<in> {a}. \<Union>y \<in> S. {x + y}) = ((+) a ` S)" by auto
then show ?thesis
using compact_closed_sums [OF compact_sing [of a] that] by auto
qed
lemma closed_translation_subtract:
"closed ((\<lambda>x. x - a) ` S)" if "closed S" for a :: "'a::real_normed_vector"
using that closed_translation [of S "- a"] by (simp cong: image_cong_simp)
lemma closure_translation:
"closure ((+) a ` s) = (+) a ` closure s" for a :: "'a::real_normed_vector"
proof -
have *: "(+) a ` (- s) = - (+) a ` s"
by (auto intro!: image_eqI [where x = "x - a" for x])
show ?thesis
using interior_translation [of a "- s", symmetric]
by (simp add: closure_interior translation_Compl *)
qed
lemma closure_translation_subtract:
"closure ((\<lambda>x. x - a) ` s) = (\<lambda>x. x - a) ` closure s" for a :: "'a::real_normed_vector"
using closure_translation [of "- a" s] by (simp cong: image_cong_simp)
lemma frontier_translation:
"frontier ((+) a ` s) = (+) a ` frontier s" for a :: "'a::real_normed_vector"
by (auto simp add: frontier_def translation_diff interior_translation closure_translation)
lemma frontier_translation_subtract:
"frontier ((+) a ` s) = (+) a ` frontier s" for a :: "'a::real_normed_vector"
by (auto simp add: frontier_def translation_diff interior_translation closure_translation)
lemma sphere_translation:
"sphere (a + c) r = (+) a ` sphere c r" for a :: "'n::real_normed_vector"
by (auto simp: dist_norm algebra_simps intro!: image_eqI [where x = "x - a" for x])
lemma sphere_translation_subtract:
"sphere (c - a) r = (\<lambda>x. x - a) ` sphere c r" for a :: "'n::real_normed_vector"
using sphere_translation [of "- a" c] by (simp cong: image_cong_simp)
lemma cball_translation:
"cball (a + c) r = (+) a ` cball c r" for a :: "'n::real_normed_vector"
by (auto simp: dist_norm algebra_simps intro!: image_eqI [where x = "x - a" for x])
lemma cball_translation_subtract:
"cball (c - a) r = (\<lambda>x. x - a) ` cball c r" for a :: "'n::real_normed_vector"
using cball_translation [of "- a" c] by (simp cong: image_cong_simp)
lemma ball_translation:
"ball (a + c) r = (+) a ` ball c r" for a :: "'n::real_normed_vector"
by (auto simp: dist_norm algebra_simps intro!: image_eqI [where x = "x - a" for x])
lemma ball_translation_subtract:
"ball (c - a) r = (\<lambda>x. x - a) ` ball c r" for a :: "'n::real_normed_vector"
using ball_translation [of "- a" c] by (simp cong: image_cong_simp)
subsection\<^marker>\<open>tag unimportant\<close>\<open>Homeomorphisms\<close>
lemma homeomorphic_scaling:
fixes s :: "'a::real_normed_vector set"
assumes "c \<noteq> 0"
shows "s homeomorphic ((\<lambda>x. c *\<^sub>R x) ` s)"
unfolding homeomorphic_minimal
apply (rule_tac x="\<lambda>x. c *\<^sub>R x" in exI)
apply (rule_tac x="\<lambda>x. (1 / c) *\<^sub>R x" in exI)
using assms
apply (auto simp: continuous_intros)
done
lemma homeomorphic_translation:
fixes s :: "'a::real_normed_vector set"
shows "s homeomorphic ((\<lambda>x. a + x) ` s)"
unfolding homeomorphic_minimal
apply (rule_tac x="\<lambda>x. a + x" in exI)
apply (rule_tac x="\<lambda>x. -a + x" in exI)
using continuous_on_add [OF continuous_on_const continuous_on_id, of s a]
continuous_on_add [OF continuous_on_const continuous_on_id, of "plus a ` s" "- a"]
apply auto
done
lemma homeomorphic_affinity:
fixes s :: "'a::real_normed_vector set"
assumes "c \<noteq> 0"
shows "s homeomorphic ((\<lambda>x. a + c *\<^sub>R x) ` s)"
proof -
have *: "(+) a ` (*\<^sub>R) c ` s = (\<lambda>x. a + c *\<^sub>R x) ` s" by auto
show ?thesis
using homeomorphic_trans
using homeomorphic_scaling[OF assms, of s]
using homeomorphic_translation[of "(\<lambda>x. c *\<^sub>R x) ` s" a]
unfolding *
by auto
qed
lemma homeomorphic_balls:
fixes a b ::"'a::real_normed_vector"
assumes "0 < d" "0 < e"
shows "(ball a d) homeomorphic (ball b e)" (is ?th)
and "(cball a d) homeomorphic (cball b e)" (is ?cth)
proof -
show ?th unfolding homeomorphic_minimal
apply(rule_tac x="\<lambda>x. b + (e/d) *\<^sub>R (x - a)" in exI)
apply(rule_tac x="\<lambda>x. a + (d/e) *\<^sub>R (x - b)" in exI)
using assms
apply (auto intro!: continuous_intros
simp: dist_commute dist_norm pos_divide_less_eq mult_strict_left_mono)
done
show ?cth unfolding homeomorphic_minimal
apply(rule_tac x="\<lambda>x. b + (e/d) *\<^sub>R (x - a)" in exI)
apply(rule_tac x="\<lambda>x. a + (d/e) *\<^sub>R (x - b)" in exI)
using assms
apply (auto intro!: continuous_intros
simp: dist_commute dist_norm pos_divide_le_eq mult_strict_left_mono)
done
qed
lemma homeomorphic_spheres:
fixes a b ::"'a::real_normed_vector"
assumes "0 < d" "0 < e"
shows "(sphere a d) homeomorphic (sphere b e)"
unfolding homeomorphic_minimal
apply(rule_tac x="\<lambda>x. b + (e/d) *\<^sub>R (x - a)" in exI)
apply(rule_tac x="\<lambda>x. a + (d/e) *\<^sub>R (x - b)" in exI)
using assms
apply (auto intro!: continuous_intros
simp: dist_commute dist_norm pos_divide_less_eq mult_strict_left_mono)
done
lemma homeomorphic_ball01_UNIV:
"ball (0::'a::real_normed_vector) 1 homeomorphic (UNIV:: 'a set)"
(is "?B homeomorphic ?U")
proof
have "x \<in> (\<lambda>z. z /\<^sub>R (1 - norm z)) ` ball 0 1" for x::'a
apply (rule_tac x="x /\<^sub>R (1 + norm x)" in image_eqI)
apply (auto simp: field_split_simps)
using norm_ge_zero [of x] apply linarith+
done
then show "(\<lambda>z::'a. z /\<^sub>R (1 - norm z)) ` ?B = ?U"
by blast
have "x \<in> range (\<lambda>z. (1 / (1 + norm z)) *\<^sub>R z)" if "norm x < 1" for x::'a
apply (rule_tac x="x /\<^sub>R (1 - norm x)" in image_eqI)
using that apply (auto simp: field_split_simps)
done
then show "(\<lambda>z::'a. z /\<^sub>R (1 + norm z)) ` ?U = ?B"
by (force simp: field_split_simps dest: add_less_zeroD)
show "continuous_on (ball 0 1) (\<lambda>z. z /\<^sub>R (1 - norm z))"
by (rule continuous_intros | force)+
show "continuous_on UNIV (\<lambda>z. z /\<^sub>R (1 + norm z))"
apply (intro continuous_intros)
apply (metis le_add_same_cancel1 norm_ge_zero not_le zero_less_one)
done
show "\<And>x. x \<in> ball 0 1 \<Longrightarrow>
x /\<^sub>R (1 - norm x) /\<^sub>R (1 + norm (x /\<^sub>R (1 - norm x))) = x"
by (auto simp: field_split_simps)
show "\<And>y. y /\<^sub>R (1 + norm y) /\<^sub>R (1 - norm (y /\<^sub>R (1 + norm y))) = y"
apply (auto simp: field_split_simps)
apply (metis le_add_same_cancel1 norm_ge_zero not_le zero_less_one)
done
qed
proposition homeomorphic_ball_UNIV:
fixes a ::"'a::real_normed_vector"
assumes "0 < r" shows "ball a r homeomorphic (UNIV:: 'a set)"
using assms homeomorphic_ball01_UNIV homeomorphic_balls(1) homeomorphic_trans zero_less_one by blast
subsection\<^marker>\<open>tag unimportant\<close> \<open>Discrete\<close>
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
subsection\<^marker>\<open>tag unimportant\<close> \<open>Completeness of "Isometry" (up to constant bounds)\<close>
lemma cauchy_isometric:\<comment> \<open>TODO: rename lemma to \<open>Cauchy_isometric\<close>\<close>
assumes e: "e > 0"
and s: "subspace s"
and f: "bounded_linear f"
and normf: "\<forall>x\<in>s. norm (f x) \<ge> e * norm x"
and xs: "\<forall>n. x n \<in> s"
and cf: "Cauchy (f \<circ> x)"
shows "Cauchy x"
proof -
interpret f: bounded_linear f by fact
have "\<exists>N. \<forall>n\<ge>N. norm (x n - x N) < d" if "d > 0" for d :: real
proof -
from that obtain N where N: "\<forall>n\<ge>N. norm (f (x n) - f (x N)) < e * d"
using cf[unfolded Cauchy_def o_def dist_norm, THEN spec[where x="e*d"]] e
by auto
have "norm (x n - x N) < d" if "n \<ge> N" for n
proof -
have "e * norm (x n - x N) \<le> norm (f (x n - x N))"
using subspace_diff[OF s, of "x n" "x N"]
using xs[THEN spec[where x=N]] and xs[THEN spec[where x=n]]
using normf[THEN bspec[where x="x n - x N"]]
by auto
also have "norm (f (x n - x N)) < e * d"
using \<open>N \<le> n\<close> N unfolding f.diff[symmetric] by auto
finally show ?thesis
using \<open>e>0\<close> by simp
qed
then show ?thesis by auto
qed
then show ?thesis
by (simp add: Cauchy_altdef2 dist_norm)
qed
lemma complete_isometric_image:
assumes "0 < e"
and s: "subspace s"
and f: "bounded_linear f"
and normf: "\<forall>x\<in>s. norm(f x) \<ge> e * norm(x)"
and cs: "complete s"
shows "complete (f ` s)"
proof -
have "\<exists>l\<in>f ` s. (g \<longlongrightarrow> l) sequentially"
if as:"\<forall>n::nat. g n \<in> f ` s" and cfg:"Cauchy g" for g
proof -
from that obtain x where "\<forall>n. x n \<in> s \<and> g n = f (x n)"
using choice[of "\<lambda> n xa. xa \<in> s \<and> g n = f xa"] by auto
then have x: "\<forall>n. x n \<in> s" "\<forall>n. g n = f (x n)" by auto
then have "f \<circ> x = g" by (simp add: fun_eq_iff)
then obtain l where "l\<in>s" and l:"(x \<longlongrightarrow> l) sequentially"
using cs[unfolded complete_def, THEN spec[where x=x]]
using cauchy_isometric[OF \<open>0 < e\<close> s f normf] and cfg and x(1)
by auto
then show ?thesis
using linear_continuous_at[OF f, unfolded continuous_at_sequentially, THEN spec[where x=x], of l]
by (auto simp: \<open>f \<circ> x = g\<close>)
qed
then show ?thesis
unfolding complete_def by auto
qed
subsection \<open>Connected Normed Spaces\<close>
lemma compact_components:
fixes s :: "'a::heine_borel set"
shows "\<lbrakk>compact s; c \<in> components s\<rbrakk> \<Longrightarrow> compact c"
by (meson bounded_subset closed_components in_components_subset compact_eq_bounded_closed)
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 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> f constant_on S))" (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> f constant_on S))" (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> f constant_on S))" (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 "f constant_on S"
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 "f constant_on S"
using assms continuous_finite_range_constant_eq by blast
end
diff --git a/src/HOL/Analysis/Embed_Measure.thy b/src/HOL/Analysis/Embed_Measure.thy
--- a/src/HOL/Analysis/Embed_Measure.thy
+++ b/src/HOL/Analysis/Embed_Measure.thy
@@ -1,406 +1,406 @@
(* 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>Embedding Measure Spaces with a Function\<close>
theory Embed_Measure
imports Binary_Product_Measure
begin
text \<open>
Given a measure space on some carrier set \<open>\<Omega>\<close> and a function \<open>f\<close>, we can define a push-forward
measure on the carrier set \<open>f(\<Omega>)\<close> whose \<open>\<sigma>\<close>-algebra is the one generated by mapping \<open>f\<close> over
the original sigma algebra.
This is useful e.\,g.\ when \<open>f\<close> is injective, i.\,e.\ it is some kind of ``tagging'' function.
For instance, suppose we have some algebraaic datatype of values with various constructors,
including a constructor \<open>RealVal\<close> for real numbers. Then \<open>embed_measure\<close> allows us to lift a
measure on real numbers to the appropriate subset of that algebraic datatype.
\<close>
definition\<^marker>\<open>tag important\<close> 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 Setcompr_eq_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_compose)
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_compose 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 ((`) f`A)" by auto
moreover
from inj and A_props have "(`) f`A \<subseteq> sets (embed_measure M f)"
by (auto simp: sets_embed_measure)
moreover
from A_props and inj have "\<Union>((`) 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>(`) 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) = ((`) 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 ((`) 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]
+ note complete_lattice_class.Sup_insert[simp del] ccSup_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 (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 (\<le>) (f ` Y)"
and countable: "countable B"
shows "(\<integral>\<^sup>+ x. (SUP i\<in>Y. f i x) \<partial>count_space B) = (SUP i\<in>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\<in>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\<in>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\<in>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\<in>Y. \<integral>\<^sup>+ x. ?f i x \<partial>count_space UNIV)"
proof(rule nn_integral_monotone_convergence_SUP_nat)
show "Complete_Partial_Order.chain (\<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\<in>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\<in>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 = "\<lambda>A f. Sup (f ` A)"] ext nn_integral_cong_AE)(simp_all add: AE_count_space countable)
finally show ?thesis .
qed
end
diff --git a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy
--- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy
+++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy
@@ -1,5080 +1,5080 @@
(* Title: HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy
Author: Johannes Hölzl, TU München
Author: Robert Himmelmann, TU München
Huge cleanup by LCP
*)
theory Equivalence_Lebesgue_Henstock_Integration
imports
Lebesgue_Measure
Henstock_Kurzweil_Integration
Complete_Measure
Set_Integral
Homeomorphism
Cartesian_Euclidean_Space
begin
lemma le_left_mono: "x \<le> y \<Longrightarrow> y \<le> a \<longrightarrow> x \<le> (a::'a::preorder)"
by (auto intro: order_trans)
lemma ball_trans:
assumes "y \<in> ball z q" "r + q \<le> s" shows "ball y r \<subseteq> ball z s"
using assms by metric
lemma has_integral_implies_lebesgue_measurable_cbox:
fixes f :: "'a :: euclidean_space \<Rightarrow> real"
assumes f: "(f has_integral I) (cbox x y)"
shows "f \<in> lebesgue_on (cbox x y) \<rightarrow>\<^sub>M borel"
proof (rule cld_measure.borel_measurable_cld)
let ?L = "lebesgue_on (cbox x y)"
let ?\<mu> = "emeasure ?L"
let ?\<mu>' = "outer_measure_of ?L"
interpret L: finite_measure ?L
proof
show "?\<mu> (space ?L) \<noteq> \<infinity>"
by (simp add: emeasure_restrict_space space_restrict_space emeasure_lborel_cbox_eq)
qed
show "cld_measure ?L"
proof
fix B A assume "B \<subseteq> A" "A \<in> null_sets ?L"
then show "B \<in> sets ?L"
using null_sets_completion_subset[OF \<open>B \<subseteq> A\<close>, of lborel]
by (auto simp add: null_sets_restrict_space sets_restrict_space_iff intro: )
next
fix A assume "A \<subseteq> space ?L" "\<And>B. B \<in> sets ?L \<Longrightarrow> ?\<mu> B < \<infinity> \<Longrightarrow> A \<inter> B \<in> sets ?L"
from this(1) this(2)[of "space ?L"] show "A \<in> sets ?L"
by (auto simp: Int_absorb2 less_top[symmetric])
qed auto
then interpret cld_measure ?L
.
have content_eq_L: "A \<in> sets borel \<Longrightarrow> A \<subseteq> cbox x y \<Longrightarrow> content A = measure ?L A" for A
by (subst measure_restrict_space) (auto simp: measure_def)
fix E and a b :: real assume "E \<in> sets ?L" "a < b" "0 < ?\<mu> E" "?\<mu> E < \<infinity>"
then obtain M :: real where "?\<mu> E = M" "0 < M"
by (cases "?\<mu> E") auto
define e where "e = M / (4 + 2 / (b - a))"
from \<open>a < b\<close> \<open>0<M\<close> have "0 < e"
by (auto intro!: divide_pos_pos simp: field_simps e_def)
have "e < M / (3 + 2 / (b - a))"
using \<open>a < b\<close> \<open>0 < M\<close>
unfolding e_def by (intro divide_strict_left_mono add_strict_right_mono mult_pos_pos) (auto simp: field_simps)
then have "2 * e < (b - a) * (M - e * 3)"
using \<open>0<M\<close> \<open>0 < e\<close> \<open>a < b\<close> by (simp add: field_simps)
have e_less_M: "e < M / 1"
unfolding e_def using \<open>a < b\<close> \<open>0<M\<close> by (intro divide_strict_left_mono) (auto simp: field_simps)
obtain d
where "gauge d"
and integral_f: "\<forall>p. p tagged_division_of cbox x y \<and> d fine p \<longrightarrow>
norm ((\<Sum>(x,k) \<in> p. content k *\<^sub>R f x) - I) < e"
using \<open>0<e\<close> f unfolding has_integral by auto
define C where "C X m = X \<inter> {x. ball x (1/Suc m) \<subseteq> d x}" for X m
have "incseq (C X)" for X
unfolding C_def [abs_def]
by (intro monoI Collect_mono conj_mono imp_refl le_left_mono subset_ball divide_left_mono Int_mono) auto
{ fix X assume "X \<subseteq> space ?L" and eq: "?\<mu>' X = ?\<mu> E"
have "(SUP m. outer_measure_of ?L (C X m)) = outer_measure_of ?L (\<Union>m. C X m)"
using \<open>X \<subseteq> space ?L\<close> by (intro SUP_outer_measure_of_incseq \<open>incseq (C X)\<close>) (auto simp: C_def)
also have "(\<Union>m. C X m) = X"
proof -
{ fix x
obtain e where "0 < e" "ball x e \<subseteq> d x"
using gaugeD[OF \<open>gauge d\<close>, of x] unfolding open_contains_ball by auto
moreover
obtain n where "1 / (1 + real n) < e"
using reals_Archimedean[OF \<open>0<e\<close>] by (auto simp: inverse_eq_divide)
then have "ball x (1 / (1 + real n)) \<subseteq> ball x e"
by (intro subset_ball) auto
ultimately have "\<exists>n. ball x (1 / (1 + real n)) \<subseteq> d x"
by blast }
then show ?thesis
by (auto simp: C_def)
qed
finally have "(SUP m. outer_measure_of ?L (C X m)) = ?\<mu> E"
using eq by auto
also have "\<dots> > M - e"
using \<open>0 < M\<close> \<open>?\<mu> E = M\<close> \<open>0<e\<close> by (auto intro!: ennreal_lessI)
finally have "\<exists>m. M - e < outer_measure_of ?L (C X m)"
unfolding less_SUP_iff by auto }
note C = this
let ?E = "{x\<in>E. f x \<le> a}" and ?F = "{x\<in>E. b \<le> f x}"
have "\<not> (?\<mu>' ?E = ?\<mu> E \<and> ?\<mu>' ?F = ?\<mu> E)"
proof
assume eq: "?\<mu>' ?E = ?\<mu> E \<and> ?\<mu>' ?F = ?\<mu> E"
with C[of ?E] C[of ?F] \<open>E \<in> sets ?L\<close>[THEN sets.sets_into_space] obtain ma mb
where "M - e < outer_measure_of ?L (C ?E ma)" "M - e < outer_measure_of ?L (C ?F mb)"
by auto
moreover define m where "m = max ma mb"
ultimately have M_minus_e: "M - e < outer_measure_of ?L (C ?E m)" "M - e < outer_measure_of ?L (C ?F m)"
using
incseqD[OF \<open>incseq (C ?E)\<close>, of ma m, THEN outer_measure_of_mono]
incseqD[OF \<open>incseq (C ?F)\<close>, of mb m, THEN outer_measure_of_mono]
by (auto intro: less_le_trans)
define d' where "d' x = d x \<inter> ball x (1 / (3 * Suc m))" for x
have "gauge d'"
unfolding d'_def by (intro gauge_Int \<open>gauge d\<close> gauge_ball) auto
then obtain p where p: "p tagged_division_of cbox x y" "d' fine p"
by (rule fine_division_exists)
then have "d fine p"
unfolding d'_def[abs_def] fine_def by auto
define s where "s = {(x::'a, k). k \<inter> (C ?E m) \<noteq> {} \<and> k \<inter> (C ?F m) \<noteq> {}}"
define T where "T E k = (SOME x. x \<in> k \<inter> C E m)" for E k
let ?A = "(\<lambda>(x, k). (T ?E k, k)) ` (p \<inter> s) \<union> (p - s)"
let ?B = "(\<lambda>(x, k). (T ?F k, k)) ` (p \<inter> s) \<union> (p - s)"
{ fix X assume X_eq: "X = ?E \<or> X = ?F"
let ?T = "(\<lambda>(x, k). (T X k, k))"
let ?p = "?T ` (p \<inter> s) \<union> (p - s)"
have in_s: "(x, k) \<in> s \<Longrightarrow> T X k \<in> k \<inter> C X m" for x k
using someI_ex[of "\<lambda>x. x \<in> k \<inter> C X m"] X_eq unfolding ex_in_conv by (auto simp: T_def s_def)
{ fix x k assume "(x, k) \<in> p" "(x, k) \<in> s"
have k: "k \<subseteq> ball x (1 / (3 * Suc m))"
using \<open>d' fine p\<close>[THEN fineD, OF \<open>(x, k) \<in> p\<close>] by (auto simp: d'_def)
then have "x \<in> ball (T X k) (1 / (3 * Suc m))"
using in_s[OF \<open>(x, k) \<in> s\<close>] by (auto simp: C_def subset_eq dist_commute)
then have "ball x (1 / (3 * Suc m)) \<subseteq> ball (T X k) (1 / Suc m)"
by (rule ball_trans) (auto simp: field_split_simps)
with k in_s[OF \<open>(x, k) \<in> s\<close>] have "k \<subseteq> d (T X k)"
by (auto simp: C_def) }
then have "d fine ?p"
using \<open>d fine p\<close> by (auto intro!: fineI)
moreover
have "?p tagged_division_of cbox x y"
proof (rule tagged_division_ofI)
show "finite ?p"
using p(1) by auto
next
fix z k assume *: "(z, k) \<in> ?p"
then consider "(z, k) \<in> p" "(z, k) \<notin> s"
| x' where "(x', k) \<in> p" "(x', k) \<in> s" "z = T X k"
by (auto simp: T_def)
then have "z \<in> k \<and> k \<subseteq> cbox x y \<and> (\<exists>a b. k = cbox a b)"
using p(1) by cases (auto dest: in_s)
then show "z \<in> k" "k \<subseteq> cbox x y" "\<exists>a b. k = cbox a b"
by auto
next
fix z k z' k' assume "(z, k) \<in> ?p" "(z', k') \<in> ?p" "(z, k) \<noteq> (z', k')"
with tagged_division_ofD(5)[OF p(1), of _ k _ k']
show "interior k \<inter> interior k' = {}"
by (auto simp: T_def dest: in_s)
next
have "{k. \<exists>x. (x, k) \<in> ?p} = {k. \<exists>x. (x, k) \<in> p}"
by (auto simp: T_def image_iff Bex_def)
then show "\<Union>{k. \<exists>x. (x, k) \<in> ?p} = cbox x y"
using p(1) by auto
qed
ultimately have I: "norm ((\<Sum>(x,k) \<in> ?p. content k *\<^sub>R f x) - I) < e"
using integral_f by auto
have "(\<Sum>(x,k) \<in> ?p. content k *\<^sub>R f x) =
(\<Sum>(x,k) \<in> ?T ` (p \<inter> s). content k *\<^sub>R f x) + (\<Sum>(x,k) \<in> p - s. content k *\<^sub>R f x)"
using p(1)[THEN tagged_division_ofD(1)]
by (safe intro!: sum.union_inter_neutral) (auto simp: s_def T_def)
also have "(\<Sum>(x,k) \<in> ?T ` (p \<inter> s). content k *\<^sub>R f x) = (\<Sum>(x,k) \<in> p \<inter> s. content k *\<^sub>R f (T X k))"
proof (subst sum.reindex_nontrivial, safe)
fix x1 x2 k assume 1: "(x1, k) \<in> p" "(x1, k) \<in> s" and 2: "(x2, k) \<in> p" "(x2, k) \<in> s"
and eq: "content k *\<^sub>R f (T X k) \<noteq> 0"
with tagged_division_ofD(5)[OF p(1), of x1 k x2 k] tagged_division_ofD(4)[OF p(1), of x1 k]
show "x1 = x2"
by (auto simp: content_eq_0_interior)
qed (use p in \<open>auto intro!: sum.cong\<close>)
finally have eq: "(\<Sum>(x,k) \<in> ?p. content k *\<^sub>R f x) =
(\<Sum>(x,k) \<in> p \<inter> s. content k *\<^sub>R f (T X k)) + (\<Sum>(x,k) \<in> p - s. content k *\<^sub>R f x)" .
have in_T: "(x, k) \<in> s \<Longrightarrow> T X k \<in> X" for x k
using in_s[of x k] by (auto simp: C_def)
note I eq in_T }
note parts = this
have p_in_L: "(x, k) \<in> p \<Longrightarrow> k \<in> sets ?L" for x k
using tagged_division_ofD(3, 4)[OF p(1), of x k] by (auto simp: sets_restrict_space)
have [simp]: "finite p"
using tagged_division_ofD(1)[OF p(1)] .
have "(M - 3*e) * (b - a) \<le> (\<Sum>(x,k) \<in> p \<inter> s. content k) * (b - a)"
proof (intro mult_right_mono)
have fin: "?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}}) < \<infinity>" for X
using \<open>?\<mu> E < \<infinity>\<close> by (rule le_less_trans[rotated]) (auto intro!: emeasure_mono \<open>E \<in> sets ?L\<close>)
have sets: "(E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}}) \<in> sets ?L" for X
using tagged_division_ofD(1)[OF p(1)] by (intro sets.Diff \<open>E \<in> sets ?L\<close> sets.finite_Union sets.Int) (auto intro: p_in_L)
{ fix X assume "X \<subseteq> E" "M - e < ?\<mu>' (C X m)"
have "M - e \<le> ?\<mu>' (C X m)"
by (rule less_imp_le) fact
also have "\<dots> \<le> ?\<mu>' (E - (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}}))"
proof (intro outer_measure_of_mono subsetI)
fix v assume "v \<in> C X m"
then have "v \<in> cbox x y" "v \<in> E"
using \<open>E \<subseteq> space ?L\<close> \<open>X \<subseteq> E\<close> by (auto simp: space_restrict_space C_def)
then obtain z k where "(z, k) \<in> p" "v \<in> k"
using tagged_division_ofD(6)[OF p(1), symmetric] by auto
then show "v \<in> E - E \<inter> (\<Union>{k\<in>snd`p. k \<inter> C X m = {}})"
using \<open>v \<in> C X m\<close> \<open>v \<in> E\<close> by auto
qed
also have "\<dots> = ?\<mu> E - ?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}})"
using \<open>E \<in> sets ?L\<close> fin[of X] sets[of X] by (auto intro!: emeasure_Diff)
finally have "?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}}) \<le> e"
using \<open>0 < e\<close> e_less_M apply (cases "?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}})")
by (auto simp add: \<open>?\<mu> E = M\<close> ennreal_minus ennreal_le_iff2)
note this }
note upper_bound = this
have "?\<mu> (E \<inter> \<Union>(snd`(p - s))) =
?\<mu> ((E \<inter> \<Union>{k\<in>snd`p. k \<inter> C ?E m = {}}) \<union> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C ?F m = {}}))"
by (intro arg_cong[where f="?\<mu>"]) (auto simp: s_def image_def Bex_def)
also have "\<dots> \<le> ?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C ?E m = {}}) + ?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C ?F m = {}})"
using sets[of ?E] sets[of ?F] M_minus_e by (intro emeasure_subadditive) auto
also have "\<dots> \<le> e + ennreal e"
using upper_bound[of ?E] upper_bound[of ?F] M_minus_e by (intro add_mono) auto
finally have "?\<mu> E - 2*e \<le> ?\<mu> (E - (E \<inter> \<Union>(snd`(p - s))))"
using \<open>0 < e\<close> \<open>E \<in> sets ?L\<close> tagged_division_ofD(1)[OF p(1)]
by (subst emeasure_Diff)
(auto simp: top_unique simp flip: ennreal_plus
intro!: sets.Int sets.finite_UN ennreal_mono_minus intro: p_in_L)
also have "\<dots> \<le> ?\<mu> (\<Union>x\<in>p \<inter> s. snd x)"
proof (safe intro!: emeasure_mono subsetI)
fix v assume "v \<in> E" and not: "v \<notin> (\<Union>x\<in>p \<inter> s. snd x)"
then have "v \<in> cbox x y"
using \<open>E \<subseteq> space ?L\<close> by (auto simp: space_restrict_space)
then obtain z k where "(z, k) \<in> p" "v \<in> k"
using tagged_division_ofD(6)[OF p(1), symmetric] by auto
with not show "v \<in> \<Union>(snd ` (p - s))"
by (auto intro!: bexI[of _ "(z, k)"] elim: ballE[of _ _ "(z, k)"])
qed (auto intro!: sets.Int sets.finite_UN ennreal_mono_minus intro: p_in_L)
also have "\<dots> = measure ?L (\<Union>x\<in>p \<inter> s. snd x)"
by (auto intro!: emeasure_eq_ennreal_measure)
finally have "M - 2 * e \<le> measure ?L (\<Union>x\<in>p \<inter> s. snd x)"
unfolding \<open>?\<mu> E = M\<close> using \<open>0 < e\<close> by (simp add: ennreal_minus)
also have "measure ?L (\<Union>x\<in>p \<inter> s. snd x) = content (\<Union>x\<in>p \<inter> s. snd x)"
using tagged_division_ofD(1,3,4) [OF p(1)]
by (intro content_eq_L[symmetric])
(fastforce intro!: sets.finite_UN UN_least del: subsetI)+
also have "content (\<Union>x\<in>p \<inter> s. snd x) \<le> (\<Sum>k\<in>p \<inter> s. content (snd k))"
using p(1) by (auto simp: emeasure_lborel_cbox_eq intro!: measure_subadditive_finite
dest!: p(1)[THEN tagged_division_ofD(4)])
finally show "M - 3 * e \<le> (\<Sum>(x, y)\<in>p \<inter> s. content y)"
using \<open>0 < e\<close> by (simp add: split_beta)
qed (use \<open>a < b\<close> in auto)
also have "\<dots> = (\<Sum>(x,k) \<in> p \<inter> s. content k * (b - a))"
by (simp add: sum_distrib_right split_beta')
also have "\<dots> \<le> (\<Sum>(x,k) \<in> p \<inter> s. content k * (f (T ?F k) - f (T ?E k)))"
using parts(3) by (auto intro!: sum_mono mult_left_mono diff_mono)
also have "\<dots> = (\<Sum>(x,k) \<in> p \<inter> s. content k * f (T ?F k)) - (\<Sum>(x,k) \<in> p \<inter> s. content k * f (T ?E k))"
by (auto intro!: sum.cong simp: field_simps sum_subtractf[symmetric])
also have "\<dots> = (\<Sum>(x,k) \<in> ?B. content k *\<^sub>R f x) - (\<Sum>(x,k) \<in> ?A. content k *\<^sub>R f x)"
by (subst (1 2) parts) auto
also have "\<dots> \<le> norm ((\<Sum>(x,k) \<in> ?B. content k *\<^sub>R f x) - (\<Sum>(x,k) \<in> ?A. content k *\<^sub>R f x))"
by auto
also have "\<dots> \<le> e + e"
using parts(1)[of ?E] parts(1)[of ?F] by (intro norm_diff_triangle_le[of _ I]) auto
finally show False
using \<open>2 * e < (b - a) * (M - e * 3)\<close> by (auto simp: field_simps)
qed
moreover have "?\<mu>' ?E \<le> ?\<mu> E" "?\<mu>' ?F \<le> ?\<mu> E"
unfolding outer_measure_of_eq[OF \<open>E \<in> sets ?L\<close>, symmetric] by (auto intro!: outer_measure_of_mono)
ultimately show "min (?\<mu>' ?E) (?\<mu>' ?F) < ?\<mu> E"
unfolding min_less_iff_disj by (auto simp: less_le)
qed
lemma has_integral_implies_lebesgue_measurable_real:
fixes f :: "'a :: euclidean_space \<Rightarrow> real"
assumes f: "(f has_integral I) \<Omega>"
shows "(\<lambda>x. f x * indicator \<Omega> x) \<in> lebesgue \<rightarrow>\<^sub>M borel"
proof -
define B :: "nat \<Rightarrow> 'a set" where "B n = cbox (- real n *\<^sub>R One) (real n *\<^sub>R One)" for n
show "(\<lambda>x. f x * indicator \<Omega> x) \<in> lebesgue \<rightarrow>\<^sub>M borel"
proof (rule measurable_piecewise_restrict)
have "(\<Union>n. box (- real n *\<^sub>R One) (real n *\<^sub>R One)) \<subseteq> \<Union>(B ` UNIV)"
unfolding B_def by (intro UN_mono box_subset_cbox order_refl)
then show "countable (range B)" "space lebesgue \<subseteq> \<Union>(B ` UNIV)"
by (auto simp: B_def UN_box_eq_UNIV)
next
fix \<Omega>' assume "\<Omega>' \<in> range B"
then obtain n where \<Omega>': "\<Omega>' = B n" by auto
then show "\<Omega>' \<inter> space lebesgue \<in> sets lebesgue"
by (auto simp: B_def)
have "f integrable_on \<Omega>"
using f by auto
then have "(\<lambda>x. f x * indicator \<Omega> x) integrable_on \<Omega>"
by (auto simp: integrable_on_def cong: has_integral_cong)
then have "(\<lambda>x. f x * indicator \<Omega> x) integrable_on (\<Omega> \<union> B n)"
by (rule integrable_on_superset) auto
then have "(\<lambda>x. f x * indicator \<Omega> x) integrable_on B n"
unfolding B_def by (rule integrable_on_subcbox) auto
then show "(\<lambda>x. f x * indicator \<Omega> x) \<in> lebesgue_on \<Omega>' \<rightarrow>\<^sub>M borel"
unfolding B_def \<Omega>' by (auto intro: has_integral_implies_lebesgue_measurable_cbox simp: integrable_on_def)
qed
qed
lemma has_integral_implies_lebesgue_measurable:
fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
assumes f: "(f has_integral I) \<Omega>"
shows "(\<lambda>x. indicator \<Omega> x *\<^sub>R f x) \<in> lebesgue \<rightarrow>\<^sub>M borel"
proof (intro borel_measurable_euclidean_space[where 'c='b, THEN iffD2] ballI)
fix i :: "'b" assume "i \<in> Basis"
have "(\<lambda>x. (f x \<bullet> i) * indicator \<Omega> x) \<in> borel_measurable (completion lborel)"
using has_integral_linear[OF f bounded_linear_inner_left, of i]
by (intro has_integral_implies_lebesgue_measurable_real) (auto simp: comp_def)
then show "(\<lambda>x. indicator \<Omega> x *\<^sub>R f x \<bullet> i) \<in> borel_measurable (completion lborel)"
by (simp add: ac_simps)
qed
subsection \<open>Equivalence Lebesgue integral on \<^const>\<open>lborel\<close> and HK-integral\<close>
lemma has_integral_measure_lborel:
fixes A :: "'a::euclidean_space set"
assumes A[measurable]: "A \<in> sets borel" and finite: "emeasure lborel A < \<infinity>"
shows "((\<lambda>x. 1) has_integral measure lborel A) A"
proof -
{ fix l u :: 'a
have "((\<lambda>x. 1) has_integral measure lborel (box l u)) (box l u)"
proof cases
assume "\<forall>b\<in>Basis. l \<bullet> b \<le> u \<bullet> b"
then show ?thesis
apply simp
apply (subst has_integral_restrict[symmetric, OF box_subset_cbox])
apply (subst has_integral_spike_interior_eq[where g="\<lambda>_. 1"])
using has_integral_const[of "1::real" l u]
apply (simp_all add: inner_diff_left[symmetric] content_cbox_cases)
done
next
assume "\<not> (\<forall>b\<in>Basis. l \<bullet> b \<le> u \<bullet> b)"
then have "box l u = {}"
unfolding box_eq_empty by (auto simp: not_le intro: less_imp_le)
then show ?thesis
by simp
qed }
note has_integral_box = this
{ fix a b :: 'a let ?M = "\<lambda>A. measure lborel (A \<inter> box a b)"
have "Int_stable (range (\<lambda>(a, b). box a b))"
by (auto simp: Int_stable_def box_Int_box)
moreover have "(range (\<lambda>(a, b). box a b)) \<subseteq> Pow UNIV"
by auto
moreover have "A \<in> sigma_sets UNIV (range (\<lambda>(a, b). box a b))"
using A unfolding borel_eq_box by simp
ultimately have "((\<lambda>x. 1) has_integral ?M A) (A \<inter> box a b)"
proof (induction rule: sigma_sets_induct_disjoint)
case (basic A) then show ?case
by (auto simp: box_Int_box has_integral_box)
next
case empty then show ?case
by simp
next
case (compl A)
then have [measurable]: "A \<in> sets borel"
by (simp add: borel_eq_box)
have "((\<lambda>x. 1) has_integral ?M (box a b)) (box a b)"
by (simp add: has_integral_box)
moreover have "((\<lambda>x. if x \<in> A \<inter> box a b then 1 else 0) has_integral ?M A) (box a b)"
by (subst has_integral_restrict) (auto intro: compl)
ultimately have "((\<lambda>x. 1 - (if x \<in> A \<inter> box a b then 1 else 0)) has_integral ?M (box a b) - ?M A) (box a b)"
by (rule has_integral_diff)
then have "((\<lambda>x. (if x \<in> (UNIV - A) \<inter> box a b then 1 else 0)) has_integral ?M (box a b) - ?M A) (box a b)"
by (rule has_integral_cong[THEN iffD1, rotated 1]) auto
then have "((\<lambda>x. 1) has_integral ?M (box a b) - ?M A) ((UNIV - A) \<inter> box a b)"
by (subst (asm) has_integral_restrict) auto
also have "?M (box a b) - ?M A = ?M (UNIV - A)"
by (subst measure_Diff[symmetric]) (auto simp: emeasure_lborel_box_eq Diff_Int_distrib2)
finally show ?case .
next
case (union F)
then have [measurable]: "\<And>i. F i \<in> sets borel"
by (simp add: borel_eq_box subset_eq)
have "((\<lambda>x. if x \<in> \<Union>(F ` UNIV) \<inter> box a b then 1 else 0) has_integral ?M (\<Union>i. F i)) (box a b)"
proof (rule has_integral_monotone_convergence_increasing)
let ?f = "\<lambda>k x. \<Sum>i<k. if x \<in> F i \<inter> box a b then 1 else 0 :: real"
show "\<And>k. (?f k has_integral (\<Sum>i<k. ?M (F i))) (box a b)"
using union.IH by (auto intro!: has_integral_sum simp del: Int_iff)
show "\<And>k x. ?f k x \<le> ?f (Suc k) x"
by (intro sum_mono2) auto
from union(1) have *: "\<And>x i j. x \<in> F i \<Longrightarrow> x \<in> F j \<longleftrightarrow> j = i"
by (auto simp add: disjoint_family_on_def)
show "\<And>x. (\<lambda>k. ?f k x) \<longlonglongrightarrow> (if x \<in> \<Union>(F ` UNIV) \<inter> box a b then 1 else 0)"
apply (auto simp: * sum.If_cases Iio_Int_singleton)
apply (rule_tac k="Suc xa" in LIMSEQ_offset)
apply simp
done
have *: "emeasure lborel ((\<Union>x. F x) \<inter> box a b) \<le> emeasure lborel (box a b)"
by (intro emeasure_mono) auto
with union(1) show "(\<lambda>k. \<Sum>i<k. ?M (F i)) \<longlonglongrightarrow> ?M (\<Union>i. F i)"
unfolding sums_def[symmetric] UN_extend_simps
by (intro measure_UNION) (auto simp: disjoint_family_on_def emeasure_lborel_box_eq top_unique)
qed
then show ?case
by (subst (asm) has_integral_restrict) auto
qed }
note * = this
show ?thesis
proof (rule has_integral_monotone_convergence_increasing)
let ?B = "\<lambda>n::nat. box (- real n *\<^sub>R One) (real n *\<^sub>R One) :: 'a set"
let ?f = "\<lambda>n::nat. \<lambda>x. if x \<in> A \<inter> ?B n then 1 else 0 :: real"
let ?M = "\<lambda>n. measure lborel (A \<inter> ?B n)"
show "\<And>n::nat. (?f n has_integral ?M n) A"
using * by (subst has_integral_restrict) simp_all
show "\<And>k x. ?f k x \<le> ?f (Suc k) x"
by (auto simp: box_def)
{ fix x assume "x \<in> A"
moreover have "(\<lambda>k. indicator (A \<inter> ?B k) x :: real) \<longlonglongrightarrow> indicator (\<Union>k::nat. A \<inter> ?B k) x"
by (intro LIMSEQ_indicator_incseq) (auto simp: incseq_def box_def)
ultimately show "(\<lambda>k. if x \<in> A \<inter> ?B k then 1 else 0::real) \<longlonglongrightarrow> 1"
by (simp add: indicator_def UN_box_eq_UNIV) }
have "(\<lambda>n. emeasure lborel (A \<inter> ?B n)) \<longlonglongrightarrow> emeasure lborel (\<Union>n::nat. A \<inter> ?B n)"
by (intro Lim_emeasure_incseq) (auto simp: incseq_def box_def)
also have "(\<lambda>n. emeasure lborel (A \<inter> ?B n)) = (\<lambda>n. measure lborel (A \<inter> ?B n))"
proof (intro ext emeasure_eq_ennreal_measure)
fix n have "emeasure lborel (A \<inter> ?B n) \<le> emeasure lborel (?B n)"
by (intro emeasure_mono) auto
then show "emeasure lborel (A \<inter> ?B n) \<noteq> top"
by (auto simp: top_unique)
qed
finally show "(\<lambda>n. measure lborel (A \<inter> ?B n)) \<longlonglongrightarrow> measure lborel A"
using emeasure_eq_ennreal_measure[of lborel A] finite
by (simp add: UN_box_eq_UNIV less_top)
qed
qed
lemma nn_integral_has_integral:
fixes f::"'a::euclidean_space \<Rightarrow> real"
assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lborel) = ennreal r" "0 \<le> r"
shows "(f has_integral r) UNIV"
using f proof (induct f arbitrary: r rule: borel_measurable_induct_real)
case (set A)
then have "((\<lambda>x. 1) has_integral measure lborel A) A"
by (intro has_integral_measure_lborel) (auto simp: ennreal_indicator)
with set show ?case
by (simp add: ennreal_indicator measure_def) (simp add: indicator_def)
next
case (mult g c)
then have "ennreal c * (\<integral>\<^sup>+ x. g x \<partial>lborel) = ennreal r"
by (subst nn_integral_cmult[symmetric]) (auto simp: ennreal_mult)
with \<open>0 \<le> r\<close> \<open>0 \<le> c\<close>
obtain r' where "(c = 0 \<and> r = 0) \<or> (0 \<le> r' \<and> (\<integral>\<^sup>+ x. ennreal (g x) \<partial>lborel) = ennreal r' \<and> r = c * r')"
by (cases "\<integral>\<^sup>+ x. ennreal (g x) \<partial>lborel" rule: ennreal_cases)
(auto split: if_split_asm simp: ennreal_mult_top ennreal_mult[symmetric])
with mult show ?case
by (auto intro!: has_integral_cmult_real)
next
case (add g h)
then have "(\<integral>\<^sup>+ x. h x + g x \<partial>lborel) = (\<integral>\<^sup>+ x. h x \<partial>lborel) + (\<integral>\<^sup>+ x. g x \<partial>lborel)"
by (simp add: nn_integral_add)
with add obtain a b where "0 \<le> a" "0 \<le> b" "(\<integral>\<^sup>+ x. h x \<partial>lborel) = ennreal a" "(\<integral>\<^sup>+ x. g x \<partial>lborel) = ennreal b" "r = a + b"
by (cases "\<integral>\<^sup>+ x. h x \<partial>lborel" "\<integral>\<^sup>+ x. g x \<partial>lborel" rule: ennreal2_cases)
(auto simp: add_top nn_integral_add top_add simp flip: ennreal_plus)
with add show ?case
by (auto intro!: has_integral_add)
next
case (seq U)
note seq(1)[measurable] and f[measurable]
{ fix i x
have "U i x \<le> f x"
using seq(5)
apply (rule LIMSEQ_le_const)
using seq(4)
apply (auto intro!: exI[of _ i] simp: incseq_def le_fun_def)
done }
note U_le_f = this
{ fix i
have "(\<integral>\<^sup>+x. U i x \<partial>lborel) \<le> (\<integral>\<^sup>+x. f x \<partial>lborel)"
using seq(2) f(2) U_le_f by (intro nn_integral_mono) simp
then obtain p where "(\<integral>\<^sup>+x. U i x \<partial>lborel) = ennreal p" "p \<le> r" "0 \<le> p"
using seq(6) \<open>0\<le>r\<close> by (cases "\<integral>\<^sup>+x. U i x \<partial>lborel" rule: ennreal_cases) (auto simp: top_unique)
moreover note seq
ultimately have "\<exists>p. (\<integral>\<^sup>+x. U i x \<partial>lborel) = ennreal p \<and> 0 \<le> p \<and> p \<le> r \<and> (U i has_integral p) UNIV"
by auto }
then obtain p where p: "\<And>i. (\<integral>\<^sup>+x. ennreal (U i x) \<partial>lborel) = ennreal (p i)"
and bnd: "\<And>i. p i \<le> r" "\<And>i. 0 \<le> p i"
and U_int: "\<And>i.(U i has_integral (p i)) UNIV" by metis
have int_eq: "\<And>i. integral UNIV (U i) = p i" using U_int by (rule integral_unique)
have *: "f integrable_on UNIV \<and> (\<lambda>k. integral UNIV (U k)) \<longlonglongrightarrow> integral UNIV f"
proof (rule monotone_convergence_increasing)
show "\<And>k. U k integrable_on UNIV" using U_int by auto
show "\<And>k x. x\<in>UNIV \<Longrightarrow> U k x \<le> U (Suc k) x" using \<open>incseq U\<close> by (auto simp: incseq_def le_fun_def)
then show "bounded (range (\<lambda>k. integral UNIV (U k)))"
using bnd int_eq by (auto simp: bounded_real intro!: exI[of _ r])
show "\<And>x. x\<in>UNIV \<Longrightarrow> (\<lambda>k. U k x) \<longlonglongrightarrow> f x"
using seq by auto
qed
moreover have "(\<lambda>i. (\<integral>\<^sup>+x. U i x \<partial>lborel)) \<longlonglongrightarrow> (\<integral>\<^sup>+x. f x \<partial>lborel)"
using seq f(2) U_le_f by (intro nn_integral_dominated_convergence[where w=f]) auto
ultimately have "integral UNIV f = r"
by (auto simp add: bnd int_eq p seq intro: LIMSEQ_unique)
with * show ?case
by (simp add: has_integral_integral)
qed
lemma nn_integral_lborel_eq_integral:
fixes f::"'a::euclidean_space \<Rightarrow> real"
assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lborel) < \<infinity>"
shows "(\<integral>\<^sup>+x. f x \<partial>lborel) = integral UNIV f"
proof -
from f(3) obtain r where r: "(\<integral>\<^sup>+x. f x \<partial>lborel) = ennreal r" "0 \<le> r"
by (cases "\<integral>\<^sup>+x. f x \<partial>lborel" rule: ennreal_cases) auto
then show ?thesis
using nn_integral_has_integral[OF f(1,2) r] by (simp add: integral_unique)
qed
lemma nn_integral_integrable_on:
fixes f::"'a::euclidean_space \<Rightarrow> real"
assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lborel) < \<infinity>"
shows "f integrable_on UNIV"
proof -
from f(3) obtain r where r: "(\<integral>\<^sup>+x. f x \<partial>lborel) = ennreal r" "0 \<le> r"
by (cases "\<integral>\<^sup>+x. f x \<partial>lborel" rule: ennreal_cases) auto
then show ?thesis
by (intro has_integral_integrable[where i=r] nn_integral_has_integral[where r=r] f)
qed
lemma nn_integral_has_integral_lborel:
fixes f :: "'a::euclidean_space \<Rightarrow> real"
assumes f_borel: "f \<in> borel_measurable borel" and nonneg: "\<And>x. 0 \<le> f x"
assumes I: "(f has_integral I) UNIV"
shows "integral\<^sup>N lborel f = I"
proof -
from f_borel have "(\<lambda>x. ennreal (f x)) \<in> borel_measurable lborel" by auto
from borel_measurable_implies_simple_function_sequence'[OF this]
obtain F where F: "\<And>i. simple_function lborel (F i)" "incseq F"
"\<And>i x. F i x < top" "\<And>x. (SUP i. F i x) = ennreal (f x)"
by blast
then have [measurable]: "\<And>i. F i \<in> borel_measurable lborel"
by (metis borel_measurable_simple_function)
let ?B = "\<lambda>i::nat. box (- (real i *\<^sub>R One)) (real i *\<^sub>R One) :: 'a set"
have "0 \<le> I"
using I by (rule has_integral_nonneg) (simp add: nonneg)
have F_le_f: "enn2real (F i x) \<le> f x" for i x
using F(3,4)[where x=x] nonneg SUP_upper[of i UNIV "\<lambda>i. F i x"]
by (cases "F i x" rule: ennreal_cases) auto
let ?F = "\<lambda>i x. F i x * indicator (?B i) x"
have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>lborel) = (SUP i. integral\<^sup>N lborel (\<lambda>x. ?F i x))"
proof (subst nn_integral_monotone_convergence_SUP[symmetric])
{ fix x
obtain j where j: "x \<in> ?B j"
using UN_box_eq_UNIV by auto
have "ennreal (f x) = (SUP i. F i x)"
using F(4)[of x] nonneg[of x] by (simp add: max_def)
also have "\<dots> = (SUP i. ?F i x)"
proof (rule SUP_eq)
fix i show "\<exists>j\<in>UNIV. F i x \<le> ?F j x"
using j F(2)
by (intro bexI[of _ "max i j"])
(auto split: split_max split_indicator simp: incseq_def le_fun_def box_def)
qed (auto intro!: F split: split_indicator)
finally have "ennreal (f x) = (SUP i. ?F i x)" . }
then show "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>lborel) = (\<integral>\<^sup>+ x. (SUP i. ?F i x) \<partial>lborel)"
by simp
qed (insert F, auto simp: incseq_def le_fun_def box_def split: split_indicator)
also have "\<dots> \<le> ennreal I"
proof (rule SUP_least)
fix i :: nat
have finite_F: "(\<integral>\<^sup>+ x. ennreal (enn2real (F i x) * indicator (?B i) x) \<partial>lborel) < \<infinity>"
proof (rule nn_integral_bound_simple_function)
have "emeasure lborel {x \<in> space lborel. ennreal (enn2real (F i x) * indicator (?B i) x) \<noteq> 0} \<le>
emeasure lborel (?B i)"
by (intro emeasure_mono) (auto split: split_indicator)
then show "emeasure lborel {x \<in> space lborel. ennreal (enn2real (F i x) * indicator (?B i) x) \<noteq> 0} < \<infinity>"
by (auto simp: less_top[symmetric] top_unique)
qed (auto split: split_indicator
intro!: F simple_function_compose1[where g="enn2real"] simple_function_ennreal)
have int_F: "(\<lambda>x. enn2real (F i x) * indicator (?B i) x) integrable_on UNIV"
using F(4) finite_F
by (intro nn_integral_integrable_on) (auto split: split_indicator simp: enn2real_nonneg)
have "(\<integral>\<^sup>+ x. F i x * indicator (?B i) x \<partial>lborel) =
(\<integral>\<^sup>+ x. ennreal (enn2real (F i x) * indicator (?B i) x) \<partial>lborel)"
using F(3,4)
by (intro nn_integral_cong) (auto simp: image_iff eq_commute split: split_indicator)
also have "\<dots> = ennreal (integral UNIV (\<lambda>x. enn2real (F i x) * indicator (?B i) x))"
using F
by (intro nn_integral_lborel_eq_integral[OF _ _ finite_F])
(auto split: split_indicator intro: enn2real_nonneg)
also have "\<dots> \<le> ennreal I"
by (auto intro!: has_integral_le[OF integrable_integral[OF int_F] I] nonneg F_le_f
simp: \<open>0 \<le> I\<close> split: split_indicator )
finally show "(\<integral>\<^sup>+ x. F i x * indicator (?B i) x \<partial>lborel) \<le> ennreal I" .
qed
finally have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>lborel) < \<infinity>"
by (auto simp: less_top[symmetric] top_unique)
from nn_integral_lborel_eq_integral[OF assms(1,2) this] I show ?thesis
by (simp add: integral_unique)
qed
lemma has_integral_iff_emeasure_lborel:
fixes A :: "'a::euclidean_space set"
assumes A[measurable]: "A \<in> sets borel" and [simp]: "0 \<le> r"
shows "((\<lambda>x. 1) has_integral r) A \<longleftrightarrow> emeasure lborel A = ennreal r"
proof (cases "emeasure lborel A = \<infinity>")
case emeasure_A: True
have "\<not> (\<lambda>x. 1::real) integrable_on A"
proof
assume int: "(\<lambda>x. 1::real) integrable_on A"
then have "(indicator A::'a \<Rightarrow> real) integrable_on UNIV"
unfolding indicator_def[abs_def] integrable_restrict_UNIV .
then obtain r where "((indicator A::'a\<Rightarrow>real) has_integral r) UNIV"
by auto
from nn_integral_has_integral_lborel[OF _ _ this] emeasure_A show False
by (simp add: ennreal_indicator)
qed
with emeasure_A show ?thesis
by auto
next
case False
then have "((\<lambda>x. 1) has_integral measure lborel A) A"
by (simp add: has_integral_measure_lborel less_top)
with False show ?thesis
by (auto simp: emeasure_eq_ennreal_measure has_integral_unique)
qed
lemma ennreal_max_0: "ennreal (max 0 x) = ennreal x"
by (auto simp: max_def ennreal_neg)
lemma has_integral_integral_real:
fixes f::"'a::euclidean_space \<Rightarrow> real"
assumes f: "integrable lborel f"
shows "(f has_integral (integral\<^sup>L lborel f)) UNIV"
proof -
from integrableE[OF f] obtain r q
where "0 \<le> r" "0 \<le> q"
and r: "(\<integral>\<^sup>+ x. ennreal (max 0 (f x)) \<partial>lborel) = ennreal r"
and q: "(\<integral>\<^sup>+ x. ennreal (max 0 (- f x)) \<partial>lborel) = ennreal q"
and f: "f \<in> borel_measurable lborel" and eq: "integral\<^sup>L lborel f = r - q"
unfolding ennreal_max_0 by auto
then have "((\<lambda>x. max 0 (f x)) has_integral r) UNIV" "((\<lambda>x. max 0 (- f x)) has_integral q) UNIV"
using nn_integral_has_integral[OF _ _ r] nn_integral_has_integral[OF _ _ q] by auto
note has_integral_diff[OF this]
moreover have "(\<lambda>x. max 0 (f x) - max 0 (- f x)) = f"
by auto
ultimately show ?thesis
by (simp add: eq)
qed
lemma has_integral_AE:
assumes ae: "AE x in lborel. x \<in> \<Omega> \<longrightarrow> f x = g x"
shows "(f has_integral x) \<Omega> = (g has_integral x) \<Omega>"
proof -
from ae obtain N
where N: "N \<in> sets borel" "emeasure lborel N = 0" "{x. \<not> (x \<in> \<Omega> \<longrightarrow> f x = g x)} \<subseteq> N"
by (auto elim!: AE_E)
then have not_N: "AE x in lborel. x \<notin> N"
by (simp add: AE_iff_measurable)
show ?thesis
proof (rule has_integral_spike_eq[symmetric])
show "\<And>x. x\<in>\<Omega> - N \<Longrightarrow> f x = g x" using N(3) by auto
show "negligible N"
unfolding negligible_def
proof (intro allI)
fix a b :: "'a"
let ?F = "\<lambda>x::'a. if x \<in> cbox a b then indicator N x else 0 :: real"
have "integrable lborel ?F = integrable lborel (\<lambda>x::'a. 0::real)"
using not_N N(1) by (intro integrable_cong_AE) auto
moreover have "(LINT x|lborel. ?F x) = (LINT x::'a|lborel. 0::real)"
using not_N N(1) by (intro integral_cong_AE) auto
ultimately have "(?F has_integral 0) UNIV"
using has_integral_integral_real[of ?F] by simp
then show "(indicator N has_integral (0::real)) (cbox a b)"
unfolding has_integral_restrict_UNIV .
qed
qed
qed
lemma nn_integral_has_integral_lebesgue:
fixes f :: "'a::euclidean_space \<Rightarrow> real"
assumes nonneg: "\<And>x. 0 \<le> f x" and I: "(f has_integral I) \<Omega>"
shows "integral\<^sup>N lborel (\<lambda>x. indicator \<Omega> x * f x) = I"
proof -
from I have "(\<lambda>x. indicator \<Omega> x *\<^sub>R f x) \<in> lebesgue \<rightarrow>\<^sub>M borel"
by (rule has_integral_implies_lebesgue_measurable)
then obtain f' :: "'a \<Rightarrow> real"
where [measurable]: "f' \<in> borel \<rightarrow>\<^sub>M borel" and eq: "AE x in lborel. indicator \<Omega> x * f x = f' x"
by (auto dest: completion_ex_borel_measurable_real)
from I have "((\<lambda>x. abs (indicator \<Omega> x * f x)) has_integral I) UNIV"
using nonneg by (simp add: indicator_def if_distrib[of "\<lambda>x. x * f y" for y] cong: if_cong)
also have "((\<lambda>x. abs (indicator \<Omega> x * f x)) has_integral I) UNIV \<longleftrightarrow> ((\<lambda>x. abs (f' x)) has_integral I) UNIV"
using eq by (intro has_integral_AE) auto
finally have "integral\<^sup>N lborel (\<lambda>x. abs (f' x)) = I"
by (rule nn_integral_has_integral_lborel[rotated 2]) auto
also have "integral\<^sup>N lborel (\<lambda>x. abs (f' x)) = integral\<^sup>N lborel (\<lambda>x. abs (indicator \<Omega> x * f x))"
using eq by (intro nn_integral_cong_AE) auto
finally show ?thesis
using nonneg by auto
qed
lemma has_integral_iff_nn_integral_lebesgue:
assumes f: "\<And>x. 0 \<le> f x"
shows "(f has_integral r) UNIV \<longleftrightarrow> (f \<in> lebesgue \<rightarrow>\<^sub>M borel \<and> integral\<^sup>N lebesgue f = r \<and> 0 \<le> r)" (is "?I = ?N")
proof
assume ?I
have "0 \<le> r"
using has_integral_nonneg[OF \<open>?I\<close>] f by auto
then show ?N
using nn_integral_has_integral_lebesgue[OF f \<open>?I\<close>]
has_integral_implies_lebesgue_measurable[OF \<open>?I\<close>]
by (auto simp: nn_integral_completion)
next
assume ?N
then obtain f' where f': "f' \<in> borel \<rightarrow>\<^sub>M borel" "AE x in lborel. f x = f' x"
by (auto dest: completion_ex_borel_measurable_real)
moreover have "(\<integral>\<^sup>+ x. ennreal \<bar>f' x\<bar> \<partial>lborel) = (\<integral>\<^sup>+ x. ennreal \<bar>f x\<bar> \<partial>lborel)"
using f' by (intro nn_integral_cong_AE) auto
moreover have "((\<lambda>x. \<bar>f' x\<bar>) has_integral r) UNIV \<longleftrightarrow> ((\<lambda>x. \<bar>f x\<bar>) has_integral r) UNIV"
using f' by (intro has_integral_AE) auto
moreover note nn_integral_has_integral[of "\<lambda>x. \<bar>f' x\<bar>" r] \<open>?N\<close>
ultimately show ?I
using f by (auto simp: nn_integral_completion)
qed
context
fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
begin
lemma has_integral_integral_lborel:
assumes f: "integrable lborel f"
shows "(f has_integral (integral\<^sup>L lborel f)) UNIV"
proof -
have "((\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) has_integral (\<Sum>b\<in>Basis. integral\<^sup>L lborel (\<lambda>x. f x \<bullet> b) *\<^sub>R b)) UNIV"
using f by (intro has_integral_sum finite_Basis ballI has_integral_scaleR_left has_integral_integral_real) auto
also have eq_f: "(\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) = f"
by (simp add: fun_eq_iff euclidean_representation)
also have "(\<Sum>b\<in>Basis. integral\<^sup>L lborel (\<lambda>x. f x \<bullet> b) *\<^sub>R b) = integral\<^sup>L lborel f"
using f by (subst (2) eq_f[symmetric]) simp
finally show ?thesis .
qed
lemma integrable_on_lborel: "integrable lborel f \<Longrightarrow> f integrable_on UNIV"
using has_integral_integral_lborel by auto
lemma integral_lborel: "integrable lborel f \<Longrightarrow> integral UNIV f = (\<integral>x. f x \<partial>lborel)"
using has_integral_integral_lborel by auto
end
context
begin
private lemma has_integral_integral_lebesgue_real:
fixes f :: "'a::euclidean_space \<Rightarrow> real"
assumes f: "integrable lebesgue f"
shows "(f has_integral (integral\<^sup>L lebesgue f)) UNIV"
proof -
obtain f' where f': "f' \<in> borel \<rightarrow>\<^sub>M borel" "AE x in lborel. f x = f' x"
using completion_ex_borel_measurable_real[OF borel_measurable_integrable[OF f]] by auto
moreover have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>lborel) = (\<integral>\<^sup>+ x. ennreal (norm (f' x)) \<partial>lborel)"
using f' by (intro nn_integral_cong_AE) auto
ultimately have "integrable lborel f'"
using f by (auto simp: integrable_iff_bounded nn_integral_completion cong: nn_integral_cong_AE)
note has_integral_integral_real[OF this]
moreover have "integral\<^sup>L lebesgue f = integral\<^sup>L lebesgue f'"
using f' f by (intro integral_cong_AE) (auto intro: AE_completion measurable_completion)
moreover have "integral\<^sup>L lebesgue f' = integral\<^sup>L lborel f'"
using f' by (simp add: integral_completion)
moreover have "(f' has_integral integral\<^sup>L lborel f') UNIV \<longleftrightarrow> (f has_integral integral\<^sup>L lborel f') UNIV"
using f' by (intro has_integral_AE) auto
ultimately show ?thesis
by auto
qed
lemma has_integral_integral_lebesgue:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes f: "integrable lebesgue f"
shows "(f has_integral (integral\<^sup>L lebesgue f)) UNIV"
proof -
have "((\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) has_integral (\<Sum>b\<in>Basis. integral\<^sup>L lebesgue (\<lambda>x. f x \<bullet> b) *\<^sub>R b)) UNIV"
using f by (intro has_integral_sum finite_Basis ballI has_integral_scaleR_left has_integral_integral_lebesgue_real) auto
also have eq_f: "(\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) = f"
by (simp add: fun_eq_iff euclidean_representation)
also have "(\<Sum>b\<in>Basis. integral\<^sup>L lebesgue (\<lambda>x. f x \<bullet> b) *\<^sub>R b) = integral\<^sup>L lebesgue f"
using f by (subst (2) eq_f[symmetric]) simp
finally show ?thesis .
qed
lemma has_integral_integral_lebesgue_on:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "integrable (lebesgue_on S) f" "S \<in> sets lebesgue"
shows "(f has_integral (integral\<^sup>L (lebesgue_on S) f)) S"
proof -
let ?f = "\<lambda>x. if x \<in> S then f x else 0"
have "integrable lebesgue (\<lambda>x. indicat_real S x *\<^sub>R f x)"
using indicator_scaleR_eq_if [of S _ f] assms
by (metis (full_types) integrable_restrict_space sets.Int_space_eq2)
then have "integrable lebesgue ?f"
using indicator_scaleR_eq_if [of S _ f] assms by auto
then have "(?f has_integral (integral\<^sup>L lebesgue ?f)) UNIV"
by (rule has_integral_integral_lebesgue)
then have "(f has_integral (integral\<^sup>L lebesgue ?f)) S"
using has_integral_restrict_UNIV by blast
moreover
have "S \<inter> space lebesgue \<in> sets lebesgue"
by (simp add: assms)
then have "(integral\<^sup>L lebesgue ?f) = (integral\<^sup>L (lebesgue_on S) f)"
by (simp add: integral_restrict_space indicator_scaleR_eq_if)
ultimately show ?thesis
by auto
qed
lemma lebesgue_integral_eq_integral:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "integrable (lebesgue_on S) f" "S \<in> sets lebesgue"
shows "integral\<^sup>L (lebesgue_on S) f = integral S f"
by (metis has_integral_integral_lebesgue_on assms integral_unique)
lemma integrable_on_lebesgue:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
shows "integrable lebesgue f \<Longrightarrow> f integrable_on UNIV"
using has_integral_integral_lebesgue by auto
lemma integral_lebesgue:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
shows "integrable lebesgue f \<Longrightarrow> integral UNIV f = (\<integral>x. f x \<partial>lebesgue)"
using has_integral_integral_lebesgue by auto
end
subsection \<open>Absolute integrability (this is the same as Lebesgue integrability)\<close>
translations
"LBINT x. f" == "CONST lebesgue_integral CONST lborel (\<lambda>x. f)"
translations
"LBINT x:A. f" == "CONST set_lebesgue_integral CONST lborel A (\<lambda>x. f)"
lemma set_integral_reflect:
fixes S and f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
shows "(LBINT x : S. f x) = (LBINT x : {x. - x \<in> S}. f (- x))"
unfolding set_lebesgue_integral_def
by (subst lborel_integral_real_affine[where c="-1" and t=0])
(auto intro!: Bochner_Integration.integral_cong split: split_indicator)
lemma borel_integrable_atLeastAtMost':
fixes f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}"
assumes f: "continuous_on {a..b} f"
shows "set_integrable lborel {a..b} f"
unfolding set_integrable_def
by (intro borel_integrable_compact compact_Icc f)
lemma integral_FTC_atLeastAtMost:
fixes f :: "real \<Rightarrow> 'a :: euclidean_space"
assumes "a \<le> b"
and F: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (F has_vector_derivative f x) (at x within {a .. b})"
and f: "continuous_on {a .. b} f"
shows "integral\<^sup>L lborel (\<lambda>x. indicator {a .. b} x *\<^sub>R f x) = F b - F a"
proof -
let ?f = "\<lambda>x. indicator {a .. b} x *\<^sub>R f x"
have "(?f has_integral (\<integral>x. ?f x \<partial>lborel)) UNIV"
using borel_integrable_atLeastAtMost'[OF f]
unfolding set_integrable_def by (rule has_integral_integral_lborel)
moreover
have "(f has_integral F b - F a) {a .. b}"
by (intro fundamental_theorem_of_calculus ballI assms) auto
then have "(?f has_integral F b - F a) {a .. b}"
by (subst has_integral_cong[where g=f]) auto
then have "(?f has_integral F b - F a) UNIV"
by (intro has_integral_on_superset[where T=UNIV and S="{a..b}"]) auto
ultimately show "integral\<^sup>L lborel ?f = F b - F a"
by (rule has_integral_unique)
qed
lemma set_borel_integral_eq_integral:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "set_integrable lborel S f"
shows "f integrable_on S" "LINT x : S | lborel. f x = integral S f"
proof -
let ?f = "\<lambda>x. indicator S x *\<^sub>R f x"
have "(?f has_integral LINT x : S | lborel. f x) UNIV"
using assms has_integral_integral_lborel
unfolding set_integrable_def set_lebesgue_integral_def by blast
hence 1: "(f has_integral (set_lebesgue_integral lborel S f)) S"
apply (subst has_integral_restrict_UNIV [symmetric])
apply (rule has_integral_eq)
by auto
thus "f integrable_on S"
by (auto simp add: integrable_on_def)
with 1 have "(f has_integral (integral S f)) S"
by (intro integrable_integral, auto simp add: integrable_on_def)
thus "LINT x : S | lborel. f x = integral S f"
by (intro has_integral_unique [OF 1])
qed
lemma has_integral_set_lebesgue:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes f: "set_integrable lebesgue S f"
shows "(f has_integral (LINT x:S|lebesgue. f x)) S"
using has_integral_integral_lebesgue f
by (fastforce simp add: set_integrable_def set_lebesgue_integral_def indicator_def if_distrib[of "\<lambda>x. x *\<^sub>R f _"] cong: if_cong)
lemma set_lebesgue_integral_eq_integral:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes f: "set_integrable lebesgue S f"
shows "f integrable_on S" "LINT x:S | lebesgue. f x = integral S f"
using has_integral_set_lebesgue[OF f] by (auto simp: integral_unique integrable_on_def)
lemma lmeasurable_iff_has_integral:
"S \<in> lmeasurable \<longleftrightarrow> ((indicator S) has_integral measure lebesgue S) UNIV"
by (subst has_integral_iff_nn_integral_lebesgue)
(auto simp: ennreal_indicator emeasure_eq_measure2 borel_measurable_indicator_iff intro!: fmeasurableI)
abbreviation
absolutely_integrable_on :: "('a::euclidean_space \<Rightarrow> 'b::{banach, second_countable_topology}) \<Rightarrow> 'a set \<Rightarrow> bool"
(infixr "absolutely'_integrable'_on" 46)
where "f absolutely_integrable_on s \<equiv> set_integrable lebesgue s f"
lemma absolutely_integrable_zero [simp]: "(\<lambda>x. 0) absolutely_integrable_on S"
by (simp add: set_integrable_def)
lemma absolutely_integrable_on_def:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
shows "f absolutely_integrable_on S \<longleftrightarrow> f integrable_on S \<and> (\<lambda>x. norm (f x)) integrable_on S"
proof safe
assume f: "f absolutely_integrable_on S"
then have nf: "integrable lebesgue (\<lambda>x. norm (indicator S x *\<^sub>R f x))"
using integrable_norm set_integrable_def by blast
show "f integrable_on S"
by (rule set_lebesgue_integral_eq_integral[OF f])
have "(\<lambda>x. norm (indicator S x *\<^sub>R f x)) = (\<lambda>x. if x \<in> S then norm (f x) else 0)"
by auto
with integrable_on_lebesgue[OF nf] show "(\<lambda>x. norm (f x)) integrable_on S"
by (simp add: integrable_restrict_UNIV)
next
assume f: "f integrable_on S" and nf: "(\<lambda>x. norm (f x)) integrable_on S"
show "f absolutely_integrable_on S"
unfolding set_integrable_def
proof (rule integrableI_bounded)
show "(\<lambda>x. indicator S x *\<^sub>R f x) \<in> borel_measurable lebesgue"
using f has_integral_implies_lebesgue_measurable[of f _ S] by (auto simp: integrable_on_def)
show "(\<integral>\<^sup>+ x. ennreal (norm (indicator S x *\<^sub>R f x)) \<partial>lebesgue) < \<infinity>"
using nf nn_integral_has_integral_lebesgue[of "\<lambda>x. norm (f x)" _ S]
by (auto simp: integrable_on_def nn_integral_completion)
qed
qed
lemma integrable_on_lebesgue_on:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes f: "integrable (lebesgue_on S) f" and S: "S \<in> sets lebesgue"
shows "f integrable_on S"
proof -
have "integrable lebesgue (\<lambda>x. indicator S x *\<^sub>R f x)"
using S f inf_top.comm_neutral integrable_restrict_space by blast
then show ?thesis
using absolutely_integrable_on_def set_integrable_def by blast
qed
lemma absolutely_integrable_imp_integrable:
assumes "f absolutely_integrable_on S" "S \<in> sets lebesgue"
shows "integrable (lebesgue_on S) f"
by (meson assms integrable_restrict_space set_integrable_def sets.Int sets.top)
lemma absolutely_integrable_on_null [intro]:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
shows "content (cbox a b) = 0 \<Longrightarrow> f absolutely_integrable_on (cbox a b)"
by (auto simp: absolutely_integrable_on_def)
lemma absolutely_integrable_on_open_interval:
fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
shows "f absolutely_integrable_on box a b \<longleftrightarrow>
f absolutely_integrable_on cbox a b"
by (auto simp: integrable_on_open_interval absolutely_integrable_on_def)
lemma absolutely_integrable_restrict_UNIV:
"(\<lambda>x. if x \<in> S then f x else 0) absolutely_integrable_on UNIV \<longleftrightarrow> f absolutely_integrable_on S"
unfolding set_integrable_def
by (intro arg_cong2[where f=integrable]) auto
lemma absolutely_integrable_onI:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
shows "f integrable_on S \<Longrightarrow> (\<lambda>x. norm (f x)) integrable_on S \<Longrightarrow> f absolutely_integrable_on S"
unfolding absolutely_integrable_on_def by auto
lemma nonnegative_absolutely_integrable_1:
fixes f :: "'a :: euclidean_space \<Rightarrow> real"
assumes f: "f integrable_on A" and "\<And>x. x \<in> A \<Longrightarrow> 0 \<le> f x"
shows "f absolutely_integrable_on A"
by (rule absolutely_integrable_onI [OF f]) (use assms in \<open>simp add: integrable_eq\<close>)
lemma absolutely_integrable_on_iff_nonneg:
fixes f :: "'a :: euclidean_space \<Rightarrow> real"
assumes "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f x" shows "f absolutely_integrable_on S \<longleftrightarrow> f integrable_on S"
proof -
{ assume "f integrable_on S"
then have "(\<lambda>x. if x \<in> S then f x else 0) integrable_on UNIV"
by (simp add: integrable_restrict_UNIV)
then have "(\<lambda>x. if x \<in> S then f x else 0) absolutely_integrable_on UNIV"
using \<open>f integrable_on S\<close> absolutely_integrable_restrict_UNIV assms nonnegative_absolutely_integrable_1 by blast
then have "f absolutely_integrable_on S"
using absolutely_integrable_restrict_UNIV by blast
}
then show ?thesis
unfolding absolutely_integrable_on_def by auto
qed
lemma absolutely_integrable_on_scaleR_iff:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
shows
"(\<lambda>x. c *\<^sub>R f x) absolutely_integrable_on S \<longleftrightarrow>
c = 0 \<or> f absolutely_integrable_on S"
proof (cases "c=0")
case False
then show ?thesis
unfolding absolutely_integrable_on_def
by (simp add: norm_mult)
qed auto
lemma absolutely_integrable_spike:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "f absolutely_integrable_on T" and S: "negligible S" "\<And>x. x \<in> T - S \<Longrightarrow> g x = f x"
shows "g absolutely_integrable_on T"
using assms integrable_spike
unfolding absolutely_integrable_on_def by metis
lemma absolutely_integrable_negligible:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "negligible S"
shows "f absolutely_integrable_on S"
using assms by (simp add: absolutely_integrable_on_def integrable_negligible)
lemma absolutely_integrable_spike_eq:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "negligible S" "\<And>x. x \<in> T - S \<Longrightarrow> g x = f x"
shows "(f absolutely_integrable_on T \<longleftrightarrow> g absolutely_integrable_on T)"
using assms by (blast intro: absolutely_integrable_spike sym)
lemma absolutely_integrable_spike_set_eq:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "negligible {x \<in> S - T. f x \<noteq> 0}" "negligible {x \<in> T - S. f x \<noteq> 0}"
shows "(f absolutely_integrable_on S \<longleftrightarrow> f absolutely_integrable_on T)"
proof -
have "(\<lambda>x. if x \<in> S then f x else 0) absolutely_integrable_on UNIV \<longleftrightarrow>
(\<lambda>x. if x \<in> T then f x else 0) absolutely_integrable_on UNIV"
proof (rule absolutely_integrable_spike_eq)
show "negligible ({x \<in> S - T. f x \<noteq> 0} \<union> {x \<in> T - S. f x \<noteq> 0})"
by (rule negligible_Un [OF assms])
qed auto
with absolutely_integrable_restrict_UNIV show ?thesis
by blast
qed
lemma absolutely_integrable_spike_set:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes f: "f absolutely_integrable_on S" and neg: "negligible {x \<in> S - T. f x \<noteq> 0}" "negligible {x \<in> T - S. f x \<noteq> 0}"
shows "f absolutely_integrable_on T"
using absolutely_integrable_spike_set_eq f neg by blast
lemma absolutely_integrable_reflect[simp]:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
shows "(\<lambda>x. f(-x)) absolutely_integrable_on cbox (-b) (-a) \<longleftrightarrow> f absolutely_integrable_on cbox a b"
apply (auto simp: absolutely_integrable_on_def dest: integrable_reflect [THEN iffD1])
unfolding integrable_on_def by auto
lemma absolutely_integrable_reflect_real[simp]:
fixes f :: "real \<Rightarrow> 'b::euclidean_space"
shows "(\<lambda>x. f(-x)) absolutely_integrable_on {-b .. -a} \<longleftrightarrow> f absolutely_integrable_on {a..b::real}"
unfolding box_real[symmetric] by (rule absolutely_integrable_reflect)
lemma absolutely_integrable_on_subcbox:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
shows "\<lbrakk>f absolutely_integrable_on S; cbox a b \<subseteq> S\<rbrakk> \<Longrightarrow> f absolutely_integrable_on cbox a b"
by (meson absolutely_integrable_on_def integrable_on_subcbox)
lemma absolutely_integrable_on_subinterval:
fixes f :: "real \<Rightarrow> 'b::euclidean_space"
shows "\<lbrakk>f absolutely_integrable_on S; {a..b} \<subseteq> S\<rbrakk> \<Longrightarrow> f absolutely_integrable_on {a..b}"
using absolutely_integrable_on_subcbox by fastforce
lemma integrable_subinterval:
fixes f :: "real \<Rightarrow> 'a::euclidean_space"
assumes "integrable (lebesgue_on {a..b}) f"
and "{c..d} \<subseteq> {a..b}"
shows "integrable (lebesgue_on {c..d}) f"
proof (rule absolutely_integrable_imp_integrable)
show "f absolutely_integrable_on {c..d}"
proof -
have "f integrable_on {c..d}"
using assms integrable_on_lebesgue_on integrable_on_subinterval by fastforce
moreover have "(\<lambda>x. norm (f x)) integrable_on {c..d}"
proof (rule integrable_on_subinterval)
show "(\<lambda>x. norm (f x)) integrable_on {a..b}"
by (simp add: assms integrable_on_lebesgue_on)
qed (use assms in auto)
ultimately show ?thesis
by (auto simp: absolutely_integrable_on_def)
qed
qed auto
lemma indefinite_integral_continuous_real:
fixes f :: "real \<Rightarrow> 'a::euclidean_space"
assumes "integrable (lebesgue_on {a..b}) f"
shows "continuous_on {a..b} (\<lambda>x. integral\<^sup>L (lebesgue_on {a..x}) f)"
proof -
have "f integrable_on {a..b}"
by (simp add: assms integrable_on_lebesgue_on)
then have "continuous_on {a..b} (\<lambda>x. integral {a..x} f)"
using indefinite_integral_continuous_1 by blast
moreover have "integral\<^sup>L (lebesgue_on {a..x}) f = integral {a..x} f" if "a \<le> x" "x \<le> b" for x
proof -
have "{a..x} \<subseteq> {a..b}"
using that by auto
then have "integrable (lebesgue_on {a..x}) f"
using integrable_subinterval assms by blast
then show "integral\<^sup>L (lebesgue_on {a..x}) f = integral {a..x} f"
by (simp add: lebesgue_integral_eq_integral)
qed
ultimately show ?thesis
by (metis (no_types, lifting) atLeastAtMost_iff continuous_on_cong)
qed
lemma lmeasurable_iff_integrable_on: "S \<in> lmeasurable \<longleftrightarrow> (\<lambda>x. 1::real) integrable_on S"
by (subst absolutely_integrable_on_iff_nonneg[symmetric])
(simp_all add: lmeasurable_iff_integrable set_integrable_def)
lemma lmeasure_integral_UNIV: "S \<in> lmeasurable \<Longrightarrow> measure lebesgue S = integral UNIV (indicator S)"
by (simp add: lmeasurable_iff_has_integral integral_unique)
lemma lmeasure_integral: "S \<in> lmeasurable \<Longrightarrow> measure lebesgue S = integral S (\<lambda>x. 1::real)"
by (fastforce simp add: lmeasure_integral_UNIV indicator_def[abs_def] lmeasurable_iff_integrable_on)
lemma integrable_on_const: "S \<in> lmeasurable \<Longrightarrow> (\<lambda>x. c) integrable_on S"
unfolding lmeasurable_iff_integrable
by (metis (mono_tags, lifting) integrable_eq integrable_on_scaleR_left lmeasurable_iff_integrable lmeasurable_iff_integrable_on scaleR_one)
lemma integral_indicator:
assumes "(S \<inter> T) \<in> lmeasurable"
shows "integral T (indicator S) = measure lebesgue (S \<inter> T)"
proof -
have "integral UNIV (indicator (S \<inter> T)) = integral UNIV (\<lambda>a. if a \<in> S \<inter> T then 1::real else 0)"
by (meson indicator_def)
moreover
have "(indicator (S \<inter> T) has_integral measure lebesgue (S \<inter> T)) UNIV"
using assms by (simp add: lmeasurable_iff_has_integral)
ultimately have "integral UNIV (\<lambda>x. if x \<in> S \<inter> T then 1 else 0) = measure lebesgue (S \<inter> T)"
by (metis (no_types) integral_unique)
then show ?thesis
using integral_restrict_Int [of UNIV "S \<inter> T" "\<lambda>x. 1::real"]
apply (simp add: integral_restrict_Int [symmetric])
by (meson indicator_def)
qed
lemma measurable_integrable:
fixes S :: "'a::euclidean_space set"
shows "S \<in> lmeasurable \<longleftrightarrow> (indicat_real S) integrable_on UNIV"
by (auto simp: lmeasurable_iff_integrable absolutely_integrable_on_iff_nonneg [symmetric] set_integrable_def)
lemma integrable_on_indicator:
fixes S :: "'a::euclidean_space set"
shows "indicat_real S integrable_on T \<longleftrightarrow> (S \<inter> T) \<in> lmeasurable"
unfolding measurable_integrable
unfolding integrable_restrict_UNIV [of T, symmetric]
by (fastforce simp add: indicator_def elim: integrable_eq)
lemma
assumes \<D>: "\<D> division_of S"
shows lmeasurable_division: "S \<in> lmeasurable" (is ?l)
and content_division: "(\<Sum>k\<in>\<D>. measure lebesgue k) = measure lebesgue S" (is ?m)
proof -
{ fix d1 d2 assume *: "d1 \<in> \<D>" "d2 \<in> \<D>" "d1 \<noteq> d2"
then obtain a b c d where "d1 = cbox a b" "d2 = cbox c d"
using division_ofD(4)[OF \<D>] by blast
with division_ofD(5)[OF \<D> *]
have "d1 \<in> sets lborel" "d2 \<in> sets lborel" "d1 \<inter> d2 \<subseteq> (cbox a b - box a b) \<union> (cbox c d - box c d)"
by auto
moreover have "(cbox a b - box a b) \<union> (cbox c d - box c d) \<in> null_sets lborel"
by (intro null_sets.Un null_sets_cbox_Diff_box)
ultimately have "d1 \<inter> d2 \<in> null_sets lborel"
by (blast intro: null_sets_subset) }
then show ?l ?m
unfolding division_ofD(6)[OF \<D>, symmetric]
using division_ofD(1,4)[OF \<D>]
by (auto intro!: measure_Union_AE[symmetric] simp: completion.AE_iff_null_sets Int_def[symmetric] pairwise_def null_sets_def)
qed
lemma has_measure_limit:
assumes "S \<in> lmeasurable" "e > 0"
obtains B where "B > 0"
"\<And>a b. ball 0 B \<subseteq> cbox a b \<Longrightarrow> \<bar>measure lebesgue (S \<inter> cbox a b) - measure lebesgue S\<bar> < e"
using assms unfolding lmeasurable_iff_has_integral has_integral_alt'
by (force simp: integral_indicator integrable_on_indicator)
lemma lmeasurable_iff_indicator_has_integral:
fixes S :: "'a::euclidean_space set"
shows "S \<in> lmeasurable \<and> m = measure lebesgue S \<longleftrightarrow> (indicat_real S has_integral m) UNIV"
by (metis has_integral_iff lmeasurable_iff_has_integral measurable_integrable)
lemma has_measure_limit_iff:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
shows "S \<in> lmeasurable \<and> m = measure lebesgue S \<longleftrightarrow>
(\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
(S \<inter> cbox a b) \<in> lmeasurable \<and> \<bar>measure lebesgue (S \<inter> cbox a b) - m\<bar> < e)" (is "?lhs = ?rhs")
proof
assume ?lhs then show ?rhs
by (meson has_measure_limit fmeasurable.Int lmeasurable_cbox)
next
assume RHS [rule_format]: ?rhs
show ?lhs
apply (simp add: lmeasurable_iff_indicator_has_integral has_integral' [where i=m])
using RHS
by (metis (full_types) integral_indicator integrable_integral integrable_on_indicator)
qed
subsection\<open>Applications to Negligibility\<close>
lemma negligible_iff_null_sets: "negligible S \<longleftrightarrow> S \<in> null_sets lebesgue"
proof
assume "negligible S"
then have "(indicator S has_integral (0::real)) UNIV"
by (auto simp: negligible)
then show "S \<in> null_sets lebesgue"
by (subst (asm) has_integral_iff_nn_integral_lebesgue)
(auto simp: borel_measurable_indicator_iff nn_integral_0_iff_AE AE_iff_null_sets indicator_eq_0_iff)
next
assume S: "S \<in> null_sets lebesgue"
show "negligible S"
unfolding negligible_def
proof (safe intro!: has_integral_iff_nn_integral_lebesgue[THEN iffD2]
has_integral_restrict_UNIV[where s="cbox _ _", THEN iffD1])
fix a b
show "(\<lambda>x. if x \<in> cbox a b then indicator S x else 0) \<in> lebesgue \<rightarrow>\<^sub>M borel"
using S by (auto intro!: measurable_If)
then show "(\<integral>\<^sup>+ x. ennreal (if x \<in> cbox a b then indicator S x else 0) \<partial>lebesgue) = ennreal 0"
using S[THEN AE_not_in] by (auto intro!: nn_integral_0_iff_AE[THEN iffD2])
qed auto
qed
corollary eventually_ae_filter_negligible:
"eventually P (ae_filter lebesgue) \<longleftrightarrow> (\<exists>N. negligible N \<and> {x. \<not> P x} \<subseteq> N)"
by (auto simp: completion.AE_iff_null_sets negligible_iff_null_sets null_sets_completion_subset)
lemma starlike_negligible:
assumes "closed S"
and eq1: "\<And>c x. (a + c *\<^sub>R x) \<in> S \<Longrightarrow> 0 \<le> c \<Longrightarrow> a + x \<in> S \<Longrightarrow> c = 1"
shows "negligible S"
proof -
have "negligible ((+) (-a) ` S)"
proof (subst negligible_on_intervals, intro allI)
fix u v
show "negligible ((+) (- a) ` S \<inter> cbox u v)"
using \<open>closed S\<close> eq1 by (auto simp add: negligible_iff_null_sets algebra_simps
intro!: closed_translation_subtract starlike_negligible_compact cong: image_cong_simp)
(metis add_diff_eq diff_add_cancel scale_right_diff_distrib)
qed
then show ?thesis
by (rule negligible_translation_rev)
qed
lemma starlike_negligible_strong:
assumes "closed S"
and star: "\<And>c x. \<lbrakk>0 \<le> c; c < 1; a+x \<in> S\<rbrakk> \<Longrightarrow> a + c *\<^sub>R x \<notin> S"
shows "negligible S"
proof -
show ?thesis
proof (rule starlike_negligible [OF \<open>closed S\<close>, of a])
fix c x
assume cx: "a + c *\<^sub>R x \<in> S" "0 \<le> c" "a + x \<in> S"
with star have "\<not> (c < 1)" by auto
moreover have "\<not> (c > 1)"
using star [of "1/c" "c *\<^sub>R x"] cx by force
ultimately show "c = 1" by arith
qed
qed
lemma negligible_hyperplane:
assumes "a \<noteq> 0 \<or> b \<noteq> 0" shows "negligible {x. a \<bullet> x = b}"
proof -
obtain x where x: "a \<bullet> x \<noteq> b"
using assms
apply auto
apply (metis inner_eq_zero_iff inner_zero_right)
using inner_zero_right by fastforce
show ?thesis
apply (rule starlike_negligible [OF closed_hyperplane, of x])
using x apply (auto simp: algebra_simps)
done
qed
lemma negligible_lowdim:
fixes S :: "'N :: euclidean_space set"
assumes "dim S < DIM('N)"
shows "negligible S"
proof -
obtain a where "a \<noteq> 0" and a: "span S \<subseteq> {x. a \<bullet> x = 0}"
using lowdim_subset_hyperplane [OF assms] by blast
have "negligible (span S)"
using \<open>a \<noteq> 0\<close> a negligible_hyperplane by (blast intro: negligible_subset)
then show ?thesis
using span_base by (blast intro: negligible_subset)
qed
proposition negligible_convex_frontier:
fixes S :: "'N :: euclidean_space set"
assumes "convex S"
shows "negligible(frontier S)"
proof -
have nf: "negligible(frontier S)" if "convex S" "0 \<in> S" for S :: "'N set"
proof -
obtain B where "B \<subseteq> S" and indB: "independent B"
and spanB: "S \<subseteq> span B" and cardB: "card B = dim S"
by (metis basis_exists)
consider "dim S < DIM('N)" | "dim S = DIM('N)"
using dim_subset_UNIV le_eq_less_or_eq by auto
then show ?thesis
proof cases
case 1
show ?thesis
by (rule negligible_subset [of "closure S"])
(simp_all add: frontier_def negligible_lowdim 1)
next
case 2
obtain a where a: "a \<in> interior S"
apply (rule interior_simplex_nonempty [OF indB])
apply (simp add: indB independent_finite)
apply (simp add: cardB 2)
apply (metis \<open>B \<subseteq> S\<close> \<open>0 \<in> S\<close> \<open>convex S\<close> insert_absorb insert_subset interior_mono subset_hull)
done
show ?thesis
proof (rule starlike_negligible_strong [where a=a])
fix c::real and x
have eq: "a + c *\<^sub>R x = (a + x) - (1 - c) *\<^sub>R ((a + x) - a)"
by (simp add: algebra_simps)
assume "0 \<le> c" "c < 1" "a + x \<in> frontier S"
then show "a + c *\<^sub>R x \<notin> frontier S"
apply (clarsimp simp: frontier_def)
apply (subst eq)
apply (rule mem_interior_closure_convex_shrink [OF \<open>convex S\<close> a, of _ "1-c"], auto)
done
qed auto
qed
qed
show ?thesis
proof (cases "S = {}")
case True then show ?thesis by auto
next
case False
then obtain a where "a \<in> S" by auto
show ?thesis
using nf [of "(\<lambda>x. -a + x) ` S"]
by (metis \<open>a \<in> S\<close> add.left_inverse assms convex_translation_eq frontier_translation
image_eqI negligible_translation_rev)
qed
qed
corollary negligible_sphere: "negligible (sphere a e)"
using frontier_cball negligible_convex_frontier convex_cball
by (blast intro: negligible_subset)
lemma non_negligible_UNIV [simp]: "\<not> negligible UNIV"
unfolding negligible_iff_null_sets by (auto simp: null_sets_def)
lemma negligible_interval:
"negligible (cbox a b) \<longleftrightarrow> box a b = {}" "negligible (box a b) \<longleftrightarrow> box a b = {}"
by (auto simp: negligible_iff_null_sets null_sets_def prod_nonneg inner_diff_left box_eq_empty
not_le emeasure_lborel_cbox_eq emeasure_lborel_box_eq
intro: eq_refl antisym less_imp_le)
proposition open_not_negligible:
assumes "open S" "S \<noteq> {}"
shows "\<not> negligible S"
proof
assume neg: "negligible S"
obtain a where "a \<in> S"
using \<open>S \<noteq> {}\<close> by blast
then obtain e where "e > 0" "cball a e \<subseteq> S"
using \<open>open S\<close> open_contains_cball_eq by blast
let ?p = "a - (e / DIM('a)) *\<^sub>R One" let ?q = "a + (e / DIM('a)) *\<^sub>R One"
have "cbox ?p ?q \<subseteq> cball a e"
proof (clarsimp simp: mem_box dist_norm)
fix x
assume "\<forall>i\<in>Basis. ?p \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> ?q \<bullet> i"
then have ax: "\<bar>(a - x) \<bullet> i\<bar> \<le> e / real DIM('a)" if "i \<in> Basis" for i
using that by (auto simp: algebra_simps)
have "norm (a - x) \<le> (\<Sum>i\<in>Basis. \<bar>(a - x) \<bullet> i\<bar>)"
by (rule norm_le_l1)
also have "\<dots> \<le> DIM('a) * (e / real DIM('a))"
by (intro sum_bounded_above ax)
also have "\<dots> = e"
by simp
finally show "norm (a - x) \<le> e" .
qed
then have "negligible (cbox ?p ?q)"
by (meson \<open>cball a e \<subseteq> S\<close> neg negligible_subset)
with \<open>e > 0\<close> show False
by (simp add: negligible_interval box_eq_empty algebra_simps field_split_simps mult_le_0_iff)
qed
lemma negligible_convex_interior:
"convex S \<Longrightarrow> (negligible S \<longleftrightarrow> interior S = {})"
apply safe
apply (metis interior_subset negligible_subset open_interior open_not_negligible)
apply (metis Diff_empty closure_subset frontier_def negligible_convex_frontier negligible_subset)
done
lemma measure_eq_0_null_sets: "S \<in> null_sets M \<Longrightarrow> measure M S = 0"
by (auto simp: measure_def null_sets_def)
lemma negligible_imp_measure0: "negligible S \<Longrightarrow> measure lebesgue S = 0"
by (simp add: measure_eq_0_null_sets negligible_iff_null_sets)
lemma negligible_iff_emeasure0: "S \<in> sets lebesgue \<Longrightarrow> (negligible S \<longleftrightarrow> emeasure lebesgue S = 0)"
by (auto simp: measure_eq_0_null_sets negligible_iff_null_sets)
lemma negligible_iff_measure0: "S \<in> lmeasurable \<Longrightarrow> (negligible S \<longleftrightarrow> measure lebesgue S = 0)"
apply (auto simp: measure_eq_0_null_sets negligible_iff_null_sets)
by (metis completion.null_sets_outer subsetI)
lemma negligible_imp_sets: "negligible S \<Longrightarrow> S \<in> sets lebesgue"
by (simp add: negligible_iff_null_sets null_setsD2)
lemma negligible_imp_measurable: "negligible S \<Longrightarrow> S \<in> lmeasurable"
by (simp add: fmeasurableI_null_sets negligible_iff_null_sets)
lemma negligible_iff_measure: "negligible S \<longleftrightarrow> S \<in> lmeasurable \<and> measure lebesgue S = 0"
by (fastforce simp: negligible_iff_measure0 negligible_imp_measurable dest: negligible_imp_measure0)
lemma negligible_outer:
"negligible S \<longleftrightarrow> (\<forall>e>0. \<exists>T. S \<subseteq> T \<and> T \<in> lmeasurable \<and> measure lebesgue T < e)" (is "_ = ?rhs")
proof
assume "negligible S" then show ?rhs
by (metis negligible_iff_measure order_refl)
next
assume ?rhs then show "negligible S"
by (meson completion.null_sets_outer negligible_iff_null_sets)
qed
lemma negligible_outer_le:
"negligible S \<longleftrightarrow> (\<forall>e>0. \<exists>T. S \<subseteq> T \<and> T \<in> lmeasurable \<and> measure lebesgue T \<le> e)" (is "_ = ?rhs")
proof
assume "negligible S" then show ?rhs
by (metis dual_order.strict_implies_order negligible_imp_measurable negligible_imp_measure0 order_refl)
next
assume ?rhs then show "negligible S"
by (metis le_less_trans negligible_outer field_lbound_gt_zero)
qed
lemma negligible_UNIV: "negligible S \<longleftrightarrow> (indicat_real S has_integral 0) UNIV" (is "_=?rhs")
proof
assume ?rhs
then show "negligible S"
apply (auto simp: negligible_def has_integral_iff integrable_on_indicator)
by (metis negligible integral_unique lmeasure_integral_UNIV negligible_iff_measure0)
qed (simp add: negligible)
lemma sets_negligible_symdiff:
"\<lbrakk>S \<in> sets lebesgue; negligible((S - T) \<union> (T - S))\<rbrakk> \<Longrightarrow> T \<in> sets lebesgue"
by (metis Diff_Diff_Int Int_Diff_Un inf_commute negligible_Un_eq negligible_imp_sets sets.Diff sets.Un)
lemma lmeasurable_negligible_symdiff:
"\<lbrakk>S \<in> lmeasurable; negligible((S - T) \<union> (T - S))\<rbrakk> \<Longrightarrow> T \<in> lmeasurable"
using integrable_spike_set_eq lmeasurable_iff_integrable_on by blast
lemma measure_Un3_negligible:
assumes meas: "S \<in> lmeasurable" "T \<in> lmeasurable" "U \<in> lmeasurable"
and neg: "negligible(S \<inter> T)" "negligible(S \<inter> U)" "negligible(T \<inter> U)" and V: "S \<union> T \<union> U = V"
shows "measure lebesgue V = measure lebesgue S + measure lebesgue T + measure lebesgue U"
proof -
have [simp]: "measure lebesgue (S \<inter> T) = 0"
using neg(1) negligible_imp_measure0 by blast
have [simp]: "measure lebesgue (S \<inter> U \<union> T \<inter> U) = 0"
using neg(2) neg(3) negligible_Un negligible_imp_measure0 by blast
have "measure lebesgue V = measure lebesgue (S \<union> T \<union> U)"
using V by simp
also have "\<dots> = measure lebesgue S + measure lebesgue T + measure lebesgue U"
by (simp add: measure_Un3 meas fmeasurable.Un Int_Un_distrib2)
finally show ?thesis .
qed
lemma measure_translate_add:
assumes meas: "S \<in> lmeasurable" "T \<in> lmeasurable"
and U: "S \<union> ((+)a ` T) = U" and neg: "negligible(S \<inter> ((+)a ` T))"
shows "measure lebesgue S + measure lebesgue T = measure lebesgue U"
proof -
have [simp]: "measure lebesgue (S \<inter> (+) a ` T) = 0"
using neg negligible_imp_measure0 by blast
have "measure lebesgue (S \<union> ((+)a ` T)) = measure lebesgue S + measure lebesgue T"
by (simp add: measure_Un3 meas measurable_translation measure_translation fmeasurable.Un)
then show ?thesis
using U by auto
qed
lemma measure_negligible_symdiff:
assumes S: "S \<in> lmeasurable"
and neg: "negligible (S - T \<union> (T - S))"
shows "measure lebesgue T = measure lebesgue S"
proof -
have "measure lebesgue (S - T) = 0"
using neg negligible_Un_eq negligible_imp_measure0 by blast
then show ?thesis
by (metis S Un_commute add.right_neutral lmeasurable_negligible_symdiff measure_Un2 neg negligible_Un_eq negligible_imp_measure0)
qed
lemma measure_closure:
assumes "bounded S" and neg: "negligible (frontier S)"
shows "measure lebesgue (closure S) = measure lebesgue S"
proof -
have "measure lebesgue (frontier S) = 0"
by (metis neg negligible_imp_measure0)
then show ?thesis
by (metis assms lmeasurable_iff_integrable_on eq_iff_diff_eq_0 has_integral_interior integrable_on_def integral_unique lmeasurable_interior lmeasure_integral measure_frontier)
qed
lemma measure_interior:
"\<lbrakk>bounded S; negligible(frontier S)\<rbrakk> \<Longrightarrow> measure lebesgue (interior S) = measure lebesgue S"
using measure_closure measure_frontier negligible_imp_measure0 by fastforce
lemma measurable_Jordan:
assumes "bounded S" and neg: "negligible (frontier S)"
shows "S \<in> lmeasurable"
proof -
have "closure S \<in> lmeasurable"
by (metis lmeasurable_closure \<open>bounded S\<close>)
moreover have "interior S \<in> lmeasurable"
by (simp add: lmeasurable_interior \<open>bounded S\<close>)
moreover have "interior S \<subseteq> S"
by (simp add: interior_subset)
ultimately show ?thesis
using assms by (metis (full_types) closure_subset completion.complete_sets_sandwich_fmeasurable measure_closure measure_interior)
qed
lemma measurable_convex: "\<lbrakk>convex S; bounded S\<rbrakk> \<Longrightarrow> S \<in> lmeasurable"
by (simp add: measurable_Jordan negligible_convex_frontier)
lemma content_cball_conv_ball: "content (cball c r) = content (ball c r)"
proof -
have "ball c r - cball c r \<union> (cball c r - ball c r) = sphere c r"
by auto
hence "measure lebesgue (cball c r) = measure lebesgue (ball c r)"
using negligible_sphere[of c r] by (intro measure_negligible_symdiff) simp_all
thus ?thesis by simp
qed
subsection\<open>Negligibility of image under non-injective linear map\<close>
lemma negligible_Union_nat:
assumes "\<And>n::nat. negligible(S n)"
shows "negligible(\<Union>n. S n)"
proof -
have "negligible (\<Union>m\<le>k. S m)" for k
using assms by blast
then have 0: "integral UNIV (indicat_real (\<Union>m\<le>k. S m)) = 0"
and 1: "(indicat_real (\<Union>m\<le>k. S m)) integrable_on UNIV" for k
by (auto simp: negligible has_integral_iff)
have 2: "\<And>k x. indicat_real (\<Union>m\<le>k. S m) x \<le> (indicat_real (\<Union>m\<le>Suc k. S m) x)"
by (simp add: indicator_def)
have 3: "\<And>x. (\<lambda>k. indicat_real (\<Union>m\<le>k. S m) x) \<longlonglongrightarrow> (indicat_real (\<Union>n. S n) x)"
by (force simp: indicator_def eventually_sequentially intro: tendsto_eventually)
have 4: "bounded (range (\<lambda>k. integral UNIV (indicat_real (\<Union>m\<le>k. S m))))"
by (simp add: 0)
have *: "indicat_real (\<Union>n. S n) integrable_on UNIV \<and>
(\<lambda>k. integral UNIV (indicat_real (\<Union>m\<le>k. S m))) \<longlonglongrightarrow> (integral UNIV (indicat_real (\<Union>n. S n)))"
by (intro monotone_convergence_increasing 1 2 3 4)
then have "integral UNIV (indicat_real (\<Union>n. S n)) = (0::real)"
using LIMSEQ_unique by (auto simp: 0)
then show ?thesis
using * by (simp add: negligible_UNIV has_integral_iff)
qed
lemma negligible_linear_singular_image:
fixes f :: "'n::euclidean_space \<Rightarrow> 'n"
assumes "linear f" "\<not> inj f"
shows "negligible (f ` S)"
proof -
obtain a where "a \<noteq> 0" "\<And>S. f ` S \<subseteq> {x. a \<bullet> x = 0}"
using assms linear_singular_image_hyperplane by blast
then show "negligible (f ` S)"
by (metis negligible_hyperplane negligible_subset)
qed
lemma measure_negligible_finite_Union:
assumes "finite \<F>"
and meas: "\<And>S. S \<in> \<F> \<Longrightarrow> S \<in> lmeasurable"
and djointish: "pairwise (\<lambda>S T. negligible (S \<inter> T)) \<F>"
shows "measure lebesgue (\<Union>\<F>) = (\<Sum>S\<in>\<F>. measure lebesgue S)"
using assms
proof (induction)
case empty
then show ?case
by auto
next
case (insert S \<F>)
then have "S \<in> lmeasurable" "\<Union>\<F> \<in> lmeasurable" "pairwise (\<lambda>S T. negligible (S \<inter> T)) \<F>"
by (simp_all add: fmeasurable.finite_Union insert.hyps(1) insert.prems(1) pairwise_insert subsetI)
then show ?case
proof (simp add: measure_Un3 insert)
have *: "\<And>T. T \<in> (\<inter>) S ` \<F> \<Longrightarrow> negligible T"
using insert by (force simp: pairwise_def)
have "negligible(S \<inter> \<Union>\<F>)"
unfolding Int_Union
by (rule negligible_Union) (simp_all add: * insert.hyps(1))
then show "measure lebesgue (S \<inter> \<Union>\<F>) = 0"
using negligible_imp_measure0 by blast
qed
qed
lemma measure_negligible_finite_Union_image:
assumes "finite S"
and meas: "\<And>x. x \<in> S \<Longrightarrow> f x \<in> lmeasurable"
and djointish: "pairwise (\<lambda>x y. negligible (f x \<inter> f y)) S"
shows "measure lebesgue (\<Union>(f ` S)) = (\<Sum>x\<in>S. measure lebesgue (f x))"
proof -
have "measure lebesgue (\<Union>(f ` S)) = sum (measure lebesgue) (f ` S)"
using assms by (auto simp: pairwise_mono pairwise_image intro: measure_negligible_finite_Union)
also have "\<dots> = sum (measure lebesgue \<circ> f) S"
using djointish [unfolded pairwise_def] by (metis inf.idem negligible_imp_measure0 sum.reindex_nontrivial [OF \<open>finite S\<close>])
also have "\<dots> = (\<Sum>x\<in>S. measure lebesgue (f x))"
by simp
finally show ?thesis .
qed
subsection \<open>Negligibility of a Lipschitz image of a negligible set\<close>
text\<open>The bound will be eliminated by a sort of onion argument\<close>
lemma locally_Lipschitz_negl_bounded:
fixes f :: "'M::euclidean_space \<Rightarrow> 'N::euclidean_space"
assumes MleN: "DIM('M) \<le> DIM('N)" "0 < B" "bounded S" "negligible S"
and lips: "\<And>x. x \<in> S
\<Longrightarrow> \<exists>T. open T \<and> x \<in> T \<and>
(\<forall>y \<in> S \<inter> T. norm(f y - f x) \<le> B * norm(y - x))"
shows "negligible (f ` S)"
unfolding negligible_iff_null_sets
proof (clarsimp simp: completion.null_sets_outer)
fix e::real
assume "0 < e"
have "S \<in> lmeasurable"
using \<open>negligible S\<close> by (simp add: negligible_iff_null_sets fmeasurableI_null_sets)
then have "S \<in> sets lebesgue"
by blast
have e22: "0 < e/2 / (2 * B * real DIM('M)) ^ DIM('N)"
using \<open>0 < e\<close> \<open>0 < B\<close> by (simp add: field_split_simps)
obtain T where "open T" "S \<subseteq> T" "(T - S) \<in> lmeasurable"
"measure lebesgue (T - S) < e/2 / (2 * B * DIM('M)) ^ DIM('N)"
using sets_lebesgue_outer_open [OF \<open>S \<in> sets lebesgue\<close> e22]
by (metis emeasure_eq_measure2 ennreal_leI linorder_not_le)
then have T: "measure lebesgue T \<le> e/2 / (2 * B * DIM('M)) ^ DIM('N)"
using \<open>negligible S\<close> by (simp add: measure_Diff_null_set negligible_iff_null_sets)
have "\<exists>r. 0 < r \<and> r \<le> 1/2 \<and>
(x \<in> S \<longrightarrow> (\<forall>y. norm(y - x) < r
\<longrightarrow> y \<in> T \<and> (y \<in> S \<longrightarrow> norm(f y - f x) \<le> B * norm(y - x))))"
for x
proof (cases "x \<in> S")
case True
obtain U where "open U" "x \<in> U" and U: "\<And>y. y \<in> S \<inter> U \<Longrightarrow> norm(f y - f x) \<le> B * norm(y - x)"
using lips [OF \<open>x \<in> S\<close>] by auto
have "x \<in> T \<inter> U"
using \<open>S \<subseteq> T\<close> \<open>x \<in> U\<close> \<open>x \<in> S\<close> by auto
then obtain \<epsilon> where "0 < \<epsilon>" "ball x \<epsilon> \<subseteq> T \<inter> U"
by (metis \<open>open T\<close> \<open>open U\<close> openE open_Int)
then show ?thesis
apply (rule_tac x="min (1/2) \<epsilon>" in exI)
apply (simp del: divide_const_simps)
apply (intro allI impI conjI)
apply (metis dist_commute dist_norm mem_ball subsetCE)
by (metis Int_iff subsetCE U dist_norm mem_ball norm_minus_commute)
next
case False
then show ?thesis
by (rule_tac x="1/4" in exI) auto
qed
then obtain R where R12: "\<And>x. 0 < R x \<and> R x \<le> 1/2"
and RT: "\<And>x y. \<lbrakk>x \<in> S; norm(y - x) < R x\<rbrakk> \<Longrightarrow> y \<in> T"
and RB: "\<And>x y. \<lbrakk>x \<in> S; y \<in> S; norm(y - x) < R x\<rbrakk> \<Longrightarrow> norm(f y - f x) \<le> B * norm(y - x)"
by metis+
then have gaugeR: "gauge (\<lambda>x. ball x (R x))"
by (simp add: gauge_def)
obtain c where c: "S \<subseteq> cbox (-c *\<^sub>R One) (c *\<^sub>R One)" "box (-c *\<^sub>R One:: 'M) (c *\<^sub>R One) \<noteq> {}"
proof -
obtain B where B: "\<And>x. x \<in> S \<Longrightarrow> norm x \<le> B"
using \<open>bounded S\<close> bounded_iff by blast
show ?thesis
apply (rule_tac c = "abs B + 1" in that)
using norm_bound_Basis_le Basis_le_norm
apply (fastforce simp: box_eq_empty mem_box dest!: B intro: order_trans)+
done
qed
obtain \<D> where "countable \<D>"
and Dsub: "\<Union>\<D> \<subseteq> cbox (-c *\<^sub>R One) (c *\<^sub>R One)"
and cbox: "\<And>K. K \<in> \<D> \<Longrightarrow> interior K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
and pw: "pairwise (\<lambda>A B. interior A \<inter> interior B = {}) \<D>"
and Ksub: "\<And>K. K \<in> \<D> \<Longrightarrow> \<exists>x \<in> S \<inter> K. K \<subseteq> (\<lambda>x. ball x (R x)) x"
and exN: "\<And>u v. cbox u v \<in> \<D> \<Longrightarrow> \<exists>n. \<forall>i \<in> Basis. v \<bullet> i - u \<bullet> i = (2*c) / 2^n"
and "S \<subseteq> \<Union>\<D>"
using covering_lemma [OF c gaugeR] by force
have "\<exists>u v z. K = cbox u v \<and> box u v \<noteq> {} \<and> z \<in> S \<and> z \<in> cbox u v \<and>
cbox u v \<subseteq> ball z (R z)" if "K \<in> \<D>" for K
proof -
obtain u v where "K = cbox u v"
using \<open>K \<in> \<D>\<close> cbox by blast
with that show ?thesis
apply (rule_tac x=u in exI)
apply (rule_tac x=v in exI)
apply (metis Int_iff interior_cbox cbox Ksub)
done
qed
then obtain uf vf zf
where uvz: "\<And>K. K \<in> \<D> \<Longrightarrow>
K = cbox (uf K) (vf K) \<and> box (uf K) (vf K) \<noteq> {} \<and> zf K \<in> S \<and>
zf K \<in> cbox (uf K) (vf K) \<and> cbox (uf K) (vf K) \<subseteq> ball (zf K) (R (zf K))"
by metis
define prj1 where "prj1 \<equiv> \<lambda>x::'M. x \<bullet> (SOME i. i \<in> Basis)"
define fbx where "fbx \<equiv> \<lambda>D. cbox (f(zf D) - (B * DIM('M) * (prj1(vf D - uf D))) *\<^sub>R One::'N)
(f(zf D) + (B * DIM('M) * prj1(vf D - uf D)) *\<^sub>R One)"
have vu_pos: "0 < prj1 (vf X - uf X)" if "X \<in> \<D>" for X
using uvz [OF that] by (simp add: prj1_def box_ne_empty SOME_Basis inner_diff_left)
have prj1_idem: "prj1 (vf X - uf X) = (vf X - uf X) \<bullet> i" if "X \<in> \<D>" "i \<in> Basis" for X i
proof -
have "cbox (uf X) (vf X) \<in> \<D>"
using uvz \<open>X \<in> \<D>\<close> by auto
with exN obtain n where "\<And>i. i \<in> Basis \<Longrightarrow> vf X \<bullet> i - uf X \<bullet> i = (2*c) / 2^n"
by blast
then show ?thesis
by (simp add: \<open>i \<in> Basis\<close> SOME_Basis inner_diff prj1_def)
qed
have countbl: "countable (fbx ` \<D>)"
using \<open>countable \<D>\<close> by blast
have "(\<Sum>k\<in>fbx`\<D>'. measure lebesgue k) \<le> e/2" if "\<D>' \<subseteq> \<D>" "finite \<D>'" for \<D>'
proof -
have BM_ge0: "0 \<le> B * (DIM('M) * prj1 (vf X - uf X))" if "X \<in> \<D>'" for X
using \<open>0 < B\<close> \<open>\<D>' \<subseteq> \<D>\<close> that vu_pos by fastforce
have "{} \<notin> \<D>'"
using cbox \<open>\<D>' \<subseteq> \<D>\<close> interior_empty by blast
have "(\<Sum>k\<in>fbx`\<D>'. measure lebesgue k) \<le> sum (measure lebesgue o fbx) \<D>'"
by (rule sum_image_le [OF \<open>finite \<D>'\<close>]) (force simp: fbx_def)
also have "\<dots> \<le> (\<Sum>X\<in>\<D>'. (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X)"
proof (rule sum_mono)
fix X assume "X \<in> \<D>'"
then have "X \<in> \<D>" using \<open>\<D>' \<subseteq> \<D>\<close> by blast
then have ufvf: "cbox (uf X) (vf X) = X"
using uvz by blast
have "prj1 (vf X - uf X) ^ DIM('M) = (\<Prod>i::'M \<in> Basis. prj1 (vf X - uf X))"
by (rule prod_constant [symmetric])
also have "\<dots> = (\<Prod>i\<in>Basis. vf X \<bullet> i - uf X \<bullet> i)"
apply (rule prod.cong [OF refl])
by (simp add: \<open>X \<in> \<D>\<close> inner_diff_left prj1_idem)
finally have prj1_eq: "prj1 (vf X - uf X) ^ DIM('M) = (\<Prod>i\<in>Basis. vf X \<bullet> i - uf X \<bullet> i)" .
have "uf X \<in> cbox (uf X) (vf X)" "vf X \<in> cbox (uf X) (vf X)"
using uvz [OF \<open>X \<in> \<D>\<close>] by (force simp: mem_box)+
moreover have "cbox (uf X) (vf X) \<subseteq> ball (zf X) (1/2)"
by (meson R12 order_trans subset_ball uvz [OF \<open>X \<in> \<D>\<close>])
ultimately have "uf X \<in> ball (zf X) (1/2)" "vf X \<in> ball (zf X) (1/2)"
by auto
then have "dist (vf X) (uf X) \<le> 1"
unfolding mem_ball
by (metis dist_commute dist_triangle_half_l dual_order.order_iff_strict)
then have 1: "prj1 (vf X - uf X) \<le> 1"
unfolding prj1_def dist_norm using Basis_le_norm SOME_Basis order_trans by fastforce
have 0: "0 \<le> prj1 (vf X - uf X)"
using \<open>X \<in> \<D>\<close> prj1_def vu_pos by fastforce
have "(measure lebesgue \<circ> fbx) X \<le> (2 * B * DIM('M)) ^ DIM('N) * content (cbox (uf X) (vf X))"
apply (simp add: fbx_def content_cbox_cases algebra_simps BM_ge0 \<open>X \<in> \<D>'\<close>)
apply (simp add: power_mult_distrib \<open>0 < B\<close> prj1_eq [symmetric])
using MleN 0 1 uvz \<open>X \<in> \<D>\<close>
apply (fastforce simp add: box_ne_empty power_decreasing)
done
also have "\<dots> = (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X"
by (subst (3) ufvf[symmetric]) simp
finally show "(measure lebesgue \<circ> fbx) X \<le> (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X" .
qed
also have "\<dots> = (2 * B * DIM('M)) ^ DIM('N) * sum (measure lebesgue) \<D>'"
by (simp add: sum_distrib_left)
also have "\<dots> \<le> e/2"
proof -
have div: "\<D>' division_of \<Union>\<D>'"
apply (auto simp: \<open>finite \<D>'\<close> \<open>{} \<notin> \<D>'\<close> division_of_def)
using cbox that apply blast
using pairwise_subset [OF pw \<open>\<D>' \<subseteq> \<D>\<close>] unfolding pairwise_def apply force+
done
have le_meaT: "measure lebesgue (\<Union>\<D>') \<le> measure lebesgue T"
proof (rule measure_mono_fmeasurable)
show "(\<Union>\<D>') \<in> sets lebesgue"
using div lmeasurable_division by auto
have "\<Union>\<D>' \<subseteq> \<Union>\<D>"
using \<open>\<D>' \<subseteq> \<D>\<close> by blast
also have "... \<subseteq> T"
proof (clarify)
fix x D
assume "x \<in> D" "D \<in> \<D>"
show "x \<in> T"
using Ksub [OF \<open>D \<in> \<D>\<close>]
by (metis \<open>x \<in> D\<close> Int_iff dist_norm mem_ball norm_minus_commute subsetD RT)
qed
finally show "\<Union>\<D>' \<subseteq> T" .
show "T \<in> lmeasurable"
using \<open>S \<in> lmeasurable\<close> \<open>S \<subseteq> T\<close> \<open>T - S \<in> lmeasurable\<close> fmeasurable_Diff_D by blast
qed
have "sum (measure lebesgue) \<D>' = sum content \<D>'"
using \<open>\<D>' \<subseteq> \<D>\<close> cbox by (force intro: sum.cong)
then have "(2 * B * DIM('M)) ^ DIM('N) * sum (measure lebesgue) \<D>' =
(2 * B * real DIM('M)) ^ DIM('N) * measure lebesgue (\<Union>\<D>')"
using content_division [OF div] by auto
also have "\<dots> \<le> (2 * B * real DIM('M)) ^ DIM('N) * measure lebesgue T"
apply (rule mult_left_mono [OF le_meaT])
using \<open>0 < B\<close>
apply (simp add: algebra_simps)
done
also have "\<dots> \<le> e/2"
using T \<open>0 < B\<close> by (simp add: field_simps)
finally show ?thesis .
qed
finally show ?thesis .
qed
then have e2: "sum (measure lebesgue) \<G> \<le> e/2" if "\<G> \<subseteq> fbx ` \<D>" "finite \<G>" for \<G>
by (metis finite_subset_image that)
show "\<exists>W\<in>lmeasurable. f ` S \<subseteq> W \<and> measure lebesgue W < e"
proof (intro bexI conjI)
have "\<exists>X\<in>\<D>. f y \<in> fbx X" if "y \<in> S" for y
proof -
obtain X where "y \<in> X" "X \<in> \<D>"
using \<open>S \<subseteq> \<Union>\<D>\<close> \<open>y \<in> S\<close> by auto
then have y: "y \<in> ball(zf X) (R(zf X))"
using uvz by fastforce
have conj_le_eq: "z - b \<le> y \<and> y \<le> z + b \<longleftrightarrow> abs(y - z) \<le> b" for z y b::real
by auto
have yin: "y \<in> cbox (uf X) (vf X)" and zin: "(zf X) \<in> cbox (uf X) (vf X)"
using uvz \<open>X \<in> \<D>\<close> \<open>y \<in> X\<close> by auto
have "norm (y - zf X) \<le> (\<Sum>i\<in>Basis. \<bar>(y - zf X) \<bullet> i\<bar>)"
by (rule norm_le_l1)
also have "\<dots> \<le> real DIM('M) * prj1 (vf X - uf X)"
proof (rule sum_bounded_above)
fix j::'M assume j: "j \<in> Basis"
show "\<bar>(y - zf X) \<bullet> j\<bar> \<le> prj1 (vf X - uf X)"
using yin zin j
by (fastforce simp add: mem_box prj1_idem [OF \<open>X \<in> \<D>\<close> j] inner_diff_left)
qed
finally have nole: "norm (y - zf X) \<le> DIM('M) * prj1 (vf X - uf X)"
by simp
have fle: "\<bar>f y \<bullet> i - f(zf X) \<bullet> i\<bar> \<le> B * DIM('M) * prj1 (vf X - uf X)" if "i \<in> Basis" for i
proof -
have "\<bar>f y \<bullet> i - f (zf X) \<bullet> i\<bar> = \<bar>(f y - f (zf X)) \<bullet> i\<bar>"
by (simp add: algebra_simps)
also have "\<dots> \<le> norm (f y - f (zf X))"
by (simp add: Basis_le_norm that)
also have "\<dots> \<le> B * norm(y - zf X)"
by (metis uvz RB \<open>X \<in> \<D>\<close> dist_commute dist_norm mem_ball \<open>y \<in> S\<close> y)
also have "\<dots> \<le> B * real DIM('M) * prj1 (vf X - uf X)"
using \<open>0 < B\<close> by (simp add: nole)
finally show ?thesis .
qed
show ?thesis
by (rule_tac x=X in bexI)
(auto simp: fbx_def prj1_idem mem_box conj_le_eq inner_add inner_diff fle \<open>X \<in> \<D>\<close>)
qed
then show "f ` S \<subseteq> (\<Union>D\<in>\<D>. fbx D)" by auto
next
have 1: "\<And>D. D \<in> \<D> \<Longrightarrow> fbx D \<in> lmeasurable"
by (auto simp: fbx_def)
have 2: "I' \<subseteq> \<D> \<Longrightarrow> finite I' \<Longrightarrow> measure lebesgue (\<Union>D\<in>I'. fbx D) \<le> e/2" for I'
by (rule order_trans[OF measure_Union_le e2]) (auto simp: fbx_def)
show "(\<Union>D\<in>\<D>. fbx D) \<in> lmeasurable"
by (intro fmeasurable_UN_bound[OF \<open>countable \<D>\<close> 1 2])
have "measure lebesgue (\<Union>D\<in>\<D>. fbx D) \<le> e/2"
by (intro measure_UN_bound[OF \<open>countable \<D>\<close> 1 2])
then show "measure lebesgue (\<Union>D\<in>\<D>. fbx D) < e"
using \<open>0 < e\<close> by linarith
qed
qed
proposition negligible_locally_Lipschitz_image:
fixes f :: "'M::euclidean_space \<Rightarrow> 'N::euclidean_space"
assumes MleN: "DIM('M) \<le> DIM('N)" "negligible S"
and lips: "\<And>x. x \<in> S
\<Longrightarrow> \<exists>T B. open T \<and> x \<in> T \<and>
(\<forall>y \<in> S \<inter> T. norm(f y - f x) \<le> B * norm(y - x))"
shows "negligible (f ` S)"
proof -
let ?S = "\<lambda>n. ({x \<in> S. norm x \<le> n \<and>
(\<exists>T. open T \<and> x \<in> T \<and>
(\<forall>y\<in>S \<inter> T. norm (f y - f x) \<le> (real n + 1) * norm (y - x)))})"
have negfn: "f ` ?S n \<in> null_sets lebesgue" for n::nat
unfolding negligible_iff_null_sets[symmetric]
apply (rule_tac B = "real n + 1" in locally_Lipschitz_negl_bounded)
by (auto simp: MleN bounded_iff intro: negligible_subset [OF \<open>negligible S\<close>])
have "S = (\<Union>n. ?S n)"
proof (intro set_eqI iffI)
fix x assume "x \<in> S"
with lips obtain T B where T: "open T" "x \<in> T"
and B: "\<And>y. y \<in> S \<inter> T \<Longrightarrow> norm(f y - f x) \<le> B * norm(y - x)"
by metis+
have no: "norm (f y - f x) \<le> (nat \<lceil>max B (norm x)\<rceil> + 1) * norm (y - x)" if "y \<in> S \<inter> T" for y
proof -
have "B * norm(y - x) \<le> (nat \<lceil>max B (norm x)\<rceil> + 1) * norm (y - x)"
by (meson max.cobounded1 mult_right_mono nat_ceiling_le_eq nat_le_iff_add norm_ge_zero order_trans)
then show ?thesis
using B order_trans that by blast
qed
have "x \<in> ?S (nat (ceiling (max B (norm x))))"
apply (simp add: \<open>x \<in> S \<close>, rule)
using real_nat_ceiling_ge max.bounded_iff apply blast
using T no
apply (force simp: algebra_simps)
done
then show "x \<in> (\<Union>n. ?S n)" by force
qed auto
then show ?thesis
by (rule ssubst) (auto simp: image_Union negligible_iff_null_sets intro: negfn)
qed
corollary negligible_differentiable_image_negligible:
fixes f :: "'M::euclidean_space \<Rightarrow> 'N::euclidean_space"
assumes MleN: "DIM('M) \<le> DIM('N)" "negligible S"
and diff_f: "f differentiable_on S"
shows "negligible (f ` S)"
proof -
have "\<exists>T B. open T \<and> x \<in> T \<and> (\<forall>y \<in> S \<inter> T. norm(f y - f x) \<le> B * norm(y - x))"
if "x \<in> S" for x
proof -
obtain f' where "linear f'"
and f': "\<And>e. e>0 \<Longrightarrow>
\<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)"
using diff_f \<open>x \<in> S\<close>
by (auto simp: linear_linear differentiable_on_def differentiable_def has_derivative_within_alt)
obtain B where "B > 0" and B: "\<forall>x. norm (f' x) \<le> B * norm x"
using linear_bounded_pos \<open>linear f'\<close> by blast
obtain d where "d>0"
and d: "\<And>y. \<lbrakk>y \<in> S; norm (y - x) < d\<rbrakk> \<Longrightarrow>
norm (f y - f x - f' (y - x)) \<le> norm (y - x)"
using f' [of 1] by (force simp:)
have *: "norm (f y - f x) \<le> (B + 1) * norm (y - x)"
if "y \<in> S" "norm (y - x) < d" for y
proof -
have "norm (f y - f x) -B * norm (y - x) \<le> norm (f y - f x) - norm (f' (y - x))"
by (simp add: B)
also have "\<dots> \<le> norm (f y - f x - f' (y - x))"
by (rule norm_triangle_ineq2)
also have "... \<le> norm (y - x)"
by (rule d [OF that])
finally show ?thesis
by (simp add: algebra_simps)
qed
show ?thesis
apply (rule_tac x="ball x d" in exI)
apply (rule_tac x="B+1" in exI)
using \<open>d>0\<close>
apply (auto simp: dist_norm norm_minus_commute intro!: *)
done
qed
with negligible_locally_Lipschitz_image assms show ?thesis by metis
qed
corollary negligible_differentiable_image_lowdim:
fixes f :: "'M::euclidean_space \<Rightarrow> 'N::euclidean_space"
assumes MlessN: "DIM('M) < DIM('N)" and diff_f: "f differentiable_on S"
shows "negligible (f ` S)"
proof -
have "x \<le> DIM('M) \<Longrightarrow> x \<le> DIM('N)" for x
using MlessN by linarith
obtain lift :: "'M * real \<Rightarrow> 'N" and drop :: "'N \<Rightarrow> 'M * real" and j :: 'N
where "linear lift" "linear drop" and dropl [simp]: "\<And>z. drop (lift z) = z"
and "j \<in> Basis" and j: "\<And>x. lift(x,0) \<bullet> j = 0"
using lowerdim_embeddings [OF MlessN] by metis
have "negligible {x. x\<bullet>j = 0}"
by (metis \<open>j \<in> Basis\<close> negligible_standard_hyperplane)
then have neg0S: "negligible ((\<lambda>x. lift (x, 0)) ` S)"
apply (rule negligible_subset)
by (simp add: image_subsetI j)
have diff_f': "f \<circ> fst \<circ> drop differentiable_on (\<lambda>x. lift (x, 0)) ` S"
using diff_f
apply (clarsimp simp add: differentiable_on_def)
apply (intro differentiable_chain_within linear_imp_differentiable [OF \<open>linear drop\<close>]
linear_imp_differentiable [OF linear_fst])
apply (force simp: image_comp o_def)
done
have "f = (f o fst o drop o (\<lambda>x. lift (x, 0)))"
by (simp add: o_def)
then show ?thesis
apply (rule ssubst)
apply (subst image_comp [symmetric])
apply (metis negligible_differentiable_image_negligible order_refl diff_f' neg0S)
done
qed
subsection\<open>Measurability of countable unions and intersections of various kinds.\<close>
lemma
assumes S: "\<And>n. S n \<in> lmeasurable"
and leB: "\<And>n. measure lebesgue (S n) \<le> B"
and nest: "\<And>n. S n \<subseteq> S(Suc n)"
shows measurable_nested_Union: "(\<Union>n. S n) \<in> lmeasurable"
and measure_nested_Union: "(\<lambda>n. measure lebesgue (S n)) \<longlonglongrightarrow> measure lebesgue (\<Union>n. S n)" (is ?Lim)
proof -
have 1: "\<And>n. (indicat_real (S n)) integrable_on UNIV"
using S measurable_integrable by blast
have 2: "\<And>n x::'a. indicat_real (S n) x \<le> (indicat_real (S (Suc n)) x)"
by (simp add: indicator_leI nest rev_subsetD)
have 3: "\<And>x. (\<lambda>n. indicat_real (S n) x) \<longlonglongrightarrow> (indicat_real (\<Union>(S ` UNIV)) x)"
apply (rule tendsto_eventually)
apply (simp add: indicator_def)
by (metis eventually_sequentiallyI lift_Suc_mono_le nest subsetCE)
have 4: "bounded (range (\<lambda>n. integral UNIV (indicat_real (S n))))"
using leB by (auto simp: lmeasure_integral_UNIV [symmetric] S bounded_iff)
have "(\<Union>n. S n) \<in> lmeasurable \<and> ?Lim"
apply (simp add: lmeasure_integral_UNIV S cong: conj_cong)
apply (simp add: measurable_integrable)
apply (rule monotone_convergence_increasing [OF 1 2 3 4])
done
then show "(\<Union>n. S n) \<in> lmeasurable" "?Lim"
by auto
qed
lemma
assumes S: "\<And>n. S n \<in> lmeasurable"
and djointish: "pairwise (\<lambda>m n. negligible (S m \<inter> S n)) UNIV"
and leB: "\<And>n. (\<Sum>k\<le>n. measure lebesgue (S k)) \<le> B"
shows measurable_countable_negligible_Union: "(\<Union>n. S n) \<in> lmeasurable"
and measure_countable_negligible_Union: "(\<lambda>n. (measure lebesgue (S n))) sums measure lebesgue (\<Union>n. S n)" (is ?Sums)
proof -
have 1: "\<Union> (S ` {..n}) \<in> lmeasurable" for n
using S by blast
have 2: "measure lebesgue (\<Union> (S ` {..n})) \<le> B" for n
proof -
have "measure lebesgue (\<Union> (S ` {..n})) \<le> (\<Sum>k\<le>n. measure lebesgue (S k))"
by (simp add: S fmeasurableD measure_UNION_le)
with leB show ?thesis
using order_trans by blast
qed
have 3: "\<And>n. \<Union> (S ` {..n}) \<subseteq> \<Union> (S ` {..Suc n})"
by (simp add: SUP_subset_mono)
have eqS: "(\<Union>n. S n) = (\<Union>n. \<Union> (S ` {..n}))"
using atLeastAtMost_iff by blast
also have "(\<Union>n. \<Union> (S ` {..n})) \<in> lmeasurable"
by (intro measurable_nested_Union [OF 1 2] 3)
finally show "(\<Union>n. S n) \<in> lmeasurable" .
have eqm: "(\<Sum>i\<le>n. measure lebesgue (S i)) = measure lebesgue (\<Union> (S ` {..n}))" for n
using assms by (simp add: measure_negligible_finite_Union_image pairwise_mono)
have "(\<lambda>n. (measure lebesgue (S n))) sums measure lebesgue (\<Union>n. \<Union> (S ` {..n}))"
by (simp add: sums_def' eqm atLeast0AtMost) (intro measure_nested_Union [OF 1 2] 3)
then show ?Sums
by (simp add: eqS)
qed
lemma negligible_countable_Union [intro]:
assumes "countable \<F>" and meas: "\<And>S. S \<in> \<F> \<Longrightarrow> negligible S"
shows "negligible (\<Union>\<F>)"
proof (cases "\<F> = {}")
case False
then show ?thesis
by (metis from_nat_into range_from_nat_into assms negligible_Union_nat)
qed simp
lemma
assumes S: "\<And>n. (S n) \<in> lmeasurable"
and djointish: "pairwise (\<lambda>m n. negligible (S m \<inter> S n)) UNIV"
and bo: "bounded (\<Union>n. S n)"
shows measurable_countable_negligible_Union_bounded: "(\<Union>n. S n) \<in> lmeasurable"
and measure_countable_negligible_Union_bounded: "(\<lambda>n. (measure lebesgue (S n))) sums measure lebesgue (\<Union>n. S n)" (is ?Sums)
proof -
obtain a b where ab: "(\<Union>n. S n) \<subseteq> cbox a b"
using bo bounded_subset_cbox_symmetric by metis
then have B: "(\<Sum>k\<le>n. measure lebesgue (S k)) \<le> measure lebesgue (cbox a b)" for n
proof -
have "(\<Sum>k\<le>n. measure lebesgue (S k)) = measure lebesgue (\<Union> (S ` {..n}))"
using measure_negligible_finite_Union_image [OF _ _ pairwise_subset] djointish
by (metis S finite_atMost subset_UNIV)
also have "\<dots> \<le> measure lebesgue (cbox a b)"
apply (rule measure_mono_fmeasurable)
using ab S by force+
finally show ?thesis .
qed
show "(\<Union>n. S n) \<in> lmeasurable"
by (rule measurable_countable_negligible_Union [OF S djointish B])
show ?Sums
by (rule measure_countable_negligible_Union [OF S djointish B])
qed
lemma measure_countable_Union_approachable:
assumes "countable \<D>" "e > 0" and measD: "\<And>d. d \<in> \<D> \<Longrightarrow> d \<in> lmeasurable"
and B: "\<And>D'. \<lbrakk>D' \<subseteq> \<D>; finite D'\<rbrakk> \<Longrightarrow> measure lebesgue (\<Union>D') \<le> B"
obtains D' where "D' \<subseteq> \<D>" "finite D'" "measure lebesgue (\<Union>\<D>) - e < measure lebesgue (\<Union>D')"
proof (cases "\<D> = {}")
case True
then show ?thesis
by (simp add: \<open>e > 0\<close> that)
next
case False
let ?S = "\<lambda>n. \<Union>k \<le> n. from_nat_into \<D> k"
have "(\<lambda>n. measure lebesgue (?S n)) \<longlonglongrightarrow> measure lebesgue (\<Union>n. ?S n)"
proof (rule measure_nested_Union)
show "?S n \<in> lmeasurable" for n
by (simp add: False fmeasurable.finite_UN from_nat_into measD)
show "measure lebesgue (?S n) \<le> B" for n
by (metis (mono_tags, lifting) B False finite_atMost finite_imageI from_nat_into image_iff subsetI)
show "?S n \<subseteq> ?S (Suc n)" for n
by force
qed
then obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> dist (measure lebesgue (?S n)) (measure lebesgue (\<Union>n. ?S n)) < e"
using metric_LIMSEQ_D \<open>e > 0\<close> by blast
show ?thesis
proof
show "from_nat_into \<D> ` {..N} \<subseteq> \<D>"
by (auto simp: False from_nat_into)
have eq: "(\<Union>n. \<Union>k\<le>n. from_nat_into \<D> k) = (\<Union>\<D>)"
using \<open>countable \<D>\<close> False
by (auto intro: from_nat_into dest: from_nat_into_surj [OF \<open>countable \<D>\<close>])
show "measure lebesgue (\<Union>\<D>) - e < measure lebesgue (\<Union> (from_nat_into \<D> ` {..N}))"
using N [OF order_refl]
by (auto simp: eq algebra_simps dist_norm)
qed auto
qed
subsection\<open>Negligibility is a local property\<close>
lemma locally_negligible_alt:
"negligible S \<longleftrightarrow> (\<forall>x \<in> S. \<exists>U. openin (top_of_set S) U \<and> x \<in> U \<and> negligible U)"
(is "_ = ?rhs")
proof
assume "negligible S"
then show ?rhs
using openin_subtopology_self by blast
next
assume ?rhs
then obtain U where ope: "\<And>x. x \<in> S \<Longrightarrow> openin (top_of_set S) (U x)"
and cov: "\<And>x. x \<in> S \<Longrightarrow> x \<in> U x"
and neg: "\<And>x. x \<in> S \<Longrightarrow> negligible (U x)"
by metis
obtain \<F> where "\<F> \<subseteq> U ` S" "countable \<F>" and eq: "\<Union>\<F> = \<Union>(U ` S)"
using ope by (force intro: Lindelof_openin [of "U ` S" S])
then have "negligible (\<Union>\<F>)"
by (metis imageE neg negligible_countable_Union subset_eq)
with eq have "negligible (\<Union>(U ` S))"
by metis
moreover have "S \<subseteq> \<Union>(U ` S)"
using cov by blast
ultimately show "negligible S"
using negligible_subset by blast
qed
lemma locally_negligible:
"locally negligible S \<longleftrightarrow> negligible S"
unfolding locally_def
apply safe
apply (meson negligible_subset openin_subtopology_self locally_negligible_alt)
by (meson negligible_subset openin_imp_subset order_refl)
subsection\<open>Integral bounds\<close>
lemma set_integral_norm_bound:
fixes f :: "_ \<Rightarrow> 'a :: {banach, second_countable_topology}"
shows "set_integrable M k f \<Longrightarrow> norm (LINT x:k|M. f x) \<le> LINT x:k|M. norm (f x)"
using integral_norm_bound[of M "\<lambda>x. indicator k x *\<^sub>R f x"] by (simp add: set_lebesgue_integral_def)
lemma set_integral_finite_UN_AE:
fixes f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
assumes "finite I"
and ae: "\<And>i j. i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> AE x in M. (x \<in> A i \<and> x \<in> A j) \<longrightarrow> i = j"
and [measurable]: "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets M"
and f: "\<And>i. i \<in> I \<Longrightarrow> set_integrable M (A i) f"
shows "LINT x:(\<Union>i\<in>I. A i)|M. f x = (\<Sum>i\<in>I. LINT x:A i|M. f x)"
using \<open>finite I\<close> order_refl[of I]
proof (induction I rule: finite_subset_induct')
case (insert i I')
have "AE x in M. (\<forall>j\<in>I'. x \<in> A i \<longrightarrow> x \<notin> A j)"
proof (intro AE_ball_countable[THEN iffD2] ballI)
fix j assume "j \<in> I'"
with \<open>I' \<subseteq> I\<close> \<open>i \<notin> I'\<close> have "i \<noteq> j" "j \<in> I"
by auto
then show "AE x in M. x \<in> A i \<longrightarrow> x \<notin> A j"
using ae[of i j] \<open>i \<in> I\<close> by auto
qed (use \<open>finite I'\<close> in \<open>rule countable_finite\<close>)
then have "AE x\<in>A i in M. \<forall>xa\<in>I'. x \<notin> A xa "
by auto
with insert.hyps insert.IH[symmetric]
show ?case
by (auto intro!: set_integral_Un_AE sets.finite_UN f set_integrable_UN)
qed (simp add: set_lebesgue_integral_def)
lemma set_integrable_norm:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes f: "set_integrable M k f" shows "set_integrable M k (\<lambda>x. norm (f x))"
using integrable_norm f by (force simp add: set_integrable_def)
lemma absolutely_integrable_bounded_variation:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes f: "f absolutely_integrable_on UNIV"
obtains B where "\<forall>d. d division_of (\<Union>d) \<longrightarrow> sum (\<lambda>k. norm (integral k f)) d \<le> B"
proof (rule that[of "integral UNIV (\<lambda>x. norm (f x))"]; safe)
fix d :: "'a set set" assume d: "d division_of \<Union>d"
have *: "k \<in> d \<Longrightarrow> f absolutely_integrable_on k" for k
using f[THEN set_integrable_subset, of k] division_ofD(2,4)[OF d, of k] by auto
note d' = division_ofD[OF d]
have "(\<Sum>k\<in>d. norm (integral k f)) = (\<Sum>k\<in>d. norm (LINT x:k|lebesgue. f x))"
by (intro sum.cong refl arg_cong[where f=norm] set_lebesgue_integral_eq_integral(2)[symmetric] *)
also have "\<dots> \<le> (\<Sum>k\<in>d. LINT x:k|lebesgue. norm (f x))"
by (intro sum_mono set_integral_norm_bound *)
also have "\<dots> = (\<Sum>k\<in>d. integral k (\<lambda>x. norm (f x)))"
by (intro sum.cong refl set_lebesgue_integral_eq_integral(2) set_integrable_norm *)
also have "\<dots> \<le> integral (\<Union>d) (\<lambda>x. norm (f x))"
using integrable_on_subdivision[OF d] assms f unfolding absolutely_integrable_on_def
by (subst integral_combine_division_topdown[OF _ d]) auto
also have "\<dots> \<le> integral UNIV (\<lambda>x. norm (f x))"
using integrable_on_subdivision[OF d] assms unfolding absolutely_integrable_on_def
by (intro integral_subset_le) auto
finally show "(\<Sum>k\<in>d. norm (integral k f)) \<le> integral UNIV (\<lambda>x. norm (f x))" .
qed
lemma absdiff_norm_less:
assumes "sum (\<lambda>x. norm (f x - g x)) s < e"
and "finite s"
shows "\<bar>sum (\<lambda>x. norm(f x)) s - sum (\<lambda>x. norm(g x)) s\<bar> < e"
unfolding sum_subtractf[symmetric]
apply (rule le_less_trans[OF sum_abs])
apply (rule le_less_trans[OF _ assms(1)])
apply (rule sum_mono)
apply (rule norm_triangle_ineq3)
done
proposition bounded_variation_absolutely_integrable_interval:
fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
assumes f: "f integrable_on cbox a b"
and *: "\<And>d. d division_of (cbox a b) \<Longrightarrow> sum (\<lambda>K. norm(integral K f)) d \<le> B"
shows "f absolutely_integrable_on cbox a b"
proof -
let ?f = "\<lambda>d. \<Sum>K\<in>d. norm (integral K f)" and ?D = "{d. d division_of (cbox a b)}"
have D_1: "?D \<noteq> {}"
by (rule elementary_interval[of a b]) auto
have D_2: "bdd_above (?f`?D)"
by (metis * mem_Collect_eq bdd_aboveI2)
note D = D_1 D_2
let ?S = "SUP x\<in>?D. ?f x"
have *: "\<exists>\<gamma>. gauge \<gamma> \<and>
(\<forall>p. p tagged_division_of cbox a b \<and>
\<gamma> fine p \<longrightarrow>
norm ((\<Sum>(x,k) \<in> p. content k *\<^sub>R norm (f x)) - ?S) < e)"
if e: "e > 0" for e
proof -
have "?S - e/2 < ?S" using \<open>e > 0\<close> by simp
then obtain d where d: "d division_of (cbox a b)" "?S - e/2 < (\<Sum>k\<in>d. norm (integral k f))"
unfolding less_cSUP_iff[OF D] by auto
note d' = division_ofD[OF this(1)]
have "\<exists>e>0. \<forall>i\<in>d. x \<notin> i \<longrightarrow> ball x e \<inter> i = {}" for x
proof -
have "\<exists>d'>0. \<forall>x'\<in>\<Union>{i \<in> d. x \<notin> i}. d' \<le> dist x x'"
proof (rule separate_point_closed)
show "closed (\<Union>{i \<in> d. x \<notin> i})"
using d' by force
show "x \<notin> \<Union>{i \<in> d. x \<notin> i}"
by auto
qed
then show ?thesis
by force
qed
then obtain k where k: "\<And>x. 0 < k x" "\<And>i x. \<lbrakk>i \<in> d; x \<notin> i\<rbrakk> \<Longrightarrow> ball x (k x) \<inter> i = {}"
by metis
have "e/2 > 0"
using e by auto
with Henstock_lemma[OF f]
obtain \<gamma> where g: "gauge \<gamma>"
"\<And>p. \<lbrakk>p tagged_partial_division_of cbox a b; \<gamma> fine p\<rbrakk>
\<Longrightarrow> (\<Sum>(x,k) \<in> p. norm (content k *\<^sub>R f x - integral k f)) < e/2"
by (metis (no_types, lifting))
let ?g = "\<lambda>x. \<gamma> x \<inter> ball x (k x)"
show ?thesis
proof (intro exI conjI allI impI)
show "gauge ?g"
using g(1) k(1) by (auto simp: gauge_def)
next
fix p
assume "p tagged_division_of (cbox a b) \<and> ?g fine p"
then have p: "p tagged_division_of cbox a b" "\<gamma> fine p" "(\<lambda>x. ball x (k x)) fine p"
by (auto simp: fine_Int)
note p' = tagged_division_ofD[OF p(1)]
define p' where "p' = {(x,k) | x k. \<exists>i l. x \<in> i \<and> i \<in> d \<and> (x,l) \<in> p \<and> k = i \<inter> l}"
have gp': "\<gamma> fine p'"
using p(2) by (auto simp: p'_def fine_def)
have p'': "p' tagged_division_of (cbox a b)"
proof (rule tagged_division_ofI)
show "finite p'"
proof (rule finite_subset)
show "p' \<subseteq> (\<lambda>(k, x, l). (x, k \<inter> l)) ` (d \<times> p)"
by (force simp: p'_def image_iff)
show "finite ((\<lambda>(k, x, l). (x, k \<inter> l)) ` (d \<times> p))"
by (simp add: d'(1) p'(1))
qed
next
fix x K
assume "(x, K) \<in> p'"
then have "\<exists>i l. x \<in> i \<and> i \<in> d \<and> (x, l) \<in> p \<and> K = i \<inter> l"
unfolding p'_def by auto
then obtain i l where il: "x \<in> i" "i \<in> d" "(x, l) \<in> p" "K = i \<inter> l" by blast
show "x \<in> K" and "K \<subseteq> cbox a b"
using p'(2-3)[OF il(3)] il by auto
show "\<exists>a b. K = cbox a b"
unfolding il using p'(4)[OF il(3)] d'(4)[OF il(2)] by (meson Int_interval)
next
fix x1 K1
assume "(x1, K1) \<in> p'"
then have "\<exists>i l. x1 \<in> i \<and> i \<in> d \<and> (x1, l) \<in> p \<and> K1 = i \<inter> l"
unfolding p'_def by auto
then obtain i1 l1 where il1: "x1 \<in> i1" "i1 \<in> d" "(x1, l1) \<in> p" "K1 = i1 \<inter> l1" by blast
fix x2 K2
assume "(x2,K2) \<in> p'"
then have "\<exists>i l. x2 \<in> i \<and> i \<in> d \<and> (x2, l) \<in> p \<and> K2 = i \<inter> l"
unfolding p'_def by auto
then obtain i2 l2 where il2: "x2 \<in> i2" "i2 \<in> d" "(x2, l2) \<in> p" "K2 = i2 \<inter> l2" by blast
assume "(x1, K1) \<noteq> (x2, K2)"
then have "interior i1 \<inter> interior i2 = {} \<or> interior l1 \<inter> interior l2 = {}"
using d'(5)[OF il1(2) il2(2)] p'(5)[OF il1(3) il2(3)] by (auto simp: il1 il2)
then show "interior K1 \<inter> interior K2 = {}"
unfolding il1 il2 by auto
next
have *: "\<forall>(x, X) \<in> p'. X \<subseteq> cbox a b"
unfolding p'_def using d' by blast
have "y \<in> \<Union>{K. \<exists>x. (x, K) \<in> p'}" if y: "y \<in> cbox a b" for y
proof -
obtain x l where xl: "(x, l) \<in> p" "y \<in> l"
using y unfolding p'(6)[symmetric] by auto
obtain i where i: "i \<in> d" "y \<in> i"
using y unfolding d'(6)[symmetric] by auto
have "x \<in> i"
using fineD[OF p(3) xl(1)] using k(2) i xl by auto
then show ?thesis
unfolding p'_def by (rule_tac X="i \<inter> l" in UnionI) (use i xl in auto)
qed
show "\<Union>{K. \<exists>x. (x, K) \<in> p'} = cbox a b"
proof
show "\<Union>{k. \<exists>x. (x, k) \<in> p'} \<subseteq> cbox a b"
using * by auto
next
show "cbox a b \<subseteq> \<Union>{k. \<exists>x. (x, k) \<in> p'}"
proof
fix y
assume y: "y \<in> cbox a b"
obtain x L where xl: "(x, L) \<in> p" "y \<in> L"
using y unfolding p'(6)[symmetric] by auto
obtain I where i: "I \<in> d" "y \<in> I"
using y unfolding d'(6)[symmetric] by auto
have "x \<in> I"
using fineD[OF p(3) xl(1)] using k(2) i xl by auto
then show "y \<in> \<Union>{k. \<exists>x. (x, k) \<in> p'}"
apply (rule_tac X="I \<inter> L" in UnionI)
using i xl by (auto simp: p'_def)
qed
qed
qed
then have sum_less_e2: "(\<Sum>(x,K) \<in> p'. norm (content K *\<^sub>R f x - integral K f)) < e/2"
using g(2) gp' tagged_division_of_def by blast
have "(x, I \<inter> L) \<in> p'" if x: "(x, L) \<in> p" "I \<in> d" and y: "y \<in> I" "y \<in> L"
for x I L y
proof -
have "x \<in> I"
using fineD[OF p(3) that(1)] k(2)[OF \<open>I \<in> d\<close>] y by auto
with x have "(\<exists>i l. x \<in> i \<and> i \<in> d \<and> (x, l) \<in> p \<and> I \<inter> L = i \<inter> l)"
by blast
then have "(x, I \<inter> L) \<in> p'"
by (simp add: p'_def)
with y show ?thesis by auto
qed
moreover have "\<exists>y i l. (x, K) = (y, i \<inter> l) \<and> (y, l) \<in> p \<and> i \<in> d \<and> i \<inter> l \<noteq> {}"
if xK: "(x,K) \<in> p'" for x K
proof -
obtain i l where il: "x \<in> i" "i \<in> d" "(x, l) \<in> p" "K = i \<inter> l"
using xK unfolding p'_def by auto
then show ?thesis
using p'(2) by fastforce
qed
ultimately have p'alt: "p' = {(x, I \<inter> L) | x I L. (x,L) \<in> p \<and> I \<in> d \<and> I \<inter> L \<noteq> {}}"
by auto
have sum_p': "(\<Sum>(x,K) \<in> p'. norm (integral K f)) = (\<Sum>k\<in>snd ` p'. norm (integral k f))"
apply (subst sum.over_tagged_division_lemma[OF p'',of "\<lambda>k. norm (integral k f)"])
apply (auto intro: integral_null simp: content_eq_0_interior)
done
have snd_p_div: "snd ` p division_of cbox a b"
by (rule division_of_tagged_division[OF p(1)])
note snd_p = division_ofD[OF snd_p_div]
have fin_d_sndp: "finite (d \<times> snd ` p)"
by (simp add: d'(1) snd_p(1))
have *: "\<And>sni sni' sf sf'. \<lbrakk>\<bar>sf' - sni'\<bar> < e/2; ?S - e/2 < sni; sni' \<le> ?S;
sni \<le> sni'; sf' = sf\<rbrakk> \<Longrightarrow> \<bar>sf - ?S\<bar> < e"
by arith
show "norm ((\<Sum>(x,k) \<in> p. content k *\<^sub>R norm (f x)) - ?S) < e"
unfolding real_norm_def
proof (rule *)
show "\<bar>(\<Sum>(x,K)\<in>p'. norm (content K *\<^sub>R f x)) - (\<Sum>(x,k)\<in>p'. norm (integral k f))\<bar> < e/2"
using p'' sum_less_e2 unfolding split_def by (force intro!: absdiff_norm_less)
show "(\<Sum>(x,k) \<in> p'. norm (integral k f)) \<le>?S"
by (auto simp: sum_p' division_of_tagged_division[OF p''] D intro!: cSUP_upper)
show "(\<Sum>k\<in>d. norm (integral k f)) \<le> (\<Sum>(x,k) \<in> p'. norm (integral k f))"
proof -
have *: "{k \<inter> l | k l. k \<in> d \<and> l \<in> snd ` p} = (\<lambda>(k,l). k \<inter> l) ` (d \<times> snd ` p)"
by auto
have "(\<Sum>K\<in>d. norm (integral K f)) \<le> (\<Sum>i\<in>d. \<Sum>l\<in>snd ` p. norm (integral (i \<inter> l) f))"
proof (rule sum_mono)
fix K assume k: "K \<in> d"
from d'(4)[OF this] obtain u v where uv: "K = cbox u v" by metis
define d' where "d' = {cbox u v \<inter> l |l. l \<in> snd ` p \<and> cbox u v \<inter> l \<noteq> {}}"
have uvab: "cbox u v \<subseteq> cbox a b"
using d(1) k uv by blast
have "d' division_of cbox u v"
unfolding d'_def by (rule division_inter_1 [OF snd_p_div uvab])
moreover then have "norm (\<Sum>i\<in>d'. integral i f) \<le> (\<Sum>k\<in>d'. norm (integral k f))"
by (simp add: sum_norm_le)
ultimately have "norm (integral K f) \<le> sum (\<lambda>k. norm (integral k f)) d'"
apply (subst integral_combine_division_topdown[of _ _ d'])
apply (auto simp: uv intro: integrable_on_subcbox[OF assms(1) uvab])
done
also have "\<dots> = (\<Sum>I\<in>{K \<inter> L |L. L \<in> snd ` p}. norm (integral I f))"
proof -
have *: "norm (integral I f) = 0"
if "I \<in> {cbox u v \<inter> l |l. l \<in> snd ` p}"
"I \<notin> {cbox u v \<inter> l |l. l \<in> snd ` p \<and> cbox u v \<inter> l \<noteq> {}}" for I
using that by auto
show ?thesis
apply (rule sum.mono_neutral_left)
apply (simp add: snd_p(1))
unfolding d'_def uv using * by auto
qed
also have "\<dots> = (\<Sum>l\<in>snd ` p. norm (integral (K \<inter> l) f))"
proof -
have *: "norm (integral (K \<inter> l) f) = 0"
if "l \<in> snd ` p" "y \<in> snd ` p" "l \<noteq> y" "K \<inter> l = K \<inter> y" for l y
proof -
have "interior (K \<inter> l) \<subseteq> interior (l \<inter> y)"
by (metis Int_lower2 interior_mono le_inf_iff that(4))
then have "interior (K \<inter> l) = {}"
by (simp add: snd_p(5) that)
moreover from d'(4)[OF k] snd_p(4)[OF that(1)]
obtain u1 v1 u2 v2
where uv: "K = cbox u1 u2" "l = cbox v1 v2" by metis
ultimately show ?thesis
using that integral_null
unfolding uv Int_interval content_eq_0_interior
by (metis (mono_tags, lifting) norm_eq_zero)
qed
show ?thesis
unfolding Setcompr_eq_image
apply (rule sum.reindex_nontrivial [unfolded o_def])
apply (rule finite_imageI)
apply (rule p')
using * by auto
qed
finally show "norm (integral K f) \<le> (\<Sum>l\<in>snd ` p. norm (integral (K \<inter> l) f))" .
qed
also have "\<dots> = (\<Sum>(i,l) \<in> d \<times> snd ` p. norm (integral (i\<inter>l) f))"
by (simp add: sum.cartesian_product)
also have "\<dots> = (\<Sum>x \<in> d \<times> snd ` p. norm (integral (case_prod (\<inter>) x) f))"
by (force simp: split_def intro!: sum.cong)
also have "\<dots> = (\<Sum>k\<in>{i \<inter> l |i l. i \<in> d \<and> l \<in> snd ` p}. norm (integral k f))"
proof -
have eq0: " (integral (l1 \<inter> k1) f) = 0"
if "l1 \<inter> k1 = l2 \<inter> k2" "(l1, k1) \<noteq> (l2, k2)"
"l1 \<in> d" "(j1,k1) \<in> p" "l2 \<in> d" "(j2,k2) \<in> p"
for l1 l2 k1 k2 j1 j2
proof -
obtain u1 v1 u2 v2 where uv: "l1 = cbox u1 u2" "k1 = cbox v1 v2"
using \<open>(j1, k1) \<in> p\<close> \<open>l1 \<in> d\<close> d'(4) p'(4) by blast
have "l1 \<noteq> l2 \<or> k1 \<noteq> k2"
using that by auto
then have "interior k1 \<inter> interior k2 = {} \<or> interior l1 \<inter> interior l2 = {}"
by (meson d'(5) old.prod.inject p'(5) that(3) that(4) that(5) that(6))
moreover have "interior (l1 \<inter> k1) = interior (l2 \<inter> k2)"
by (simp add: that(1))
ultimately have "interior(l1 \<inter> k1) = {}"
by auto
then show ?thesis
unfolding uv Int_interval content_eq_0_interior[symmetric] by auto
qed
show ?thesis
unfolding *
apply (rule sum.reindex_nontrivial [OF fin_d_sndp, symmetric, unfolded o_def])
apply clarsimp
by (metis eq0 fst_conv snd_conv)
qed
also have "\<dots> = (\<Sum>(x,k) \<in> p'. norm (integral k f))"
proof -
have 0: "integral (ia \<inter> snd (a, b)) f = 0"
if "ia \<inter> snd (a, b) \<notin> snd ` p'" "ia \<in> d" "(a, b) \<in> p" for ia a b
proof -
have "ia \<inter> b = {}"
using that unfolding p'alt image_iff Bex_def not_ex
apply (erule_tac x="(a, ia \<inter> b)" in allE)
apply auto
done
then show ?thesis by auto
qed
have 1: "\<exists>i l. snd (a, b) = i \<inter> l \<and> i \<in> d \<and> l \<in> snd ` p" if "(a, b) \<in> p'" for a b
using that
apply (clarsimp simp: p'_def image_iff)
by (metis (no_types, hide_lams) snd_conv)
show ?thesis
unfolding sum_p'
apply (rule sum.mono_neutral_right)
apply (metis * finite_imageI[OF fin_d_sndp])
using 0 1 by auto
qed
finally show ?thesis .
qed
show "(\<Sum>(x,k) \<in> p'. norm (content k *\<^sub>R f x)) = (\<Sum>(x,k) \<in> p. content k *\<^sub>R norm (f x))"
proof -
let ?S = "{(x, i \<inter> l) |x i l. (x, l) \<in> p \<and> i \<in> d}"
have *: "?S = (\<lambda>(xl,i). (fst xl, snd xl \<inter> i)) ` (p \<times> d)"
by force
have fin_pd: "finite (p \<times> d)"
using finite_cartesian_product[OF p'(1) d'(1)] by metis
have "(\<Sum>(x,k) \<in> p'. norm (content k *\<^sub>R f x)) = (\<Sum>(x,k) \<in> ?S. \<bar>content k\<bar> * norm (f x))"
unfolding norm_scaleR
apply (rule sum.mono_neutral_left)
apply (subst *)
apply (rule finite_imageI [OF fin_pd])
unfolding p'alt apply auto
by fastforce
also have "\<dots> = (\<Sum>((x,l),i)\<in>p \<times> d. \<bar>content (l \<inter> i)\<bar> * norm (f x))"
proof -
have "\<bar>content (l1 \<inter> k1)\<bar> * norm (f x1) = 0"
if "(x1, l1) \<in> p" "(x2, l2) \<in> p" "k1 \<in> d" "k2 \<in> d"
"x1 = x2" "l1 \<inter> k1 = l2 \<inter> k2" "x1 \<noteq> x2 \<or> l1 \<noteq> l2 \<or> k1 \<noteq> k2"
for x1 l1 k1 x2 l2 k2
proof -
obtain u1 v1 u2 v2 where uv: "k1 = cbox u1 u2" "l1 = cbox v1 v2"
by (meson \<open>(x1, l1) \<in> p\<close> \<open>k1 \<in> d\<close> d(1) division_ofD(4) p'(4))
have "l1 \<noteq> l2 \<or> k1 \<noteq> k2"
using that by auto
then have "interior k1 \<inter> interior k2 = {} \<or> interior l1 \<inter> interior l2 = {}"
apply (rule disjE)
using that p'(5) d'(5) by auto
moreover have "interior (l1 \<inter> k1) = interior (l2 \<inter> k2)"
unfolding that ..
ultimately have "interior (l1 \<inter> k1) = {}"
by auto
then show "\<bar>content (l1 \<inter> k1)\<bar> * norm (f x1) = 0"
unfolding uv Int_interval content_eq_0_interior[symmetric] by auto
qed
then show ?thesis
unfolding *
apply (subst sum.reindex_nontrivial [OF fin_pd])
unfolding split_paired_all o_def split_def prod.inject
apply force+
done
qed
also have "\<dots> = (\<Sum>(x,k) \<in> p. content k *\<^sub>R norm (f x))"
proof -
have sumeq: "(\<Sum>i\<in>d. content (l \<inter> i) * norm (f x)) = content l * norm (f x)"
if "(x, l) \<in> p" for x l
proof -
note xl = p'(2-4)[OF that]
then obtain u v where uv: "l = cbox u v" by blast
have "(\<Sum>i\<in>d. \<bar>content (l \<inter> i)\<bar>) = (\<Sum>k\<in>d. content (k \<inter> cbox u v))"
by (simp add: Int_commute uv)
also have "\<dots> = sum content {k \<inter> cbox u v| k. k \<in> d}"
proof -
have eq0: "content (k \<inter> cbox u v) = 0"
if "k \<in> d" "y \<in> d" "k \<noteq> y" and eq: "k \<inter> cbox u v = y \<inter> cbox u v" for k y
proof -
from d'(4)[OF that(1)] d'(4)[OF that(2)]
obtain \<alpha> \<beta> where \<alpha>: "k \<inter> cbox u v = cbox \<alpha> \<beta>"
by (meson Int_interval)
have "{} = interior ((k \<inter> y) \<inter> cbox u v)"
by (simp add: d'(5) that)
also have "\<dots> = interior (y \<inter> (k \<inter> cbox u v))"
by auto
also have "\<dots> = interior (k \<inter> cbox u v)"
unfolding eq by auto
finally show ?thesis
unfolding \<alpha> content_eq_0_interior ..
qed
then show ?thesis
unfolding Setcompr_eq_image
apply (rule sum.reindex_nontrivial [OF \<open>finite d\<close>, unfolded o_def, symmetric])
by auto
qed
also have "\<dots> = sum content {cbox u v \<inter> k |k. k \<in> d \<and> cbox u v \<inter> k \<noteq> {}}"
apply (rule sum.mono_neutral_right)
unfolding Setcompr_eq_image
apply (rule finite_imageI [OF \<open>finite d\<close>])
apply (fastforce simp: inf.commute)+
done
finally show "(\<Sum>i\<in>d. content (l \<inter> i) * norm (f x)) = content l * norm (f x)"
unfolding sum_distrib_right[symmetric] real_scaleR_def
apply (subst(asm) additive_content_division[OF division_inter_1[OF d(1)]])
using xl(2)[unfolded uv] unfolding uv apply auto
done
qed
show ?thesis
by (subst sum_Sigma_product[symmetric]) (auto intro!: sumeq sum.cong p' d')
qed
finally show ?thesis .
qed
qed (rule d)
qed
qed
then show ?thesis
using absolutely_integrable_onI [OF f has_integral_integrable] has_integral[of _ ?S]
by blast
qed
lemma bounded_variation_absolutely_integrable:
fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
assumes "f integrable_on UNIV"
and "\<forall>d. d division_of (\<Union>d) \<longrightarrow> sum (\<lambda>k. norm (integral k f)) d \<le> B"
shows "f absolutely_integrable_on UNIV"
proof (rule absolutely_integrable_onI, fact)
let ?f = "\<lambda>d. \<Sum>k\<in>d. norm (integral k f)" and ?D = "{d. d division_of (\<Union>d)}"
have D_1: "?D \<noteq> {}"
by (rule elementary_interval) auto
have D_2: "bdd_above (?f`?D)"
by (intro bdd_aboveI2[where M=B] assms(2)[rule_format]) simp
note D = D_1 D_2
let ?S = "SUP d\<in>?D. ?f d"
have "\<And>a b. f integrable_on cbox a b"
using assms(1) integrable_on_subcbox by blast
then have f_int: "\<And>a b. f absolutely_integrable_on cbox a b"
apply (rule bounded_variation_absolutely_integrable_interval[where B=B])
using assms(2) apply blast
done
have "((\<lambda>x. norm (f x)) has_integral ?S) UNIV"
apply (subst has_integral_alt')
apply safe
proof goal_cases
case (1 a b)
show ?case
using f_int[of a b] unfolding absolutely_integrable_on_def by auto
next
case prems: (2 e)
have "\<exists>y\<in>sum (\<lambda>k. norm (integral k f)) ` {d. d division_of \<Union>d}. \<not> y \<le> ?S - e"
proof (rule ccontr)
assume "\<not> ?thesis"
then have "?S \<le> ?S - e"
by (intro cSUP_least[OF D(1)]) auto
then show False
using prems by auto
qed
then obtain d K where ddiv: "d division_of \<Union>d" and "K = (\<Sum>k\<in>d. norm (integral k f))"
"Sup (sum (\<lambda>k. norm (integral k f)) ` {d. d division_of \<Union> d}) - e < K"
by (auto simp add: image_iff not_le)
then have d: "Sup (sum (\<lambda>k. norm (integral k f)) ` {d. d division_of \<Union> d}) - e
< (\<Sum>k\<in>d. norm (integral k f))"
by auto
note d'=division_ofD[OF ddiv]
have "bounded (\<Union>d)"
by (rule elementary_bounded,fact)
from this[unfolded bounded_pos] obtain K where
K: "0 < K" "\<forall>x\<in>\<Union>d. norm x \<le> K" by auto
show ?case
proof (intro conjI impI allI exI)
fix a b :: 'n
assume ab: "ball 0 (K + 1) \<subseteq> cbox a b"
have *: "\<And>s s1. \<lbrakk>?S - e < s1; s1 \<le> s; s < ?S + e\<rbrakk> \<Longrightarrow> \<bar>s - ?S\<bar> < e"
by arith
show "norm (integral (cbox a b) (\<lambda>x. if x \<in> UNIV then norm (f x) else 0) - ?S) < e"
unfolding real_norm_def
proof (rule * [OF d])
have "(\<Sum>k\<in>d. norm (integral k f)) \<le> sum (\<lambda>k. integral k (\<lambda>x. norm (f x))) d"
proof (intro sum_mono)
fix k assume "k \<in> d"
with d'(4) f_int show "norm (integral k f) \<le> integral k (\<lambda>x. norm (f x))"
by (force simp: absolutely_integrable_on_def integral_norm_bound_integral)
qed
also have "\<dots> = integral (\<Union>d) (\<lambda>x. norm (f x))"
apply (rule integral_combine_division_bottomup[OF ddiv, symmetric])
using absolutely_integrable_on_def d'(4) f_int by blast
also have "\<dots> \<le> integral (cbox a b) (\<lambda>x. if x \<in> UNIV then norm (f x) else 0)"
proof -
have "\<Union>d \<subseteq> cbox a b"
using K(2) ab by fastforce
then show ?thesis
using integrable_on_subdivision[OF ddiv] f_int[of a b] unfolding absolutely_integrable_on_def
by (auto intro!: integral_subset_le)
qed
finally show "(\<Sum>k\<in>d. norm (integral k f))
\<le> integral (cbox a b) (\<lambda>x. if x \<in> UNIV then norm (f x) else 0)" .
next
have "e/2>0"
using \<open>e > 0\<close> by auto
moreover
have f: "f integrable_on cbox a b" "(\<lambda>x. norm (f x)) integrable_on cbox a b"
using f_int by (auto simp: absolutely_integrable_on_def)
ultimately obtain d1 where "gauge d1"
and d1: "\<And>p. \<lbrakk>p tagged_division_of (cbox a b); d1 fine p\<rbrakk> \<Longrightarrow>
norm ((\<Sum>(x,k) \<in> p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\<lambda>x. norm (f x))) < e/2"
unfolding has_integral_integral has_integral by meson
obtain d2 where "gauge d2"
and d2: "\<And>p. \<lbrakk>p tagged_partial_division_of (cbox a b); d2 fine p\<rbrakk> \<Longrightarrow>
(\<Sum>(x,k) \<in> p. norm (content k *\<^sub>R f x - integral k f)) < e/2"
by (blast intro: Henstock_lemma [OF f(1) \<open>e/2>0\<close>])
obtain p where
p: "p tagged_division_of (cbox a b)" "d1 fine p" "d2 fine p"
by (rule fine_division_exists [OF gauge_Int [OF \<open>gauge d1\<close> \<open>gauge d2\<close>], of a b])
(auto simp add: fine_Int)
have *: "\<And>sf sf' si di. \<lbrakk>sf' = sf; si \<le> ?S; \<bar>sf - si\<bar> < e/2;
\<bar>sf' - di\<bar> < e/2\<rbrakk> \<Longrightarrow> di < ?S + e"
by arith
have "integral (cbox a b) (\<lambda>x. norm (f x)) < ?S + e"
proof (rule *)
show "\<bar>(\<Sum>(x,k)\<in>p. norm (content k *\<^sub>R f x)) - (\<Sum>(x,k)\<in>p. norm (integral k f))\<bar> < e/2"
unfolding split_def
apply (rule absdiff_norm_less)
using d2[of p] p(1,3) apply (auto simp: tagged_division_of_def split_def)
done
show "\<bar>(\<Sum>(x,k) \<in> p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\<lambda>x. norm(f x))\<bar> < e/2"
using d1[OF p(1,2)] by (simp only: real_norm_def)
show "(\<Sum>(x,k) \<in> p. content k *\<^sub>R norm (f x)) = (\<Sum>(x,k) \<in> p. norm (content k *\<^sub>R f x))"
by (auto simp: split_paired_all sum.cong [OF refl])
show "(\<Sum>(x,k) \<in> p. norm (integral k f)) \<le> ?S"
using partial_division_of_tagged_division[of p "cbox a b"] p(1)
apply (subst sum.over_tagged_division_lemma[OF p(1)])
apply (auto simp: content_eq_0_interior tagged_partial_division_of_def intro!: cSUP_upper2 D)
done
qed
then show "integral (cbox a b) (\<lambda>x. if x \<in> UNIV then norm (f x) else 0) < ?S + e"
by simp
qed
qed (insert K, auto)
qed
then show "(\<lambda>x. norm (f x)) integrable_on UNIV"
by blast
qed
subsection\<open>Outer and inner approximation of measurable sets by well-behaved sets.\<close>
proposition measurable_outer_intervals_bounded:
assumes "S \<in> lmeasurable" "S \<subseteq> cbox a b" "e > 0"
obtains \<D>
where "countable \<D>"
"\<And>K. K \<in> \<D> \<Longrightarrow> K \<subseteq> cbox a b \<and> K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
"pairwise (\<lambda>A B. interior A \<inter> interior B = {}) \<D>"
"\<And>u v. cbox u v \<in> \<D> \<Longrightarrow> \<exists>n. \<forall>i \<in> Basis. v \<bullet> i - u \<bullet> i = (b \<bullet> i - a \<bullet> i)/2^n"
"\<And>K. \<lbrakk>K \<in> \<D>; box a b \<noteq> {}\<rbrakk> \<Longrightarrow> interior K \<noteq> {}"
"S \<subseteq> \<Union>\<D>" "\<Union>\<D> \<in> lmeasurable" "measure lebesgue (\<Union>\<D>) \<le> measure lebesgue S + e"
proof (cases "box a b = {}")
case True
show ?thesis
proof (cases "cbox a b = {}")
case True
with assms have [simp]: "S = {}"
by auto
show ?thesis
proof
show "countable {}"
by simp
qed (use \<open>e > 0\<close> in auto)
next
case False
show ?thesis
proof
show "countable {cbox a b}"
by simp
show "\<And>u v. cbox u v \<in> {cbox a b} \<Longrightarrow> \<exists>n. \<forall>i\<in>Basis. v \<bullet> i - u \<bullet> i = (b \<bullet> i - a \<bullet> i)/2 ^ n"
using False by (force simp: eq_cbox intro: exI [where x=0])
show "measure lebesgue (\<Union>{cbox a b}) \<le> measure lebesgue S + e"
using assms by (simp add: sum_content.box_empty_imp [OF True])
qed (use assms \<open>cbox a b \<noteq> {}\<close> in auto)
qed
next
case False
let ?\<mu> = "measure lebesgue"
have "S \<inter> cbox a b \<in> lmeasurable"
using \<open>S \<in> lmeasurable\<close> by blast
then have indS_int: "(indicator S has_integral (?\<mu> S)) (cbox a b)"
by (metis integral_indicator \<open>S \<subseteq> cbox a b\<close> has_integral_integrable_integral inf.orderE integrable_on_indicator)
with \<open>e > 0\<close> obtain \<gamma> where "gauge \<gamma>" and \<gamma>:
"\<And>\<D>. \<lbrakk>\<D> tagged_division_of (cbox a b); \<gamma> fine \<D>\<rbrakk> \<Longrightarrow> norm ((\<Sum>(x,K)\<in>\<D>. content(K) *\<^sub>R indicator S x) - ?\<mu> S) < e"
by (force simp: has_integral)
have inteq: "integral (cbox a b) (indicat_real S) = integral UNIV (indicator S)"
using assms by (metis has_integral_iff indS_int lmeasure_integral_UNIV)
obtain \<D> where \<D>: "countable \<D>" "\<Union>\<D> \<subseteq> cbox a b"
and cbox: "\<And>K. K \<in> \<D> \<Longrightarrow> interior K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
and djointish: "pairwise (\<lambda>A B. interior A \<inter> interior B = {}) \<D>"
and covered: "\<And>K. K \<in> \<D> \<Longrightarrow> \<exists>x \<in> S \<inter> K. K \<subseteq> \<gamma> x"
and close: "\<And>u v. cbox u v \<in> \<D> \<Longrightarrow> \<exists>n. \<forall>i \<in> Basis. v\<bullet>i - u\<bullet>i = (b\<bullet>i - a\<bullet>i)/2^n"
and covers: "S \<subseteq> \<Union>\<D>"
using covering_lemma [of S a b \<gamma>] \<open>gauge \<gamma>\<close> \<open>box a b \<noteq> {}\<close> assms by force
show ?thesis
proof
show "\<And>K. K \<in> \<D> \<Longrightarrow> K \<subseteq> cbox a b \<and> K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
by (meson Sup_le_iff \<D>(2) cbox interior_empty)
have negl_int: "negligible(K \<inter> L)" if "K \<in> \<D>" "L \<in> \<D>" "K \<noteq> L" for K L
proof -
have "interior K \<inter> interior L = {}"
using djointish pairwiseD that by fastforce
moreover obtain u v x y where "K = cbox u v" "L = cbox x y"
using cbox \<open>K \<in> \<D>\<close> \<open>L \<in> \<D>\<close> by blast
ultimately show ?thesis
by (simp add: Int_interval box_Int_box negligible_interval(1))
qed
have fincase: "\<Union>\<F> \<in> lmeasurable \<and> ?\<mu> (\<Union>\<F>) \<le> ?\<mu> S + e" if "finite \<F>" "\<F> \<subseteq> \<D>" for \<F>
proof -
obtain t where t: "\<And>K. K \<in> \<F> \<Longrightarrow> t K \<in> S \<inter> K \<and> K \<subseteq> \<gamma>(t K)"
using covered \<open>\<F> \<subseteq> \<D>\<close> subsetD by metis
have "\<forall>K \<in> \<F>. \<forall>L \<in> \<F>. K \<noteq> L \<longrightarrow> interior K \<inter> interior L = {}"
using that djointish by (simp add: pairwise_def) (metis subsetD)
with cbox that \<D> have \<F>div: "\<F> division_of (\<Union>\<F>)"
by (fastforce simp: division_of_def dest: cbox)
then have 1: "\<Union>\<F> \<in> lmeasurable"
by blast
have norme: "\<And>p. \<lbrakk>p tagged_division_of cbox a b; \<gamma> fine p\<rbrakk>
\<Longrightarrow> norm ((\<Sum>(x,K)\<in>p. content K * indicator S x) - integral (cbox a b) (indicator S)) < e"
by (auto simp: lmeasure_integral_UNIV assms inteq dest: \<gamma>)
have "\<forall>x K y L. (x,K) \<in> (\<lambda>K. (t K,K)) ` \<F> \<and> (y,L) \<in> (\<lambda>K. (t K,K)) ` \<F> \<and> (x,K) \<noteq> (y,L) \<longrightarrow> interior K \<inter> interior L = {}"
using that djointish by (clarsimp simp: pairwise_def) (metis subsetD)
with that \<D> have tagged: "(\<lambda>K. (t K, K)) ` \<F> tagged_partial_division_of cbox a b"
by (auto simp: tagged_partial_division_of_def dest: t cbox)
have fine: "\<gamma> fine (\<lambda>K. (t K, K)) ` \<F>"
using t by (auto simp: fine_def)
have *: "y \<le> ?\<mu> S \<Longrightarrow> \<bar>x - y\<bar> \<le> e \<Longrightarrow> x \<le> ?\<mu> S + e" for x y
by arith
have "?\<mu> (\<Union>\<F>) \<le> ?\<mu> S + e"
proof (rule *)
have "(\<Sum>K\<in>\<F>. ?\<mu> (K \<inter> S)) = ?\<mu> (\<Union>C\<in>\<F>. C \<inter> S)"
apply (rule measure_negligible_finite_Union_image [OF \<open>finite \<F>\<close>, symmetric])
using \<F>div \<open>S \<in> lmeasurable\<close> apply blast
unfolding pairwise_def
by (metis inf.commute inf_sup_aci(3) negligible_Int subsetCE negl_int \<open>\<F> \<subseteq> \<D>\<close>)
also have "\<dots> = ?\<mu> (\<Union>\<F> \<inter> S)"
by simp
also have "\<dots> \<le> ?\<mu> S"
by (simp add: "1" \<open>S \<in> lmeasurable\<close> fmeasurableD measure_mono_fmeasurable sets.Int)
finally show "(\<Sum>K\<in>\<F>. ?\<mu> (K \<inter> S)) \<le> ?\<mu> S" .
next
have "?\<mu> (\<Union>\<F>) = sum ?\<mu> \<F>"
by (metis \<F>div content_division)
also have "\<dots> = (\<Sum>K\<in>\<F>. content K)"
using \<F>div by (force intro: sum.cong)
also have "\<dots> = (\<Sum>x\<in>\<F>. content x * indicator S (t x))"
using t by auto
finally have eq1: "?\<mu> (\<Union>\<F>) = (\<Sum>x\<in>\<F>. content x * indicator S (t x))" .
have eq2: "(\<Sum>K\<in>\<F>. ?\<mu> (K \<inter> S)) = (\<Sum>K\<in>\<F>. integral K (indicator S))"
apply (rule sum.cong [OF refl])
by (metis integral_indicator \<F>div \<open>S \<in> lmeasurable\<close> division_ofD(4) fmeasurable.Int inf.commute lmeasurable_cbox)
have "\<bar>\<Sum>(x,K)\<in>(\<lambda>K. (t K, K)) ` \<F>. content K * indicator S x - integral K (indicator S)\<bar> \<le> e"
using Henstock_lemma_part1 [of "indicator S::'a\<Rightarrow>real", OF _ \<open>e > 0\<close> \<open>gauge \<gamma>\<close> _ tagged fine]
indS_int norme by auto
then show "\<bar>?\<mu> (\<Union>\<F>) - (\<Sum>K\<in>\<F>. ?\<mu> (K \<inter> S))\<bar> \<le> e"
by (simp add: eq1 eq2 comm_monoid_add_class.sum.reindex inj_on_def sum_subtractf)
qed
with 1 show ?thesis by blast
qed
have "\<Union>\<D> \<in> lmeasurable \<and> ?\<mu> (\<Union>\<D>) \<le> ?\<mu> S + e"
proof (cases "finite \<D>")
case True
with fincase show ?thesis
by blast
next
case False
let ?T = "from_nat_into \<D>"
have T: "bij_betw ?T UNIV \<D>"
by (simp add: False \<D>(1) bij_betw_from_nat_into)
have TM: "\<And>n. ?T n \<in> lmeasurable"
by (metis False cbox finite.emptyI from_nat_into lmeasurable_cbox)
have TN: "\<And>m n. m \<noteq> n \<Longrightarrow> negligible (?T m \<inter> ?T n)"
by (simp add: False \<D>(1) from_nat_into infinite_imp_nonempty negl_int)
have TB: "(\<Sum>k\<le>n. ?\<mu> (?T k)) \<le> ?\<mu> S + e" for n
proof -
have "(\<Sum>k\<le>n. ?\<mu> (?T k)) = ?\<mu> (\<Union> (?T ` {..n}))"
by (simp add: pairwise_def TM TN measure_negligible_finite_Union_image)
also have "?\<mu> (\<Union> (?T ` {..n})) \<le> ?\<mu> S + e"
using fincase [of "?T ` {..n}"] T by (auto simp: bij_betw_def)
finally show ?thesis .
qed
have "\<Union>\<D> \<in> lmeasurable"
by (metis lmeasurable_compact T \<D>(2) bij_betw_def cbox compact_cbox countable_Un_Int(1) fmeasurableD fmeasurableI2 rangeI)
moreover
have "?\<mu> (\<Union>x. from_nat_into \<D> x) \<le> ?\<mu> S + e"
proof (rule measure_countable_Union_le [OF TM])
show "?\<mu> (\<Union>x\<le>n. from_nat_into \<D> x) \<le> ?\<mu> S + e" for n
by (metis (mono_tags, lifting) False fincase finite.emptyI finite_atMost finite_imageI from_nat_into imageE subsetI)
qed
ultimately show ?thesis by (metis T bij_betw_def)
qed
then show "\<Union>\<D> \<in> lmeasurable" "measure lebesgue (\<Union>\<D>) \<le> ?\<mu> S + e" by blast+
qed (use \<D> cbox djointish close covers in auto)
qed
subsection\<open>Transformation of measure by linear maps\<close>
lemma emeasure_lebesgue_ball_conv_unit_ball:
fixes c :: "'a :: euclidean_space"
assumes "r \<ge> 0"
shows "emeasure lebesgue (ball c r) =
ennreal (r ^ DIM('a)) * emeasure lebesgue (ball (0 :: 'a) 1)"
proof (cases "r = 0")
case False
with assms have r: "r > 0" by auto
have "emeasure lebesgue ((\<lambda>x. c + x) ` (\<lambda>x. r *\<^sub>R x) ` ball (0 :: 'a) 1) =
r ^ DIM('a) * emeasure lebesgue (ball (0 :: 'a) 1)"
unfolding image_image using emeasure_lebesgue_affine[of r c "ball 0 1"] assms
by (simp add: add_ac)
also have "(\<lambda>x. r *\<^sub>R x) ` ball 0 1 = ball (0 :: 'a) r"
using r by (subst ball_scale) auto
also have "(\<lambda>x. c + x) ` \<dots> = ball c r"
by (subst image_add_ball) (simp_all add: algebra_simps)
finally show ?thesis by simp
qed auto
lemma content_ball_conv_unit_ball:
fixes c :: "'a :: euclidean_space"
assumes "r \<ge> 0"
shows "content (ball c r) = r ^ DIM('a) * content (ball (0 :: 'a) 1)"
proof -
have "ennreal (content (ball c r)) = emeasure lebesgue (ball c r)"
using emeasure_lborel_ball_finite[of c r] by (subst emeasure_eq_ennreal_measure) auto
also have "\<dots> = ennreal (r ^ DIM('a)) * emeasure lebesgue (ball (0 :: 'a) 1)"
using assms by (intro emeasure_lebesgue_ball_conv_unit_ball) auto
also have "\<dots> = ennreal (r ^ DIM('a) * content (ball (0::'a) 1))"
using emeasure_lborel_ball_finite[of "0::'a" 1] assms
by (subst emeasure_eq_ennreal_measure) (auto simp: ennreal_mult')
finally show ?thesis
using assms by (subst (asm) ennreal_inj) auto
qed
lemma measurable_linear_image_interval:
"linear f \<Longrightarrow> f ` (cbox a b) \<in> lmeasurable"
by (metis bounded_linear_image linear_linear bounded_cbox closure_bounded_linear_image closure_cbox compact_closure lmeasurable_compact)
proposition measure_linear_sufficient:
fixes f :: "'n::euclidean_space \<Rightarrow> 'n"
assumes "linear f" and S: "S \<in> lmeasurable"
and im: "\<And>a b. measure lebesgue (f ` (cbox a b)) = m * measure lebesgue (cbox a b)"
shows "f ` S \<in> lmeasurable \<and> m * measure lebesgue S = measure lebesgue (f ` S)"
using le_less_linear [of 0 m]
proof
assume "m < 0"
then show ?thesis
using im [of 0 One] by auto
next
assume "m \<ge> 0"
let ?\<mu> = "measure lebesgue"
show ?thesis
proof (cases "inj f")
case False
then have "?\<mu> (f ` S) = 0"
using \<open>linear f\<close> negligible_imp_measure0 negligible_linear_singular_image by blast
then have "m * ?\<mu> (cbox 0 (One)) = 0"
by (metis False \<open>linear f\<close> cbox_borel content_unit im measure_completion negligible_imp_measure0 negligible_linear_singular_image sets_lborel)
then show ?thesis
using \<open>linear f\<close> negligible_linear_singular_image negligible_imp_measure0 False
by (auto simp: lmeasurable_iff_has_integral negligible_UNIV)
next
case True
then obtain h where "linear h" and hf: "\<And>x. h (f x) = x" and fh: "\<And>x. f (h x) = x"
using \<open>linear f\<close> linear_injective_isomorphism by blast
have fBS: "(f ` S) \<in> lmeasurable \<and> m * ?\<mu> S = ?\<mu> (f ` S)"
if "bounded S" "S \<in> lmeasurable" for S
proof -
obtain a b where "S \<subseteq> cbox a b"
using \<open>bounded S\<close> bounded_subset_cbox_symmetric by metis
have fUD: "(f ` \<Union>\<D>) \<in> lmeasurable \<and> ?\<mu> (f ` \<Union>\<D>) = (m * ?\<mu> (\<Union>\<D>))"
if "countable \<D>"
and cbox: "\<And>K. K \<in> \<D> \<Longrightarrow> K \<subseteq> cbox a b \<and> K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
and intint: "pairwise (\<lambda>A B. interior A \<inter> interior B = {}) \<D>"
for \<D>
proof -
have conv: "\<And>K. K \<in> \<D> \<Longrightarrow> convex K"
using cbox convex_box(1) by blast
have neg: "negligible (g ` K \<inter> g ` L)" if "linear g" "K \<in> \<D>" "L \<in> \<D>" "K \<noteq> L"
for K L and g :: "'n\<Rightarrow>'n"
proof (cases "inj g")
case True
have "negligible (frontier(g ` K \<inter> g ` L) \<union> interior(g ` K \<inter> g ` L))"
proof (rule negligible_Un)
show "negligible (frontier (g ` K \<inter> g ` L))"
by (simp add: negligible_convex_frontier convex_Int conv convex_linear_image that)
next
have "\<forall>p N. pairwise p N = (\<forall>Na. (Na::'n set) \<in> N \<longrightarrow> (\<forall>Nb. Nb \<in> N \<and> Na \<noteq> Nb \<longrightarrow> p Na Nb))"
by (metis pairwise_def)
then have "interior K \<inter> interior L = {}"
using intint that(2) that(3) that(4) by presburger
then show "negligible (interior (g ` K \<inter> g ` L))"
by (metis True empty_imp_negligible image_Int image_empty interior_Int interior_injective_linear_image that(1))
qed
moreover have "g ` K \<inter> g ` L \<subseteq> frontier (g ` K \<inter> g ` L) \<union> interior (g ` K \<inter> g ` L)"
apply (auto simp: frontier_def)
using closure_subset contra_subsetD by fastforce+
ultimately show ?thesis
by (rule negligible_subset)
next
case False
then show ?thesis
by (simp add: negligible_Int negligible_linear_singular_image \<open>linear g\<close>)
qed
have negf: "negligible ((f ` K) \<inter> (f ` L))"
and negid: "negligible (K \<inter> L)" if "K \<in> \<D>" "L \<in> \<D>" "K \<noteq> L" for K L
using neg [OF \<open>linear f\<close>] neg [OF linear_id] that by auto
show ?thesis
proof (cases "finite \<D>")
case True
then have "?\<mu> (\<Union>x\<in>\<D>. f ` x) = (\<Sum>x\<in>\<D>. ?\<mu> (f ` x))"
using \<open>linear f\<close> cbox measurable_linear_image_interval negf
by (blast intro: measure_negligible_finite_Union_image [unfolded pairwise_def])
also have "\<dots> = (\<Sum>k\<in>\<D>. m * ?\<mu> k)"
by (metis (no_types, lifting) cbox im sum.cong)
also have "\<dots> = m * ?\<mu> (\<Union>\<D>)"
unfolding sum_distrib_left [symmetric]
by (metis True cbox lmeasurable_cbox measure_negligible_finite_Union [unfolded pairwise_def] negid)
finally show ?thesis
by (metis True \<open>linear f\<close> cbox image_Union fmeasurable.finite_UN measurable_linear_image_interval)
next
case False
with \<open>countable \<D>\<close> obtain X :: "nat \<Rightarrow> 'n set" where S: "bij_betw X UNIV \<D>"
using bij_betw_from_nat_into by blast
then have eq: "(\<Union>\<D>) = (\<Union>n. X n)" "(f ` \<Union>\<D>) = (\<Union>n. f ` X n)"
by (auto simp: bij_betw_def)
have meas: "\<And>K. K \<in> \<D> \<Longrightarrow> K \<in> lmeasurable"
using cbox by blast
with S have 1: "\<And>n. X n \<in> lmeasurable"
by (auto simp: bij_betw_def)
have 2: "pairwise (\<lambda>m n. negligible (X m \<inter> X n)) UNIV"
using S unfolding bij_betw_def pairwise_def by (metis injD negid range_eqI)
have "bounded (\<Union>\<D>)"
by (meson Sup_least bounded_cbox bounded_subset cbox)
then have 3: "bounded (\<Union>n. X n)"
using S unfolding bij_betw_def by blast
have "(\<Union>n. X n) \<in> lmeasurable"
by (rule measurable_countable_negligible_Union_bounded [OF 1 2 3])
with S have f1: "\<And>n. f ` (X n) \<in> lmeasurable"
unfolding bij_betw_def by (metis assms(1) cbox measurable_linear_image_interval rangeI)
have f2: "pairwise (\<lambda>m n. negligible (f ` (X m) \<inter> f ` (X n))) UNIV"
using S unfolding bij_betw_def pairwise_def by (metis injD negf rangeI)
have "bounded (\<Union>\<D>)"
by (meson Sup_least bounded_cbox bounded_subset cbox)
then have f3: "bounded (\<Union>n. f ` X n)"
using S unfolding bij_betw_def
by (metis bounded_linear_image linear_linear assms(1) image_Union range_composition)
have "(\<lambda>n. ?\<mu> (X n)) sums ?\<mu> (\<Union>n. X n)"
by (rule measure_countable_negligible_Union_bounded [OF 1 2 3])
have meq: "?\<mu> (\<Union>n. f ` X n) = m * ?\<mu> (\<Union>(X ` UNIV))"
proof (rule sums_unique2 [OF measure_countable_negligible_Union_bounded [OF f1 f2 f3]])
have m: "\<And>n. ?\<mu> (f ` X n) = (m * ?\<mu> (X n))"
using S unfolding bij_betw_def by (metis cbox im rangeI)
show "(\<lambda>n. ?\<mu> (f ` X n)) sums (m * ?\<mu> (\<Union>(X ` UNIV)))"
unfolding m
using measure_countable_negligible_Union_bounded [OF 1 2 3] sums_mult by blast
qed
show ?thesis
using measurable_countable_negligible_Union_bounded [OF f1 f2 f3] meq
by (auto simp: eq [symmetric])
qed
qed
show ?thesis
unfolding completion.fmeasurable_measure_inner_outer_le
proof (intro conjI allI impI)
fix e :: real
assume "e > 0"
have 1: "cbox a b - S \<in> lmeasurable"
by (simp add: fmeasurable.Diff that)
have 2: "0 < e / (1 + \<bar>m\<bar>)"
using \<open>e > 0\<close> by (simp add: field_split_simps abs_add_one_gt_zero)
obtain \<D>
where "countable \<D>"
and cbox: "\<And>K. K \<in> \<D> \<Longrightarrow> K \<subseteq> cbox a b \<and> K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
and intdisj: "pairwise (\<lambda>A B. interior A \<inter> interior B = {}) \<D>"
and DD: "cbox a b - S \<subseteq> \<Union>\<D>" "\<Union>\<D> \<in> lmeasurable"
and le: "?\<mu> (\<Union>\<D>) \<le> ?\<mu> (cbox a b - S) + e/(1 + \<bar>m\<bar>)"
by (rule measurable_outer_intervals_bounded [of "cbox a b - S" a b "e/(1 + \<bar>m\<bar>)"]; use 1 2 pairwise_def in force)
have meq: "?\<mu> (cbox a b - S) = ?\<mu> (cbox a b) - ?\<mu> S"
by (simp add: measurable_measure_Diff \<open>S \<subseteq> cbox a b\<close> fmeasurableD that(2))
show "\<exists>T \<in> lmeasurable. T \<subseteq> f ` S \<and> m * ?\<mu> S - e \<le> ?\<mu> T"
proof (intro bexI conjI)
show "f ` (cbox a b) - f ` (\<Union>\<D>) \<subseteq> f ` S"
using \<open>cbox a b - S \<subseteq> \<Union>\<D>\<close> by force
have "m * ?\<mu> S - e \<le> m * (?\<mu> S - e / (1 + \<bar>m\<bar>))"
using \<open>m \<ge> 0\<close> \<open>e > 0\<close> by (simp add: field_simps)
also have "\<dots> \<le> ?\<mu> (f ` cbox a b) - ?\<mu> (f ` (\<Union>\<D>))"
using le \<open>m \<ge> 0\<close> \<open>e > 0\<close>
apply (simp add: im fUD [OF \<open>countable \<D>\<close> cbox intdisj] right_diff_distrib [symmetric])
apply (rule mult_left_mono; simp add: algebra_simps meq)
done
also have "\<dots> = ?\<mu> (f ` cbox a b - f ` \<Union>\<D>)"
apply (rule measurable_measure_Diff [symmetric])
apply (simp add: assms(1) measurable_linear_image_interval)
apply (simp add: \<open>countable \<D>\<close> cbox fUD fmeasurableD intdisj)
apply (simp add: Sup_le_iff cbox image_mono)
done
finally show "m * ?\<mu> S - e \<le> ?\<mu> (f ` cbox a b - f ` \<Union>\<D>)" .
show "f ` cbox a b - f ` \<Union>\<D> \<in> lmeasurable"
by (simp add: fUD \<open>countable \<D>\<close> \<open>linear f\<close> cbox fmeasurable.Diff intdisj measurable_linear_image_interval)
qed
next
fix e :: real
assume "e > 0"
have em: "0 < e / (1 + \<bar>m\<bar>)"
using \<open>e > 0\<close> by (simp add: field_split_simps abs_add_one_gt_zero)
obtain \<D>
where "countable \<D>"
and cbox: "\<And>K. K \<in> \<D> \<Longrightarrow> K \<subseteq> cbox a b \<and> K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
and intdisj: "pairwise (\<lambda>A B. interior A \<inter> interior B = {}) \<D>"
and DD: "S \<subseteq> \<Union>\<D>" "\<Union>\<D> \<in> lmeasurable"
and le: "?\<mu> (\<Union>\<D>) \<le> ?\<mu> S + e/(1 + \<bar>m\<bar>)"
by (rule measurable_outer_intervals_bounded [of S a b "e/(1 + \<bar>m\<bar>)"]; use \<open>S \<in> lmeasurable\<close> \<open>S \<subseteq> cbox a b\<close> em in force)
show "\<exists>U \<in> lmeasurable. f ` S \<subseteq> U \<and> ?\<mu> U \<le> m * ?\<mu> S + e"
proof (intro bexI conjI)
show "f ` S \<subseteq> f ` (\<Union>\<D>)"
by (simp add: DD(1) image_mono)
have "?\<mu> (f ` \<Union>\<D>) \<le> m * (?\<mu> S + e / (1 + \<bar>m\<bar>))"
using \<open>m \<ge> 0\<close> le mult_left_mono
by (auto simp: fUD \<open>countable \<D>\<close> \<open>linear f\<close> cbox fmeasurable.Diff intdisj measurable_linear_image_interval)
also have "\<dots> \<le> m * ?\<mu> S + e"
using \<open>m \<ge> 0\<close> \<open>e > 0\<close> by (simp add: fUD [OF \<open>countable \<D>\<close> cbox intdisj] field_simps)
finally show "?\<mu> (f ` \<Union>\<D>) \<le> m * ?\<mu> S + e" .
show "f ` \<Union>\<D> \<in> lmeasurable"
by (simp add: \<open>countable \<D>\<close> cbox fUD intdisj)
qed
qed
qed
show ?thesis
unfolding has_measure_limit_iff
proof (intro allI impI)
fix e :: real
assume "e > 0"
obtain B where "B > 0" and B:
"\<And>a b. ball 0 B \<subseteq> cbox a b \<Longrightarrow> \<bar>?\<mu> (S \<inter> cbox a b) - ?\<mu> S\<bar> < e / (1 + \<bar>m\<bar>)"
using has_measure_limit [OF S] \<open>e > 0\<close> by (metis abs_add_one_gt_zero zero_less_divide_iff)
obtain c d::'n where cd: "ball 0 B \<subseteq> cbox c d"
by (metis bounded_subset_cbox_symmetric bounded_ball)
with B have less: "\<bar>?\<mu> (S \<inter> cbox c d) - ?\<mu> S\<bar> < e / (1 + \<bar>m\<bar>)" .
obtain D where "D > 0" and D: "cbox c d \<subseteq> ball 0 D"
by (metis bounded_cbox bounded_subset_ballD)
obtain C where "C > 0" and C: "\<And>x. norm (f x) \<le> C * norm x"
using linear_bounded_pos \<open>linear f\<close> by blast
have "f ` S \<inter> cbox a b \<in> lmeasurable \<and>
\<bar>?\<mu> (f ` S \<inter> cbox a b) - m * ?\<mu> S\<bar> < e"
if "ball 0 (D*C) \<subseteq> cbox a b" for a b
proof -
have "bounded (S \<inter> h ` cbox a b)"
by (simp add: bounded_linear_image linear_linear \<open>linear h\<close> bounded_Int)
moreover have Shab: "S \<inter> h ` cbox a b \<in> lmeasurable"
by (simp add: S \<open>linear h\<close> fmeasurable.Int measurable_linear_image_interval)
moreover have fim: "f ` (S \<inter> h ` (cbox a b)) = (f ` S) \<inter> cbox a b"
by (auto simp: hf rev_image_eqI fh)
ultimately have 1: "(f ` S) \<inter> cbox a b \<in> lmeasurable"
and 2: "m * ?\<mu> (S \<inter> h ` cbox a b) = ?\<mu> ((f ` S) \<inter> cbox a b)"
using fBS [of "S \<inter> (h ` (cbox a b))"] by auto
have *: "\<lbrakk>\<bar>z - m\<bar> < e; z \<le> w; w \<le> m\<rbrakk> \<Longrightarrow> \<bar>w - m\<bar> \<le> e"
for w z m and e::real by auto
have meas_adiff: "\<bar>?\<mu> (S \<inter> h ` cbox a b) - ?\<mu> S\<bar> \<le> e / (1 + \<bar>m\<bar>)"
proof (rule * [OF less])
show "?\<mu> (S \<inter> cbox c d) \<le> ?\<mu> (S \<inter> h ` cbox a b)"
proof (rule measure_mono_fmeasurable [OF _ _ Shab])
have "f ` ball 0 D \<subseteq> ball 0 (C * D)"
using C \<open>C > 0\<close>
apply (clarsimp simp: algebra_simps)
by (meson le_less_trans linordered_comm_semiring_strict_class.comm_mult_strict_left_mono)
then have "f ` ball 0 D \<subseteq> cbox a b"
by (metis mult.commute order_trans that)
have "ball 0 D \<subseteq> h ` cbox a b"
by (metis \<open>f ` ball 0 D \<subseteq> cbox a b\<close> hf image_subset_iff subsetI)
then show "S \<inter> cbox c d \<subseteq> S \<inter> h ` cbox a b"
using D by blast
next
show "S \<inter> cbox c d \<in> sets lebesgue"
using S fmeasurable_cbox by blast
qed
next
show "?\<mu> (S \<inter> h ` cbox a b) \<le> ?\<mu> S"
by (simp add: S Shab fmeasurableD measure_mono_fmeasurable)
qed
have "\<bar>?\<mu> (f ` S \<inter> cbox a b) - m * ?\<mu> S\<bar> \<le> m * e / (1 + \<bar>m\<bar>)"
proof -
have mm: "\<bar>m\<bar> = m"
by (simp add: \<open>0 \<le> m\<close>)
then have "\<bar>?\<mu> S - ?\<mu> (S \<inter> h ` cbox a b)\<bar> * m \<le> e / (1 + m) * m"
by (metis (no_types) \<open>0 \<le> m\<close> meas_adiff abs_minus_commute mult_right_mono)
moreover have "\<forall>r. \<bar>r * m\<bar> = m * \<bar>r\<bar>"
by (metis \<open>0 \<le> m\<close> abs_mult_pos mult.commute)
ultimately show ?thesis
apply (simp add: 2 [symmetric])
by (metis (no_types) abs_minus_commute mult.commute right_diff_distrib' mm)
qed
also have "\<dots> < e"
using \<open>e > 0\<close> by (cases "m \<ge> 0") (simp_all add: field_simps)
finally have "\<bar>?\<mu> (f ` S \<inter> cbox a b) - m * ?\<mu> S\<bar> < e" .
with 1 show ?thesis by auto
qed
then show "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
f ` S \<inter> cbox a b \<in> lmeasurable \<and>
\<bar>?\<mu> (f ` S \<inter> cbox a b) - m * ?\<mu> S\<bar> < e"
using \<open>C>0\<close> \<open>D>0\<close> by (metis mult_zero_left real_mult_less_iff1)
qed
qed
qed
subsection\<open>Lemmas about absolute integrability\<close>
lemma absolutely_integrable_linear:
fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
and h :: "'n::euclidean_space \<Rightarrow> 'p::euclidean_space"
shows "f absolutely_integrable_on s \<Longrightarrow> bounded_linear h \<Longrightarrow> (h \<circ> f) absolutely_integrable_on s"
using integrable_bounded_linear[of h lebesgue "\<lambda>x. indicator s x *\<^sub>R f x"]
by (simp add: linear_simps[of h] set_integrable_def)
lemma absolutely_integrable_sum:
fixes f :: "'a \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
assumes "finite T" and "\<And>a. a \<in> T \<Longrightarrow> (f a) absolutely_integrable_on S"
shows "(\<lambda>x. sum (\<lambda>a. f a x) T) absolutely_integrable_on S"
using assms by induction auto
lemma absolutely_integrable_integrable_bound:
fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
assumes le: "\<And>x. x\<in>S \<Longrightarrow> norm (f x) \<le> g x" and f: "f integrable_on S" and g: "g integrable_on S"
shows "f absolutely_integrable_on S"
unfolding set_integrable_def
proof (rule Bochner_Integration.integrable_bound)
have "g absolutely_integrable_on S"
unfolding absolutely_integrable_on_def
proof
show "(\<lambda>x. norm (g x)) integrable_on S"
using le norm_ge_zero[of "f _"]
by (intro integrable_spike_finite[OF _ _ g, of "{}"])
(auto intro!: abs_of_nonneg intro: order_trans simp del: norm_ge_zero)
qed fact
then show "integrable lebesgue (\<lambda>x. indicat_real S x *\<^sub>R g x)"
by (simp add: set_integrable_def)
show "(\<lambda>x. indicat_real S x *\<^sub>R f x) \<in> borel_measurable lebesgue"
using f by (auto intro: has_integral_implies_lebesgue_measurable simp: integrable_on_def)
qed (use le in \<open>force intro!: always_eventually split: split_indicator\<close>)
corollary absolutely_integrable_on_const [simp]:
fixes c :: "'a::euclidean_space"
assumes "S \<in> lmeasurable"
shows "(\<lambda>x. c) absolutely_integrable_on S"
by (metis (full_types) assms absolutely_integrable_integrable_bound integrable_on_const order_refl)
lemma absolutely_integrable_continuous:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
shows "continuous_on (cbox a b) f \<Longrightarrow> f absolutely_integrable_on cbox a b"
using absolutely_integrable_integrable_bound
by (simp add: absolutely_integrable_on_def continuous_on_norm integrable_continuous)
lemma absolutely_integrable_continuous_real:
fixes f :: "real \<Rightarrow> 'b::euclidean_space"
shows "continuous_on {a..b} f \<Longrightarrow> f absolutely_integrable_on {a..b}"
by (metis absolutely_integrable_continuous box_real(2))
lemma continuous_imp_integrable:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "continuous_on (cbox a b) f"
shows "integrable (lebesgue_on (cbox a b)) f"
proof -
have "f absolutely_integrable_on cbox a b"
by (simp add: absolutely_integrable_continuous assms)
then show ?thesis
by (simp add: integrable_restrict_space set_integrable_def)
qed
lemma continuous_imp_integrable_real:
fixes f :: "real \<Rightarrow> 'b::euclidean_space"
assumes "continuous_on {a..b} f"
shows "integrable (lebesgue_on {a..b}) f"
by (metis assms continuous_imp_integrable interval_cbox)
subsection \<open>Componentwise\<close>
proposition absolutely_integrable_componentwise_iff:
shows "f absolutely_integrable_on A \<longleftrightarrow> (\<forall>b\<in>Basis. (\<lambda>x. f x \<bullet> b) absolutely_integrable_on A)"
proof -
have *: "(\<lambda>x. norm (f x)) integrable_on A \<longleftrightarrow> (\<forall>b\<in>Basis. (\<lambda>x. norm (f x \<bullet> b)) integrable_on A)"
if "f integrable_on A"
proof -
have 1: "\<And>i. \<lbrakk>(\<lambda>x. norm (f x)) integrable_on A; i \<in> Basis\<rbrakk>
\<Longrightarrow> (\<lambda>x. f x \<bullet> i) absolutely_integrable_on A"
apply (rule absolutely_integrable_integrable_bound [where g = "\<lambda>x. norm(f x)"])
using Basis_le_norm integrable_component that apply fastforce+
done
have 2: "\<forall>i\<in>Basis. (\<lambda>x. \<bar>f x \<bullet> i\<bar>) integrable_on A \<Longrightarrow> f absolutely_integrable_on A"
apply (rule absolutely_integrable_integrable_bound [where g = "\<lambda>x. \<Sum>i\<in>Basis. norm (f x \<bullet> i)"])
using norm_le_l1 that apply (force intro: integrable_sum)+
done
show ?thesis
apply auto
apply (metis (full_types) absolutely_integrable_on_def set_integrable_abs 1)
apply (metis (full_types) absolutely_integrable_on_def 2)
done
qed
show ?thesis
unfolding absolutely_integrable_on_def
by (simp add: integrable_componentwise_iff [symmetric] ball_conj_distrib * cong: conj_cong)
qed
lemma absolutely_integrable_componentwise:
shows "(\<And>b. b \<in> Basis \<Longrightarrow> (\<lambda>x. f x \<bullet> b) absolutely_integrable_on A) \<Longrightarrow> f absolutely_integrable_on A"
using absolutely_integrable_componentwise_iff by blast
lemma absolutely_integrable_component:
"f absolutely_integrable_on A \<Longrightarrow> (\<lambda>x. f x \<bullet> (b :: 'b :: euclidean_space)) absolutely_integrable_on A"
by (drule absolutely_integrable_linear[OF _ bounded_linear_inner_left[of b]]) (simp add: o_def)
lemma absolutely_integrable_scaleR_left:
fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
assumes "f absolutely_integrable_on S"
shows "(\<lambda>x. c *\<^sub>R f x) absolutely_integrable_on S"
proof -
have "(\<lambda>x. c *\<^sub>R x) o f absolutely_integrable_on S"
apply (rule absolutely_integrable_linear [OF assms])
by (simp add: bounded_linear_scaleR_right)
then show ?thesis
using assms by blast
qed
lemma absolutely_integrable_scaleR_right:
assumes "f absolutely_integrable_on S"
shows "(\<lambda>x. f x *\<^sub>R c) absolutely_integrable_on S"
using assms by blast
lemma absolutely_integrable_norm:
fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
assumes "f absolutely_integrable_on S"
shows "(norm o f) absolutely_integrable_on S"
using assms by (simp add: absolutely_integrable_on_def o_def)
lemma absolutely_integrable_abs:
fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
assumes "f absolutely_integrable_on S"
shows "(\<lambda>x. \<Sum>i\<in>Basis. \<bar>f x \<bullet> i\<bar> *\<^sub>R i) absolutely_integrable_on S"
(is "?g absolutely_integrable_on S")
proof -
have eq: "?g =
(\<lambda>x. \<Sum>i\<in>Basis. ((\<lambda>y. \<Sum>j\<in>Basis. if j = i then y *\<^sub>R j else 0) \<circ>
(\<lambda>x. norm(\<Sum>j\<in>Basis. if j = i then (x \<bullet> i) *\<^sub>R j else 0)) \<circ> f) x)"
- by (simp add: sum.delta)
+ by (simp)
have *: "(\<lambda>y. \<Sum>j\<in>Basis. if j = i then y *\<^sub>R j else 0) \<circ>
(\<lambda>x. norm (\<Sum>j\<in>Basis. if j = i then (x \<bullet> i) *\<^sub>R j else 0)) \<circ> f
absolutely_integrable_on S"
if "i \<in> Basis" for i
proof -
have "bounded_linear (\<lambda>y. \<Sum>j\<in>Basis. if j = i then y *\<^sub>R j else 0)"
by (simp add: linear_linear algebra_simps linearI)
moreover have "(\<lambda>x. norm (\<Sum>j\<in>Basis. if j = i then (x \<bullet> i) *\<^sub>R j else 0)) \<circ> f
absolutely_integrable_on S"
unfolding o_def
apply (rule absolutely_integrable_norm [unfolded o_def])
using assms \<open>i \<in> Basis\<close>
apply (auto simp: algebra_simps dest: absolutely_integrable_component[where b=i])
done
ultimately show ?thesis
by (subst comp_assoc) (blast intro: absolutely_integrable_linear)
qed
show ?thesis
apply (rule ssubst [OF eq])
apply (rule absolutely_integrable_sum)
apply (force simp: intro!: *)+
done
qed
lemma abs_absolutely_integrableI_1:
fixes f :: "'a :: euclidean_space \<Rightarrow> real"
assumes f: "f integrable_on A" and "(\<lambda>x. \<bar>f x\<bar>) integrable_on A"
shows "f absolutely_integrable_on A"
by (rule absolutely_integrable_integrable_bound [OF _ assms]) auto
lemma abs_absolutely_integrableI:
assumes f: "f integrable_on S" and fcomp: "(\<lambda>x. \<Sum>i\<in>Basis. \<bar>f x \<bullet> i\<bar> *\<^sub>R i) integrable_on S"
shows "f absolutely_integrable_on S"
proof -
have "(\<lambda>x. (f x \<bullet> i) *\<^sub>R i) absolutely_integrable_on S" if "i \<in> Basis" for i
proof -
have "(\<lambda>x. \<bar>f x \<bullet> i\<bar>) integrable_on S"
using assms integrable_component [OF fcomp, where y=i] that by simp
then have "(\<lambda>x. f x \<bullet> i) absolutely_integrable_on S"
using abs_absolutely_integrableI_1 f integrable_component by blast
then show ?thesis
by (rule absolutely_integrable_scaleR_right)
qed
then have "(\<lambda>x. \<Sum>i\<in>Basis. (f x \<bullet> i) *\<^sub>R i) absolutely_integrable_on S"
by (simp add: absolutely_integrable_sum)
then show ?thesis
by (simp add: euclidean_representation)
qed
lemma absolutely_integrable_abs_iff:
"f absolutely_integrable_on S \<longleftrightarrow>
f integrable_on S \<and> (\<lambda>x. \<Sum>i\<in>Basis. \<bar>f x \<bullet> i\<bar> *\<^sub>R i) integrable_on S"
(is "?lhs = ?rhs")
proof
assume ?lhs then show ?rhs
using absolutely_integrable_abs absolutely_integrable_on_def by blast
next
assume ?rhs
moreover
have "(\<lambda>x. if x \<in> S then \<Sum>i\<in>Basis. \<bar>f x \<bullet> i\<bar> *\<^sub>R i else 0) = (\<lambda>x. \<Sum>i\<in>Basis. \<bar>(if x \<in> S then f x else 0) \<bullet> i\<bar> *\<^sub>R i)"
by force
ultimately show ?lhs
by (simp only: absolutely_integrable_restrict_UNIV [of S, symmetric] integrable_restrict_UNIV [of S, symmetric] abs_absolutely_integrableI)
qed
lemma absolutely_integrable_max:
fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S"
shows "(\<lambda>x. \<Sum>i\<in>Basis. max (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i)
absolutely_integrable_on S"
proof -
have "(\<lambda>x. \<Sum>i\<in>Basis. max (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i) =
(\<lambda>x. (1/2) *\<^sub>R (f x + g x + (\<Sum>i\<in>Basis. \<bar>f x \<bullet> i - g x \<bullet> i\<bar> *\<^sub>R i)))"
proof (rule ext)
fix x
have "(\<Sum>i\<in>Basis. max (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i) = (\<Sum>i\<in>Basis. ((f x \<bullet> i + g x \<bullet> i + \<bar>f x \<bullet> i - g x \<bullet> i\<bar>) / 2) *\<^sub>R i)"
by (force intro: sum.cong)
also have "... = (1 / 2) *\<^sub>R (\<Sum>i\<in>Basis. (f x \<bullet> i + g x \<bullet> i + \<bar>f x \<bullet> i - g x \<bullet> i\<bar>) *\<^sub>R i)"
by (simp add: scaleR_right.sum)
also have "... = (1 / 2) *\<^sub>R (f x + g x + (\<Sum>i\<in>Basis. \<bar>f x \<bullet> i - g x \<bullet> i\<bar> *\<^sub>R i))"
by (simp add: sum.distrib algebra_simps euclidean_representation)
finally
show "(\<Sum>i\<in>Basis. max (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i) =
(1 / 2) *\<^sub>R (f x + g x + (\<Sum>i\<in>Basis. \<bar>f x \<bullet> i - g x \<bullet> i\<bar> *\<^sub>R i))" .
qed
moreover have "(\<lambda>x. (1 / 2) *\<^sub>R (f x + g x + (\<Sum>i\<in>Basis. \<bar>f x \<bullet> i - g x \<bullet> i\<bar> *\<^sub>R i)))
absolutely_integrable_on S"
apply (intro set_integral_add absolutely_integrable_scaleR_left assms)
using absolutely_integrable_abs [OF set_integral_diff(1) [OF assms]]
apply (simp add: algebra_simps)
done
ultimately show ?thesis by metis
qed
corollary absolutely_integrable_max_1:
fixes f :: "'n::euclidean_space \<Rightarrow> real"
assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S"
shows "(\<lambda>x. max (f x) (g x)) absolutely_integrable_on S"
using absolutely_integrable_max [OF assms] by simp
lemma absolutely_integrable_min:
fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S"
shows "(\<lambda>x. \<Sum>i\<in>Basis. min (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i)
absolutely_integrable_on S"
proof -
have "(\<lambda>x. \<Sum>i\<in>Basis. min (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i) =
(\<lambda>x. (1/2) *\<^sub>R (f x + g x - (\<Sum>i\<in>Basis. \<bar>f x \<bullet> i - g x \<bullet> i\<bar> *\<^sub>R i)))"
proof (rule ext)
fix x
have "(\<Sum>i\<in>Basis. min (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i) = (\<Sum>i\<in>Basis. ((f x \<bullet> i + g x \<bullet> i - \<bar>f x \<bullet> i - g x \<bullet> i\<bar>) / 2) *\<^sub>R i)"
by (force intro: sum.cong)
also have "... = (1 / 2) *\<^sub>R (\<Sum>i\<in>Basis. (f x \<bullet> i + g x \<bullet> i - \<bar>f x \<bullet> i - g x \<bullet> i\<bar>) *\<^sub>R i)"
by (simp add: scaleR_right.sum)
also have "... = (1 / 2) *\<^sub>R (f x + g x - (\<Sum>i\<in>Basis. \<bar>f x \<bullet> i - g x \<bullet> i\<bar> *\<^sub>R i))"
by (simp add: sum.distrib sum_subtractf algebra_simps euclidean_representation)
finally
show "(\<Sum>i\<in>Basis. min (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i) =
(1 / 2) *\<^sub>R (f x + g x - (\<Sum>i\<in>Basis. \<bar>f x \<bullet> i - g x \<bullet> i\<bar> *\<^sub>R i))" .
qed
moreover have "(\<lambda>x. (1 / 2) *\<^sub>R (f x + g x - (\<Sum>i\<in>Basis. \<bar>f x \<bullet> i - g x \<bullet> i\<bar> *\<^sub>R i)))
absolutely_integrable_on S"
apply (intro set_integral_add set_integral_diff absolutely_integrable_scaleR_left assms)
using absolutely_integrable_abs [OF set_integral_diff(1) [OF assms]]
apply (simp add: algebra_simps)
done
ultimately show ?thesis by metis
qed
corollary absolutely_integrable_min_1:
fixes f :: "'n::euclidean_space \<Rightarrow> real"
assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S"
shows "(\<lambda>x. min (f x) (g x)) absolutely_integrable_on S"
using absolutely_integrable_min [OF assms] by simp
lemma nonnegative_absolutely_integrable:
fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
assumes "f integrable_on A" and comp: "\<And>x b. \<lbrakk>x \<in> A; b \<in> Basis\<rbrakk> \<Longrightarrow> 0 \<le> f x \<bullet> b"
shows "f absolutely_integrable_on A"
proof -
have "(\<lambda>x. (f x \<bullet> i) *\<^sub>R i) absolutely_integrable_on A" if "i \<in> Basis" for i
proof -
have "(\<lambda>x. f x \<bullet> i) integrable_on A"
by (simp add: assms(1) integrable_component)
then have "(\<lambda>x. f x \<bullet> i) absolutely_integrable_on A"
by (metis that comp nonnegative_absolutely_integrable_1)
then show ?thesis
by (rule absolutely_integrable_scaleR_right)
qed
then have "(\<lambda>x. \<Sum>i\<in>Basis. (f x \<bullet> i) *\<^sub>R i) absolutely_integrable_on A"
by (simp add: absolutely_integrable_sum)
then show ?thesis
by (simp add: euclidean_representation)
qed
lemma absolutely_integrable_component_ubound:
fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
assumes f: "f integrable_on A" and g: "g absolutely_integrable_on A"
and comp: "\<And>x b. \<lbrakk>x \<in> A; b \<in> Basis\<rbrakk> \<Longrightarrow> f x \<bullet> b \<le> g x \<bullet> b"
shows "f absolutely_integrable_on A"
proof -
have "(\<lambda>x. g x - (g x - f x)) absolutely_integrable_on A"
apply (rule set_integral_diff [OF g nonnegative_absolutely_integrable])
using Henstock_Kurzweil_Integration.integrable_diff absolutely_integrable_on_def f g apply blast
by (simp add: comp inner_diff_left)
then show ?thesis
by simp
qed
lemma absolutely_integrable_component_lbound:
fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
assumes f: "f absolutely_integrable_on A" and g: "g integrable_on A"
and comp: "\<And>x b. \<lbrakk>x \<in> A; b \<in> Basis\<rbrakk> \<Longrightarrow> f x \<bullet> b \<le> g x \<bullet> b"
shows "g absolutely_integrable_on A"
proof -
have "(\<lambda>x. f x + (g x - f x)) absolutely_integrable_on A"
apply (rule set_integral_add [OF f nonnegative_absolutely_integrable])
using Henstock_Kurzweil_Integration.integrable_diff absolutely_integrable_on_def f g apply blast
by (simp add: comp inner_diff_left)
then show ?thesis
by simp
qed
lemma integrable_on_1_iff:
fixes f :: "'a::euclidean_space \<Rightarrow> real^1"
shows "f integrable_on S \<longleftrightarrow> (\<lambda>x. f x $ 1) integrable_on S"
by (auto simp: integrable_componentwise_iff [of f] Basis_vec_def cart_eq_inner_axis)
lemma integral_on_1_eq:
fixes f :: "'a::euclidean_space \<Rightarrow> real^1"
shows "integral S f = vec (integral S (\<lambda>x. f x $ 1))"
by (cases "f integrable_on S") (simp_all add: integrable_on_1_iff vec_eq_iff not_integrable_integral)
lemma absolutely_integrable_on_1_iff:
fixes f :: "'a::euclidean_space \<Rightarrow> real^1"
shows "f absolutely_integrable_on S \<longleftrightarrow> (\<lambda>x. f x $ 1) absolutely_integrable_on S"
unfolding absolutely_integrable_on_def
by (auto simp: integrable_on_1_iff norm_real)
lemma absolutely_integrable_absolutely_integrable_lbound:
fixes f :: "'m::euclidean_space \<Rightarrow> real"
assumes f: "f integrable_on S" and g: "g absolutely_integrable_on S"
and *: "\<And>x. x \<in> S \<Longrightarrow> g x \<le> f x"
shows "f absolutely_integrable_on S"
by (rule absolutely_integrable_component_lbound [OF g f]) (simp add: *)
lemma absolutely_integrable_absolutely_integrable_ubound:
fixes f :: "'m::euclidean_space \<Rightarrow> real"
assumes fg: "f integrable_on S" "g absolutely_integrable_on S"
and *: "\<And>x. x \<in> S \<Longrightarrow> f x \<le> g x"
shows "f absolutely_integrable_on S"
by (rule absolutely_integrable_component_ubound [OF fg]) (simp add: *)
lemma has_integral_vec1_I_cbox:
fixes f :: "real^1 \<Rightarrow> 'a::real_normed_vector"
assumes "(f has_integral y) (cbox a b)"
shows "((f \<circ> vec) has_integral y) {a$1..b$1}"
proof -
have "((\<lambda>x. f(vec x)) has_integral (1 / 1) *\<^sub>R y) ((\<lambda>x. x $ 1) ` cbox a b)"
proof (rule has_integral_twiddle)
show "\<exists>w z::real^1. vec ` cbox u v = cbox w z"
"content (vec ` cbox u v :: (real^1) set) = 1 * content (cbox u v)" for u v
unfolding vec_cbox_1_eq
by (auto simp: content_cbox_if_cart interval_eq_empty_cart)
show "\<exists>w z. (\<lambda>x. x $ 1) ` cbox u v = cbox w z" for u v :: "real^1"
using vec_nth_cbox_1_eq by blast
qed (auto simp: continuous_vec assms)
then show ?thesis
by (simp add: o_def)
qed
lemma has_integral_vec1_I:
fixes f :: "real^1 \<Rightarrow> 'a::real_normed_vector"
assumes "(f has_integral y) S"
shows "(f \<circ> vec has_integral y) ((\<lambda>x. x $ 1) ` S)"
proof -
have *: "\<exists>z. ((\<lambda>x. if x \<in> (\<lambda>x. x $ 1) ` S then (f \<circ> vec) x else 0) has_integral z) {a..b} \<and> norm (z - y) < e"
if int: "\<And>a b. ball 0 B \<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)"
and B: "ball 0 B \<subseteq> {a..b}" for e B a b
proof -
have [simp]: "(\<exists>y\<in>S. x = y $ 1) \<longleftrightarrow> vec x \<in> S" for x
by force
have B': "ball (0::real^1) B \<subseteq> cbox (vec a) (vec b)"
using B by (simp add: Basis_vec_def cart_eq_inner_axis [symmetric] mem_box norm_real subset_iff)
show ?thesis
using int [OF B'] by (auto simp: image_iff o_def cong: if_cong dest!: has_integral_vec1_I_cbox)
qed
show ?thesis
using assms
apply (subst has_integral_alt)
apply (subst (asm) has_integral_alt)
apply (simp add: has_integral_vec1_I_cbox split: if_split_asm)
apply (metis vector_one_nth)
apply (erule all_forward imp_forward asm_rl ex_forward conj_forward)+
apply (blast intro!: *)
done
qed
lemma has_integral_vec1_nth_cbox:
fixes f :: "real \<Rightarrow> 'a::real_normed_vector"
assumes "(f has_integral y) {a..b}"
shows "((\<lambda>x::real^1. f(x$1)) has_integral y) (cbox (vec a) (vec b))"
proof -
have "((\<lambda>x::real^1. f(x$1)) has_integral (1 / 1) *\<^sub>R y) (vec ` cbox a b)"
proof (rule has_integral_twiddle)
show "\<exists>w z::real. (\<lambda>x. x $ 1) ` cbox u v = cbox w z"
"content ((\<lambda>x. x $ 1) ` cbox u v) = 1 * content (cbox u v)" for u v::"real^1"
unfolding vec_cbox_1_eq by (auto simp: content_cbox_if_cart interval_eq_empty_cart)
show "\<exists>w z::real^1. vec ` cbox u v = cbox w z" for u v :: "real"
using vec_cbox_1_eq by auto
qed (auto simp: continuous_vec assms)
then show ?thesis
using vec_cbox_1_eq by auto
qed
lemma has_integral_vec1_D_cbox:
fixes f :: "real^1 \<Rightarrow> 'a::real_normed_vector"
assumes "((f \<circ> vec) has_integral y) {a$1..b$1}"
shows "(f has_integral y) (cbox a b)"
by (metis (mono_tags, lifting) assms comp_apply has_integral_eq has_integral_vec1_nth_cbox vector_one_nth)
lemma has_integral_vec1_D:
fixes f :: "real^1 \<Rightarrow> 'a::real_normed_vector"
assumes "((f \<circ> vec) has_integral y) ((\<lambda>x. x $ 1) ` S)"
shows "(f has_integral y) S"
proof -
have *: "\<exists>z. ((\<lambda>x. if x \<in> S then f x else 0) has_integral z) (cbox a b) \<and> norm (z - y) < e"
if int: "\<And>a b. ball 0 B \<subseteq> {a..b} \<Longrightarrow>
(\<exists>z. ((\<lambda>x. if x \<in> (\<lambda>x. x $ 1) ` S then (f \<circ> vec) x else 0) has_integral z) {a..b} \<and> norm (z - y) < e)"
and B: "ball 0 B \<subseteq> cbox a b" for e B and a b::"real^1"
proof -
have B': "ball 0 B \<subseteq> {a$1..b$1}"
using B
apply (simp add: subset_iff Basis_vec_def cart_eq_inner_axis [symmetric] mem_box)
apply (metis (full_types) norm_real vec_component)
done
have eq: "(\<lambda>x. if vec x \<in> S then f (vec x) else 0) = (\<lambda>x. if x \<in> S then f x else 0) \<circ> vec"
by force
have [simp]: "(\<exists>y\<in>S. x = y $ 1) \<longleftrightarrow> vec x \<in> S" for x
by force
show ?thesis
using int [OF B'] by (auto simp: image_iff eq cong: if_cong dest!: has_integral_vec1_D_cbox)
qed
show ?thesis
using assms
apply (subst has_integral_alt)
apply (subst (asm) has_integral_alt)
apply (simp add: has_integral_vec1_D_cbox eq_cbox split: if_split_asm, blast)
apply (intro conjI impI)
apply (metis vector_one_nth)
apply (erule thin_rl)
apply (erule all_forward imp_forward asm_rl ex_forward conj_forward)+
using * apply blast
done
qed
lemma integral_vec1_eq:
fixes f :: "real^1 \<Rightarrow> 'a::real_normed_vector"
shows "integral S f = integral ((\<lambda>x. x $ 1) ` S) (f \<circ> vec)"
using has_integral_vec1_I [of f] has_integral_vec1_D [of f]
by (metis has_integral_iff not_integrable_integral)
lemma absolutely_integrable_drop:
fixes f :: "real^1 \<Rightarrow> 'b::euclidean_space"
shows "f absolutely_integrable_on S \<longleftrightarrow> (f \<circ> vec) absolutely_integrable_on (\<lambda>x. x $ 1) ` S"
unfolding absolutely_integrable_on_def integrable_on_def
proof safe
fix y r
assume "(f has_integral y) S" "((\<lambda>x. norm (f x)) has_integral r) S"
then show "\<exists>y. (f \<circ> vec has_integral y) ((\<lambda>x. x $ 1) ` S)"
"\<exists>y. ((\<lambda>x. norm ((f \<circ> vec) x)) has_integral y) ((\<lambda>x. x $ 1) ` S)"
by (force simp: o_def dest!: has_integral_vec1_I)+
next
fix y :: "'b" and r :: "real"
assume "(f \<circ> vec has_integral y) ((\<lambda>x. x $ 1) ` S)"
"((\<lambda>x. norm ((f \<circ> vec) x)) has_integral r) ((\<lambda>x. x $ 1) ` S)"
then show "\<exists>y. (f has_integral y) S" "\<exists>y. ((\<lambda>x. norm (f x)) has_integral y) S"
by (force simp: o_def intro: has_integral_vec1_D)+
qed
subsection \<open>Dominated convergence\<close>
lemma dominated_convergence:
fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
assumes f: "\<And>k. (f k) integrable_on S" and h: "h integrable_on S"
and le: "\<And>k x. x \<in> S \<Longrightarrow> norm (f k x) \<le> h x"
and conv: "\<And>x. x \<in> S \<Longrightarrow> (\<lambda>k. f k x) \<longlonglongrightarrow> g x"
shows "g integrable_on S" "(\<lambda>k. integral S (f k)) \<longlonglongrightarrow> integral S g"
proof -
have 3: "h absolutely_integrable_on S"
unfolding absolutely_integrable_on_def
proof
show "(\<lambda>x. norm (h x)) integrable_on S"
proof (intro integrable_spike_finite[OF _ _ h, of "{}"] ballI)
fix x assume "x \<in> S - {}" then show "norm (h x) = h x"
by (metis Diff_empty abs_of_nonneg bot_set_def le norm_ge_zero order_trans real_norm_def)
qed auto
qed fact
have 2: "set_borel_measurable lebesgue S (f k)" for k
unfolding set_borel_measurable_def
using f by (auto intro: has_integral_implies_lebesgue_measurable simp: integrable_on_def)
then have 1: "set_borel_measurable lebesgue S g"
unfolding set_borel_measurable_def
by (rule borel_measurable_LIMSEQ_metric) (use conv in \<open>auto split: split_indicator\<close>)
have 4: "AE x in lebesgue. (\<lambda>i. indicator S x *\<^sub>R f i x) \<longlonglongrightarrow> indicator S x *\<^sub>R g x"
"AE x in lebesgue. norm (indicator S x *\<^sub>R f k x) \<le> indicator S x *\<^sub>R h x" for k
using conv le by (auto intro!: always_eventually split: split_indicator)
have g: "g absolutely_integrable_on S"
using 1 2 3 4 unfolding set_borel_measurable_def set_integrable_def
by (rule integrable_dominated_convergence)
then show "g integrable_on S"
by (auto simp: absolutely_integrable_on_def)
have "(\<lambda>k. (LINT x:S|lebesgue. f k x)) \<longlonglongrightarrow> (LINT x:S|lebesgue. g x)"
unfolding set_borel_measurable_def set_lebesgue_integral_def
using 1 2 3 4 unfolding set_borel_measurable_def set_lebesgue_integral_def set_integrable_def
by (rule integral_dominated_convergence)
then show "(\<lambda>k. integral S (f k)) \<longlonglongrightarrow> integral S g"
using g absolutely_integrable_integrable_bound[OF le f h]
by (subst (asm) (1 2) set_lebesgue_integral_eq_integral) auto
qed
lemma has_integral_dominated_convergence:
fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
assumes "\<And>k. (f k has_integral y k) S" "h integrable_on S"
"\<And>k. \<forall>x\<in>S. norm (f k x) \<le> h x" "\<forall>x\<in>S. (\<lambda>k. f k x) \<longlonglongrightarrow> g x"
and x: "y \<longlonglongrightarrow> x"
shows "(g has_integral x) S"
proof -
have int_f: "\<And>k. (f k) integrable_on S"
using assms by (auto simp: integrable_on_def)
have "(g has_integral (integral S g)) S"
by (metis assms(2-4) dominated_convergence(1) has_integral_integral int_f)
moreover have "integral S g = x"
proof (rule LIMSEQ_unique)
show "(\<lambda>i. integral S (f i)) \<longlonglongrightarrow> x"
using integral_unique[OF assms(1)] x by simp
show "(\<lambda>i. integral S (f i)) \<longlonglongrightarrow> integral S g"
by (metis assms(2) assms(3) assms(4) dominated_convergence(2) int_f)
qed
ultimately show ?thesis
by simp
qed
lemma dominated_convergence_integrable_1:
fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> real"
assumes f: "\<And>k. f k absolutely_integrable_on S"
and h: "h integrable_on S"
and normg: "\<And>x. x \<in> S \<Longrightarrow> norm(g x) \<le> (h x)"
and lim: "\<And>x. x \<in> S \<Longrightarrow> (\<lambda>k. f k x) \<longlonglongrightarrow> g x"
shows "g integrable_on S"
proof -
have habs: "h absolutely_integrable_on S"
using h normg nonnegative_absolutely_integrable_1 norm_ge_zero order_trans by blast
let ?f = "\<lambda>n x. (min (max (- h x) (f n x)) (h x))"
have h0: "h x \<ge> 0" if "x \<in> S" for x
using normg that by force
have leh: "norm (?f k x) \<le> h x" if "x \<in> S" for k x
using h0 that by force
have limf: "(\<lambda>k. ?f k x) \<longlonglongrightarrow> g x" if "x \<in> S" for x
proof -
have "\<And>e y. \<bar>f y x - g x\<bar> < e \<Longrightarrow> \<bar>min (max (- h x) (f y x)) (h x) - g x\<bar> < e"
using h0 [OF that] normg [OF that] by simp
then show ?thesis
using lim [OF that] by (auto simp add: tendsto_iff dist_norm elim!: eventually_mono)
qed
show ?thesis
proof (rule dominated_convergence [of ?f S h g])
have "(\<lambda>x. - h x) absolutely_integrable_on S"
using habs unfolding set_integrable_def by auto
then show "?f k integrable_on S" for k
by (intro set_lebesgue_integral_eq_integral absolutely_integrable_min_1 absolutely_integrable_max_1 f habs)
qed (use assms leh limf in auto)
qed
lemma dominated_convergence_integrable:
fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
assumes f: "\<And>k. f k absolutely_integrable_on S"
and h: "h integrable_on S"
and normg: "\<And>x. x \<in> S \<Longrightarrow> norm(g x) \<le> (h x)"
and lim: "\<And>x. x \<in> S \<Longrightarrow> (\<lambda>k. f k x) \<longlonglongrightarrow> g x"
shows "g integrable_on S"
using f
unfolding integrable_componentwise_iff [of g] absolutely_integrable_componentwise_iff [where f = "f k" for k]
proof clarify
fix b :: "'m"
assume fb [rule_format]: "\<And>k. \<forall>b\<in>Basis. (\<lambda>x. f k x \<bullet> b) absolutely_integrable_on S" and b: "b \<in> Basis"
show "(\<lambda>x. g x \<bullet> b) integrable_on S"
proof (rule dominated_convergence_integrable_1 [OF fb h])
fix x
assume "x \<in> S"
show "norm (g x \<bullet> b) \<le> h x"
using norm_nth_le \<open>x \<in> S\<close> b normg order.trans by blast
show "(\<lambda>k. f k x \<bullet> b) \<longlonglongrightarrow> g x \<bullet> b"
using \<open>x \<in> S\<close> b lim tendsto_componentwise_iff by fastforce
qed (use b in auto)
qed
lemma dominated_convergence_absolutely_integrable:
fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
assumes f: "\<And>k. f k absolutely_integrable_on S"
and h: "h integrable_on S"
and normg: "\<And>x. x \<in> S \<Longrightarrow> norm(g x) \<le> (h x)"
and lim: "\<And>x. x \<in> S \<Longrightarrow> (\<lambda>k. f k x) \<longlonglongrightarrow> g x"
shows "g absolutely_integrable_on S"
proof -
have "g integrable_on S"
by (rule dominated_convergence_integrable [OF assms])
with assms show ?thesis
by (blast intro: absolutely_integrable_integrable_bound [where g=h])
qed
proposition integral_countable_UN:
fixes f :: "real^'m \<Rightarrow> real^'n"
assumes f: "f absolutely_integrable_on (\<Union>(range s))"
and s: "\<And>m. s m \<in> sets lebesgue"
shows "\<And>n. f absolutely_integrable_on (\<Union>m\<le>n. s m)"
and "(\<lambda>n. integral (\<Union>m\<le>n. s m) f) \<longlonglongrightarrow> integral (\<Union>(s ` UNIV)) f" (is "?F \<longlonglongrightarrow> ?I")
proof -
show fU: "f absolutely_integrable_on (\<Union>m\<le>n. s m)" for n
using assms by (blast intro: set_integrable_subset [OF f])
have fint: "f integrable_on (\<Union> (range s))"
using absolutely_integrable_on_def f by blast
let ?h = "\<lambda>x. if x \<in> \<Union>(s ` UNIV) then norm(f x) else 0"
have "(\<lambda>n. integral UNIV (\<lambda>x. if x \<in> (\<Union>m\<le>n. s m) then f x else 0))
\<longlonglongrightarrow> integral UNIV (\<lambda>x. if x \<in> \<Union>(s ` UNIV) then f x else 0)"
proof (rule dominated_convergence)
show "(\<lambda>x. if x \<in> (\<Union>m\<le>n. s m) then f x else 0) integrable_on UNIV" for n
unfolding integrable_restrict_UNIV
using fU absolutely_integrable_on_def by blast
show "(\<lambda>x. if x \<in> \<Union>(s ` UNIV) then norm(f x) else 0) integrable_on UNIV"
by (metis (no_types) absolutely_integrable_on_def f integrable_restrict_UNIV)
show "\<And>x. (\<lambda>n. if x \<in> (\<Union>m\<le>n. s m) then f x else 0)
\<longlonglongrightarrow> (if x \<in> \<Union>(s ` UNIV) then f x else 0)"
by (force intro: tendsto_eventually eventually_sequentiallyI)
qed auto
then show "?F \<longlonglongrightarrow> ?I"
by (simp only: integral_restrict_UNIV)
qed
subsection \<open>Fundamental Theorem of Calculus for the Lebesgue integral\<close>
text \<open>
For the positive integral we replace continuity with Borel-measurability.
\<close>
lemma
fixes f :: "real \<Rightarrow> real"
assumes [measurable]: "f \<in> borel_measurable borel"
assumes f: "\<And>x. x \<in> {a..b} \<Longrightarrow> DERIV F x :> f x" "\<And>x. x \<in> {a..b} \<Longrightarrow> 0 \<le> f x" and "a \<le> b"
shows nn_integral_FTC_Icc: "(\<integral>\<^sup>+x. ennreal (f x) * indicator {a .. b} x \<partial>lborel) = F b - F a" (is ?nn)
and has_bochner_integral_FTC_Icc_nonneg:
"has_bochner_integral lborel (\<lambda>x. f x * indicator {a .. b} x) (F b - F a)" (is ?has)
and integral_FTC_Icc_nonneg: "(\<integral>x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a" (is ?eq)
and integrable_FTC_Icc_nonneg: "integrable lborel (\<lambda>x. f x * indicator {a .. b} x)" (is ?int)
proof -
have *: "(\<lambda>x. f x * indicator {a..b} x) \<in> borel_measurable borel" "\<And>x. 0 \<le> f x * indicator {a..b} x"
using f(2) by (auto split: split_indicator)
have F_mono: "a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b\<Longrightarrow> F x \<le> F y" for x y
using f by (intro DERIV_nonneg_imp_nondecreasing[of x y F]) (auto intro: order_trans)
have "(f has_integral F b - F a) {a..b}"
by (intro fundamental_theorem_of_calculus)
(auto simp: has_field_derivative_iff_has_vector_derivative[symmetric]
intro: has_field_derivative_subset[OF f(1)] \<open>a \<le> b\<close>)
then have i: "((\<lambda>x. f x * indicator {a .. b} x) has_integral F b - F a) UNIV"
unfolding indicator_def if_distrib[where f="\<lambda>x. a * x" for a]
by (simp cong del: if_weak_cong del: atLeastAtMost_iff)
then have nn: "(\<integral>\<^sup>+x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a"
by (rule nn_integral_has_integral_lborel[OF *])
then show ?has
by (rule has_bochner_integral_nn_integral[rotated 3]) (simp_all add: * F_mono \<open>a \<le> b\<close>)
then show ?eq ?int
unfolding has_bochner_integral_iff by auto
show ?nn
by (subst nn[symmetric])
(auto intro!: nn_integral_cong simp add: ennreal_mult f split: split_indicator)
qed
lemma
fixes f :: "real \<Rightarrow> 'a :: euclidean_space"
assumes "a \<le> b"
assumes "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (F has_vector_derivative f x) (at x within {a .. b})"
assumes cont: "continuous_on {a .. b} f"
shows has_bochner_integral_FTC_Icc:
"has_bochner_integral lborel (\<lambda>x. indicator {a .. b} x *\<^sub>R f x) (F b - F a)" (is ?has)
and integral_FTC_Icc: "(\<integral>x. indicator {a .. b} x *\<^sub>R f x \<partial>lborel) = F b - F a" (is ?eq)
proof -
let ?f = "\<lambda>x. indicator {a .. b} x *\<^sub>R f x"
have int: "integrable lborel ?f"
using borel_integrable_compact[OF _ cont] by auto
have "(f has_integral F b - F a) {a..b}"
using assms(1,2) by (intro fundamental_theorem_of_calculus) auto
moreover
have "(f has_integral integral\<^sup>L lborel ?f) {a..b}"
using has_integral_integral_lborel[OF int]
unfolding indicator_def if_distrib[where f="\<lambda>x. x *\<^sub>R a" for a]
by (simp cong del: if_weak_cong del: atLeastAtMost_iff)
ultimately show ?eq
by (auto dest: has_integral_unique)
then show ?has
using int by (auto simp: has_bochner_integral_iff)
qed
lemma
fixes f :: "real \<Rightarrow> real"
assumes "a \<le> b"
assumes deriv: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> DERIV F x :> f x"
assumes cont: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> isCont f x"
shows has_bochner_integral_FTC_Icc_real:
"has_bochner_integral lborel (\<lambda>x. f x * indicator {a .. b} x) (F b - F a)" (is ?has)
and integral_FTC_Icc_real: "(\<integral>x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a" (is ?eq)
proof -
have 1: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (F has_vector_derivative f x) (at x within {a .. b})"
unfolding has_field_derivative_iff_has_vector_derivative[symmetric]
using deriv by (auto intro: DERIV_subset)
have 2: "continuous_on {a .. b} f"
using cont by (intro continuous_at_imp_continuous_on) auto
show ?has ?eq
using has_bochner_integral_FTC_Icc[OF \<open>a \<le> b\<close> 1 2] integral_FTC_Icc[OF \<open>a \<le> b\<close> 1 2]
by (auto simp: mult.commute)
qed
lemma nn_integral_FTC_atLeast:
fixes f :: "real \<Rightarrow> real"
assumes f_borel: "f \<in> borel_measurable borel"
assumes f: "\<And>x. a \<le> x \<Longrightarrow> DERIV F x :> f x"
assumes nonneg: "\<And>x. a \<le> x \<Longrightarrow> 0 \<le> f x"
assumes lim: "(F \<longlongrightarrow> T) at_top"
shows "(\<integral>\<^sup>+x. ennreal (f x) * indicator {a ..} x \<partial>lborel) = T - F a"
proof -
let ?f = "\<lambda>(i::nat) (x::real). ennreal (f x) * indicator {a..a + real i} x"
let ?fR = "\<lambda>x. ennreal (f x) * indicator {a ..} x"
have F_mono: "a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> F x \<le> F y" for x y
using f nonneg by (intro DERIV_nonneg_imp_nondecreasing[of x y F]) (auto intro: order_trans)
then have F_le_T: "a \<le> x \<Longrightarrow> F x \<le> T" for x
by (intro tendsto_lowerbound[OF lim])
(auto simp: eventually_at_top_linorder)
have "(SUP i. ?f i x) = ?fR x" for x
proof (rule LIMSEQ_unique[OF LIMSEQ_SUP])
obtain n where "x - a < real n"
using reals_Archimedean2[of "x - a"] ..
then have "eventually (\<lambda>n. ?f n x = ?fR x) sequentially"
by (auto intro!: eventually_sequentiallyI[where c=n] split: split_indicator)
then show "(\<lambda>n. ?f n x) \<longlonglongrightarrow> ?fR x"
by (rule tendsto_eventually)
qed (auto simp: nonneg incseq_def le_fun_def split: split_indicator)
then have "integral\<^sup>N lborel ?fR = (\<integral>\<^sup>+ x. (SUP i. ?f i x) \<partial>lborel)"
by simp
also have "\<dots> = (SUP i. (\<integral>\<^sup>+ x. ?f i x \<partial>lborel))"
proof (rule nn_integral_monotone_convergence_SUP)
show "incseq ?f"
using nonneg by (auto simp: incseq_def le_fun_def split: split_indicator)
show "\<And>i. (?f i) \<in> borel_measurable lborel"
using f_borel by auto
qed
also have "\<dots> = (SUP i. ennreal (F (a + real i) - F a))"
by (subst nn_integral_FTC_Icc[OF f_borel f nonneg]) auto
also have "\<dots> = T - F a"
proof (rule LIMSEQ_unique[OF LIMSEQ_SUP])
have "(\<lambda>x. F (a + real x)) \<longlonglongrightarrow> T"
apply (rule filterlim_compose[OF lim filterlim_tendsto_add_at_top])
apply (rule LIMSEQ_const_iff[THEN iffD2, OF refl])
apply (rule filterlim_real_sequentially)
done
then show "(\<lambda>n. ennreal (F (a + real n) - F a)) \<longlonglongrightarrow> ennreal (T - F a)"
by (simp add: F_mono F_le_T tendsto_diff)
qed (auto simp: incseq_def intro!: ennreal_le_iff[THEN iffD2] F_mono)
finally show ?thesis .
qed
lemma integral_power:
"a \<le> b \<Longrightarrow> (\<integral>x. x^k * indicator {a..b} x \<partial>lborel) = (b^Suc k - a^Suc k) / Suc k"
proof (subst integral_FTC_Icc_real)
fix x show "DERIV (\<lambda>x. x^Suc k / Suc k) x :> x^k"
by (intro derivative_eq_intros) auto
qed (auto simp: field_simps simp del: of_nat_Suc)
subsection \<open>Integration by parts\<close>
lemma integral_by_parts_integrable:
fixes f g F G::"real \<Rightarrow> real"
assumes "a \<le> b"
assumes cont_f[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont f x"
assumes cont_g[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont g x"
assumes [intro]: "!!x. DERIV F x :> f x"
assumes [intro]: "!!x. DERIV G x :> g x"
shows "integrable lborel (\<lambda>x.((F x) * (g x) + (f x) * (G x)) * indicator {a .. b} x)"
by (auto intro!: borel_integrable_atLeastAtMost continuous_intros) (auto intro!: DERIV_isCont)
lemma integral_by_parts:
fixes f g F G::"real \<Rightarrow> real"
assumes [arith]: "a \<le> b"
assumes cont_f[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont f x"
assumes cont_g[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont g x"
assumes [intro]: "!!x. DERIV F x :> f x"
assumes [intro]: "!!x. DERIV G x :> g x"
shows "(\<integral>x. (F x * g x) * indicator {a .. b} x \<partial>lborel)
= F b * G b - F a * G a - \<integral>x. (f x * G x) * indicator {a .. b} x \<partial>lborel"
proof-
have 0: "(\<integral>x. (F x * g x + f x * G x) * indicator {a .. b} x \<partial>lborel) = F b * G b - F a * G a"
by (rule integral_FTC_Icc_real, auto intro!: derivative_eq_intros continuous_intros)
(auto intro!: DERIV_isCont)
have "(\<integral>x. (F x * g x + f x * G x) * indicator {a .. b} x \<partial>lborel) =
(\<integral>x. (F x * g x) * indicator {a .. b} x \<partial>lborel) + \<integral>x. (f x * G x) * indicator {a .. b} x \<partial>lborel"
apply (subst Bochner_Integration.integral_add[symmetric])
apply (auto intro!: borel_integrable_atLeastAtMost continuous_intros)
by (auto intro!: DERIV_isCont Bochner_Integration.integral_cong split: split_indicator)
thus ?thesis using 0 by auto
qed
lemma integral_by_parts':
fixes f g F G::"real \<Rightarrow> real"
assumes "a \<le> b"
assumes "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont f x"
assumes "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont g x"
assumes "!!x. DERIV F x :> f x"
assumes "!!x. DERIV G x :> g x"
shows "(\<integral>x. indicator {a .. b} x *\<^sub>R (F x * g x) \<partial>lborel)
= F b * G b - F a * G a - \<integral>x. indicator {a .. b} x *\<^sub>R (f x * G x) \<partial>lborel"
using integral_by_parts[OF assms] by (simp add: ac_simps)
lemma has_bochner_integral_even_function:
fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
assumes f: "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x) x"
assumes even: "\<And>x. f (- x) = f x"
shows "has_bochner_integral lborel f (2 *\<^sub>R x)"
proof -
have indicator: "\<And>x::real. indicator {..0} (- x) = indicator {0..} x"
by (auto split: split_indicator)
have "has_bochner_integral lborel (\<lambda>x. indicator {.. 0} x *\<^sub>R f x) x"
by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0])
(auto simp: indicator even f)
with f have "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x + indicator {.. 0} x *\<^sub>R f x) (x + x)"
by (rule has_bochner_integral_add)
then have "has_bochner_integral lborel f (x + x)"
by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4])
(auto split: split_indicator)
then show ?thesis
by (simp add: scaleR_2)
qed
lemma has_bochner_integral_odd_function:
fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
assumes f: "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x) x"
assumes odd: "\<And>x. f (- x) = - f x"
shows "has_bochner_integral lborel f 0"
proof -
have indicator: "\<And>x::real. indicator {..0} (- x) = indicator {0..} x"
by (auto split: split_indicator)
have "has_bochner_integral lborel (\<lambda>x. - indicator {.. 0} x *\<^sub>R f x) x"
by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0])
(auto simp: indicator odd f)
from has_bochner_integral_minus[OF this]
have "has_bochner_integral lborel (\<lambda>x. indicator {.. 0} x *\<^sub>R f x) (- x)"
by simp
with f have "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x + indicator {.. 0} x *\<^sub>R f x) (x + - x)"
by (rule has_bochner_integral_add)
then have "has_bochner_integral lborel f (x + - x)"
by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4])
(auto split: split_indicator)
then show ?thesis
by simp
qed
lemma has_integral_0_closure_imp_0:
fixes f :: "'a::euclidean_space \<Rightarrow> real"
assumes f: "continuous_on (closure S) f"
and nonneg_interior: "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f x"
and pos: "0 < emeasure lborel S"
and finite: "emeasure lborel S < \<infinity>"
and regular: "emeasure lborel (closure S) = emeasure lborel S"
and opn: "open S"
assumes int: "(f has_integral 0) (closure S)"
assumes x: "x \<in> closure S"
shows "f x = 0"
proof -
have zero: "emeasure lborel (frontier S) = 0"
using finite closure_subset regular
unfolding frontier_def
by (subst emeasure_Diff) (auto simp: frontier_def interior_open \<open>open S\<close> )
have nonneg: "0 \<le> f x" if "x \<in> closure S" for x
using continuous_ge_on_closure[OF f that nonneg_interior] by simp
have "0 = integral (closure S) f"
by (blast intro: int sym)
also
note intl = has_integral_integrable[OF int]
have af: "f absolutely_integrable_on (closure S)"
using nonneg
by (intro absolutely_integrable_onI intl integrable_eq[OF intl]) simp
then have "integral (closure S) f = set_lebesgue_integral lebesgue (closure S) f"
by (intro set_lebesgue_integral_eq_integral(2)[symmetric])
also have "\<dots> = 0 \<longleftrightarrow> (AE x in lebesgue. indicator (closure S) x *\<^sub>R f x = 0)"
unfolding set_lebesgue_integral_def
proof (rule integral_nonneg_eq_0_iff_AE)
show "integrable lebesgue (\<lambda>x. indicat_real (closure S) x *\<^sub>R f x)"
by (metis af set_integrable_def)
qed (use nonneg in \<open>auto simp: indicator_def\<close>)
also have "\<dots> \<longleftrightarrow> (AE x in lebesgue. x \<in> {x. x \<in> closure S \<longrightarrow> f x = 0})"
by (auto simp: indicator_def)
finally have "(AE x in lebesgue. x \<in> {x. x \<in> closure S \<longrightarrow> f x = 0})" by simp
moreover have "(AE x in lebesgue. x \<in> - frontier S)"
using zero
by (auto simp: eventually_ae_filter null_sets_def intro!: exI[where x="frontier S"])
ultimately have ae: "AE x \<in> S in lebesgue. x \<in> {x \<in> closure S. f x = 0}" (is ?th)
by eventually_elim (use closure_subset in \<open>auto simp: \<close>)
have "closed {0::real}" by simp
with continuous_on_closed_vimage[OF closed_closure, of S f] f
have "closed (f -` {0} \<inter> closure S)" by blast
then have "closed {x \<in> closure S. f x = 0}" by (auto simp: vimage_def Int_def conj_commute)
with \<open>open S\<close> have "x \<in> {x \<in> closure S. f x = 0}" if "x \<in> S" for x using ae that
by (rule mem_closed_if_AE_lebesgue_open)
then have "f x = 0" if "x \<in> S" for x using that by auto
from continuous_constant_on_closure[OF f this \<open>x \<in> closure S\<close>]
show "f x = 0" .
qed
lemma has_integral_0_cbox_imp_0:
fixes f :: "'a::euclidean_space \<Rightarrow> real"
assumes f: "continuous_on (cbox a b) f"
and nonneg_interior: "\<And>x. x \<in> box a b \<Longrightarrow> 0 \<le> f x"
assumes int: "(f has_integral 0) (cbox a b)"
assumes ne: "box a b \<noteq> {}"
assumes x: "x \<in> cbox a b"
shows "f x = 0"
proof -
have "0 < emeasure lborel (box a b)"
using ne x unfolding emeasure_lborel_box_eq
by (force intro!: prod_pos simp: mem_box algebra_simps)
then show ?thesis using assms
by (intro has_integral_0_closure_imp_0[of "box a b" f x])
(auto simp: emeasure_lborel_box_eq emeasure_lborel_cbox_eq algebra_simps mem_box)
qed
subsection\<open>Various common equivalent forms of function measurability\<close>
lemma indicator_sum_eq:
fixes m::real and f :: "'a \<Rightarrow> real"
assumes "\<bar>m\<bar> \<le> 2 ^ (2*n)" "m/2^n \<le> f x" "f x < (m+1)/2^n" "m \<in> \<int>"
shows "(\<Sum>k::real | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n).
k/2^n * indicator {y. k/2^n \<le> f y \<and> f y < (k+1)/2^n} x) = m/2^n"
(is "sum ?f ?S = _")
proof -
have "sum ?f ?S = sum (\<lambda>k. k/2^n * indicator {y. k/2^n \<le> f y \<and> f y < (k+1)/2^n} x) {m}"
proof (rule comm_monoid_add_class.sum.mono_neutral_right)
show "finite ?S"
by (rule finite_abs_int_segment)
show "{m} \<subseteq> {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}"
using assms by auto
show "\<forall>i\<in>{k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)} - {m}. ?f i = 0"
using assms by (auto simp: indicator_def Ints_def abs_le_iff field_split_simps)
qed
also have "\<dots> = m/2^n"
using assms by (auto simp: indicator_def not_less)
finally show ?thesis .
qed
lemma measurable_on_sf_limit_lemma1:
fixes f :: "'a::euclidean_space \<Rightarrow> real"
assumes "\<And>a b. {x \<in> S. a \<le> f x \<and> f x < b} \<in> sets (lebesgue_on S)"
obtains g where "\<And>n. g n \<in> borel_measurable (lebesgue_on S)"
"\<And>n. finite(range (g n))"
"\<And>x. (\<lambda>n. g n x) \<longlonglongrightarrow> f x"
proof
show "(\<lambda>x. sum (\<lambda>k::real. k/2^n * indicator {y. k/2^n \<le> f y \<and> f y < (k+1)/2^n} x)
{k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}) \<in> borel_measurable (lebesgue_on S)"
(is "?g \<in> _") for n
proof -
have "\<And>k. \<lbrakk>k \<in> \<int>; \<bar>k\<bar> \<le> 2 ^ (2*n)\<rbrakk>
\<Longrightarrow> Measurable.pred (lebesgue_on S) (\<lambda>x. k / (2^n) \<le> f x \<and> f x < (k+1) / (2^n))"
using assms by (force simp: pred_def space_restrict_space)
then show ?thesis
by (simp add: field_class.field_divide_inverse)
qed
show "finite (range (?g n))" for n
proof -
have "range (?g n) \<subseteq> (\<lambda>k. k/2^n) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}"
proof clarify
fix x
show "?g n x \<in> (\<lambda>k. k/2^n) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}"
proof (cases "\<exists>k::real. k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n) \<and> k/2^n \<le> (f x) \<and> (f x) < (k+1)/2^n")
case True
then show ?thesis
apply clarify
by (subst indicator_sum_eq) auto
next
case False
then have "?g n x = 0" by auto
then show ?thesis by force
qed
qed
moreover have "finite ((\<lambda>k::real. (k/2^n)) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)})"
by (simp add: finite_abs_int_segment)
ultimately show ?thesis
using finite_subset by blast
qed
show "(\<lambda>n. ?g n x) \<longlonglongrightarrow> f x" for x
proof (rule LIMSEQ_I)
fix e::real
assume "e > 0"
obtain N1 where N1: "\<bar>f x\<bar> < 2 ^ N1"
using real_arch_pow by fastforce
obtain N2 where N2: "(1/2) ^ N2 < e"
using real_arch_pow_inv \<open>e > 0\<close> by force
have "norm (?g n x - f x) < e" if n: "n \<ge> max N1 N2" for n
proof -
define m where "m \<equiv> floor(2^n * (f x))"
have "1 \<le> \<bar>2^n\<bar> * e"
using n N2 \<open>e > 0\<close> less_eq_real_def less_le_trans by (fastforce simp add: field_split_simps)
then have *: "\<lbrakk>x \<le> y; y < x + 1\<rbrakk> \<Longrightarrow> abs(x - y) < \<bar>2^n\<bar> * e" for x y::real
by linarith
have "\<bar>2^n\<bar> * \<bar>m/2^n - f x\<bar> = \<bar>2^n * (m/2^n - f x)\<bar>"
by (simp add: abs_mult)
also have "\<dots> = \<bar>real_of_int \<lfloor>2^n * f x\<rfloor> - f x * 2^n\<bar>"
by (simp add: algebra_simps m_def)
also have "\<dots> < \<bar>2^n\<bar> * e"
by (rule *; simp add: mult.commute)
finally have "\<bar>2^n\<bar> * \<bar>m/2^n - f x\<bar> < \<bar>2^n\<bar> * e" .
then have me: "\<bar>m/2^n - f x\<bar> < e"
by simp
have "\<bar>real_of_int m\<bar> \<le> 2 ^ (2*n)"
proof (cases "f x < 0")
case True
then have "-\<lfloor>f x\<rfloor> \<le> \<lfloor>(2::real) ^ N1\<rfloor>"
using N1 le_floor_iff minus_le_iff by fastforce
with n True have "\<bar>real_of_int\<lfloor>f x\<rfloor>\<bar> \<le> 2 ^ N1"
by linarith
also have "\<dots> \<le> 2^n"
using n by (simp add: m_def)
finally have "\<bar>real_of_int \<lfloor>f x\<rfloor>\<bar> * 2^n \<le> 2^n * 2^n"
by simp
moreover
have "\<bar>real_of_int \<lfloor>2^n * f x\<rfloor>\<bar> \<le> \<bar>real_of_int \<lfloor>f x\<rfloor>\<bar> * 2^n"
proof -
have "\<bar>real_of_int \<lfloor>2^n * f x\<rfloor>\<bar> = - (real_of_int \<lfloor>2^n * f x\<rfloor>)"
using True by (simp add: abs_if mult_less_0_iff)
also have "\<dots> \<le> - (real_of_int (\<lfloor>(2::real) ^ n\<rfloor> * \<lfloor>f x\<rfloor>))"
using le_mult_floor_Ints [of "(2::real)^n"] by simp
also have "\<dots> \<le> \<bar>real_of_int \<lfloor>f x\<rfloor>\<bar> * 2^n"
using True
by simp
finally show ?thesis .
qed
ultimately show ?thesis
by (metis (no_types, hide_lams) m_def order_trans power2_eq_square power_even_eq)
next
case False
with n N1 have "f x \<le> 2^n"
by (simp add: not_less) (meson less_eq_real_def one_le_numeral order_trans power_increasing)
moreover have "0 \<le> m"
using False m_def by force
ultimately show ?thesis
by (metis abs_of_nonneg floor_mono le_floor_iff m_def of_int_0_le_iff power2_eq_square power_mult real_mult_le_cancel_iff1 zero_less_numeral mult.commute zero_less_power)
qed
then have "?g n x = m/2^n"
by (rule indicator_sum_eq) (auto simp add: m_def field_split_simps, linarith)
then have "norm (?g n x - f x) = norm (m/2^n - f x)"
by simp
also have "\<dots> < e"
by (simp add: me)
finally show ?thesis .
qed
then show "\<exists>no. \<forall>n\<ge>no. norm (?g n x - f x) < e"
by blast
qed
qed
lemma borel_measurable_simple_function_limit:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
shows "f \<in> borel_measurable (lebesgue_on S) \<longleftrightarrow>
(\<exists>g. (\<forall>n. (g n) \<in> borel_measurable (lebesgue_on S)) \<and>
(\<forall>n. finite (range (g n))) \<and> (\<forall>x. (\<lambda>n. g n x) \<longlonglongrightarrow> f x))"
proof -
have "\<exists>g. (\<forall>n. (g n) \<in> borel_measurable (lebesgue_on S)) \<and>
(\<forall>n. finite (range (g n))) \<and> (\<forall>x. (\<lambda>n. g n x) \<longlonglongrightarrow> f x)"
if f: "\<And>a i. i \<in> Basis \<Longrightarrow> {x \<in> S. f x \<bullet> i < a} \<in> sets (lebesgue_on S)"
proof -
have "\<exists>g. (\<forall>n. (g n) \<in> borel_measurable (lebesgue_on S)) \<and>
(\<forall>n. finite(image (g n) UNIV)) \<and>
(\<forall>x. ((\<lambda>n. g n x) \<longlonglongrightarrow> f x \<bullet> i))" if "i \<in> Basis" for i
proof (rule measurable_on_sf_limit_lemma1 [of S "\<lambda>x. f x \<bullet> i"])
show "{x \<in> S. a \<le> f x \<bullet> i \<and> f x \<bullet> i < b} \<in> sets (lebesgue_on S)" for a b
proof -
have "{x \<in> S. a \<le> f x \<bullet> i \<and> f x \<bullet> i < b} = {x \<in> S. f x \<bullet> i < b} - {x \<in> S. a > f x \<bullet> i}"
by auto
also have "\<dots> \<in> sets (lebesgue_on S)"
using f that by blast
finally show ?thesis .
qed
qed blast
then obtain g where g:
"\<And>i n. i \<in> Basis \<Longrightarrow> g i n \<in> borel_measurable (lebesgue_on S)"
"\<And>i n. i \<in> Basis \<Longrightarrow> finite(range (g i n))"
"\<And>i x. i \<in> Basis \<Longrightarrow> ((\<lambda>n. g i n x) \<longlonglongrightarrow> f x \<bullet> i)"
by metis
show ?thesis
proof (intro conjI allI exI)
show "(\<lambda>x. \<Sum>i\<in>Basis. g i n x *\<^sub>R i) \<in> borel_measurable (lebesgue_on S)" for n
by (intro borel_measurable_sum borel_measurable_scaleR) (auto intro: g)
show "finite (range (\<lambda>x. \<Sum>i\<in>Basis. g i n x *\<^sub>R i))" for n
proof -
have "range (\<lambda>x. \<Sum>i\<in>Basis. g i n x *\<^sub>R i) \<subseteq> (\<lambda>h. \<Sum>i\<in>Basis. h i *\<^sub>R i) ` PiE Basis (\<lambda>i. range (g i n))"
proof clarify
fix x
show "(\<Sum>i\<in>Basis. g i n x *\<^sub>R i) \<in> (\<lambda>h. \<Sum>i\<in>Basis. h i *\<^sub>R i) ` (\<Pi>\<^sub>E i\<in>Basis. range (g i n))"
by (rule_tac x="\<lambda>i\<in>Basis. g i n x" in image_eqI) auto
qed
moreover have "finite(PiE Basis (\<lambda>i. range (g i n)))"
by (simp add: g finite_PiE)
ultimately show ?thesis
by (metis (mono_tags, lifting) finite_surj)
qed
show "(\<lambda>n. \<Sum>i\<in>Basis. g i n x *\<^sub>R i) \<longlonglongrightarrow> f x" for x
proof -
have "(\<lambda>n. \<Sum>i\<in>Basis. g i n x *\<^sub>R i) \<longlonglongrightarrow> (\<Sum>i\<in>Basis. (f x \<bullet> i) *\<^sub>R i)"
by (auto intro!: tendsto_sum tendsto_scaleR g)
moreover have "(\<Sum>i\<in>Basis. (f x \<bullet> i) *\<^sub>R i) = f x"
using euclidean_representation by blast
ultimately show ?thesis
by metis
qed
qed
qed
moreover have "f \<in> borel_measurable (lebesgue_on S)"
if meas_g: "\<And>n. g n \<in> borel_measurable (lebesgue_on S)"
and fin: "\<And>n. finite (range (g n))"
and to_f: "\<And>x. (\<lambda>n. g n x) \<longlonglongrightarrow> f x" for g
by (rule borel_measurable_LIMSEQ_metric [OF meas_g to_f])
ultimately show ?thesis
using borel_measurable_vimage_halfspace_component_lt by blast
qed
subsection \<open>Lebesgue sets and continuous images\<close>
proposition lebesgue_regular_inner:
assumes "S \<in> sets lebesgue"
obtains K C where "negligible K" "\<And>n::nat. compact(C n)" "S = (\<Union>n. C n) \<union> K"
proof -
have "\<exists>T. closed T \<and> T \<subseteq> S \<and> (S - T) \<in> lmeasurable \<and> emeasure lebesgue (S - T) < ennreal ((1/2)^n)" for n
using sets_lebesgue_inner_closed assms
by (metis sets_lebesgue_inner_closed zero_less_divide_1_iff zero_less_numeral zero_less_power)
then obtain C where clo: "\<And>n. closed (C n)" and subS: "\<And>n. C n \<subseteq> S"
and mea: "\<And>n. (S - C n) \<in> lmeasurable"
and less: "\<And>n. emeasure lebesgue (S - C n) < ennreal ((1/2)^n)"
by metis
have "\<exists>F. (\<forall>n::nat. compact(F n)) \<and> (\<Union>n. F n) = C m" for m::nat
by (metis clo closed_Union_compact_subsets)
then obtain D :: "[nat,nat] \<Rightarrow> 'a set" where D: "\<And>m n. compact(D m n)" "\<And>m. (\<Union>n. D m n) = C m"
by metis
let ?C = "from_nat_into (\<Union>m. range (D m))"
have "countable (\<Union>m. range (D m))"
by blast
have "range (from_nat_into (\<Union>m. range (D m))) = (\<Union>m. range (D m))"
using range_from_nat_into by simp
then have CD: "\<exists>m n. ?C k = D m n" for k
by (metis (mono_tags, lifting) UN_iff rangeE range_eqI)
show thesis
proof
show "negligible (S - (\<Union>n. C n))"
proof (clarsimp simp: negligible_outer_le)
fix e :: "real"
assume "e > 0"
then obtain n where n: "(1/2)^n < e"
using real_arch_pow_inv [of e "1/2"] by auto
show "\<exists>T. S - (\<Union>n. C n) \<subseteq> T \<and> T \<in> lmeasurable \<and> measure lebesgue T \<le> e"
proof (intro exI conjI)
show "S - (\<Union>n. C n) \<subseteq> S - C n"
by blast
show "S - C n \<in> lmeasurable"
by (simp add: mea)
show "measure lebesgue (S - C n) \<le> e"
using less [of n] n
by (simp add: emeasure_eq_measure2 less_le mea)
qed
qed
show "compact (?C n)" for n
using CD D by metis
show "S = (\<Union>n. ?C n) \<union> (S - (\<Union>n. C n))" (is "_ = ?rhs")
proof
show "S \<subseteq> ?rhs"
using D by fastforce
show "?rhs \<subseteq> S"
using subS D CD by auto (metis Sup_upper range_eqI subsetCE)
qed
qed
qed
lemma sets_lebesgue_continuous_image:
assumes T: "T \<in> sets lebesgue" and contf: "continuous_on S f"
and negim: "\<And>T. \<lbrakk>negligible T; T \<subseteq> S\<rbrakk> \<Longrightarrow> negligible(f ` T)" and "T \<subseteq> S"
shows "f ` T \<in> sets lebesgue"
proof -
obtain K C where "negligible K" and com: "\<And>n::nat. compact(C n)" and Teq: "T = (\<Union>n. C n) \<union> K"
using lebesgue_regular_inner [OF T] by metis
then have comf: "\<And>n::nat. compact(f ` C n)"
by (metis Un_subset_iff Union_upper \<open>T \<subseteq> S\<close> compact_continuous_image contf continuous_on_subset rangeI)
have "((\<Union>n. f ` C n) \<union> f ` K) \<in> sets lebesgue"
proof (rule sets.Un)
have "K \<subseteq> S"
using Teq \<open>T \<subseteq> S\<close> by blast
show "(\<Union>n. f ` C n) \<in> sets lebesgue"
proof (rule sets.countable_Union)
show "range (\<lambda>n. f ` C n) \<subseteq> sets lebesgue"
using borel_compact comf by (auto simp: borel_compact)
qed auto
show "f ` K \<in> sets lebesgue"
by (simp add: \<open>K \<subseteq> S\<close> \<open>negligible K\<close> negim negligible_imp_sets)
qed
then show ?thesis
by (simp add: Teq image_Un image_Union)
qed
lemma differentiable_image_in_sets_lebesgue:
fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
assumes S: "S \<in> sets lebesgue" and dim: "DIM('m) \<le> DIM('n)" and f: "f differentiable_on S"
shows "f`S \<in> sets lebesgue"
proof (rule sets_lebesgue_continuous_image [OF S])
show "continuous_on S f"
by (meson differentiable_imp_continuous_on f)
show "\<And>T. \<lbrakk>negligible T; T \<subseteq> S\<rbrakk> \<Longrightarrow> negligible (f ` T)"
using differentiable_on_subset f
by (auto simp: intro!: negligible_differentiable_image_negligible [OF dim])
qed auto
lemma sets_lebesgue_on_continuous_image:
assumes S: "S \<in> sets lebesgue" and X: "X \<in> sets (lebesgue_on S)" and contf: "continuous_on S f"
and negim: "\<And>T. \<lbrakk>negligible T; T \<subseteq> S\<rbrakk> \<Longrightarrow> negligible(f ` T)"
shows "f ` X \<in> sets (lebesgue_on (f ` S))"
proof -
have "X \<subseteq> S"
by (metis S X sets.Int_space_eq2 sets_restrict_space_iff)
moreover have "f ` S \<in> sets lebesgue"
using S contf negim sets_lebesgue_continuous_image by blast
moreover have "f ` X \<in> sets lebesgue"
by (metis S X contf negim sets_lebesgue_continuous_image sets_restrict_space_iff space_restrict_space space_restrict_space2)
ultimately show ?thesis
by (auto simp: sets_restrict_space_iff)
qed
lemma differentiable_image_in_sets_lebesgue_on:
fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
assumes S: "S \<in> sets lebesgue" and X: "X \<in> sets (lebesgue_on S)" and dim: "DIM('m) \<le> DIM('n)"
and f: "f differentiable_on S"
shows "f ` X \<in> sets (lebesgue_on (f`S))"
proof (rule sets_lebesgue_on_continuous_image [OF S X])
show "continuous_on S f"
by (meson differentiable_imp_continuous_on f)
show "\<And>T. \<lbrakk>negligible T; T \<subseteq> S\<rbrakk> \<Longrightarrow> negligible (f ` T)"
using differentiable_on_subset f
by (auto simp: intro!: negligible_differentiable_image_negligible [OF dim])
qed
subsection \<open>Affine lemmas\<close>
lemma borel_measurable_affine:
fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
assumes f: "f \<in> borel_measurable lebesgue" and "c \<noteq> 0"
shows "(\<lambda>x. f(t + c *\<^sub>R x)) \<in> borel_measurable lebesgue"
proof -
{ fix a b
have "{x. f x \<in> cbox a b} \<in> sets lebesgue"
using f cbox_borel lebesgue_measurable_vimage_borel by blast
then have "(\<lambda>x. (x - t) /\<^sub>R c) ` {x. f x \<in> cbox a b} \<in> sets lebesgue"
proof (rule differentiable_image_in_sets_lebesgue)
show "(\<lambda>x. (x - t) /\<^sub>R c) differentiable_on {x. f x \<in> cbox a b}"
unfolding differentiable_on_def differentiable_def
by (rule \<open>c \<noteq> 0\<close> derivative_eq_intros strip exI | simp)+
qed auto
moreover
have "{x. f(t + c *\<^sub>R x) \<in> cbox a b} = (\<lambda>x. (x-t) /\<^sub>R c) ` {x. f x \<in> cbox a b}"
using \<open>c \<noteq> 0\<close> by (auto simp: image_def)
ultimately have "{x. f(t + c *\<^sub>R x) \<in> cbox a b} \<in> sets lebesgue"
by (auto simp: borel_measurable_vimage_closed_interval) }
then show ?thesis
by (subst lebesgue_on_UNIV_eq [symmetric]; auto simp: borel_measurable_vimage_closed_interval)
qed
lemma lebesgue_integrable_real_affine:
fixes f :: "real \<Rightarrow> 'a :: euclidean_space"
assumes f: "integrable lebesgue f" and "c \<noteq> 0"
shows "integrable lebesgue (\<lambda>x. f(t + c * x))"
proof -
have "(\<lambda>x. norm (f x)) \<in> borel_measurable lebesgue"
by (simp add: borel_measurable_integrable f)
then show ?thesis
using assms borel_measurable_affine [of f c]
unfolding integrable_iff_bounded
by (subst (asm) nn_integral_real_affine_lebesgue[where c=c and t=t]) (auto simp: ennreal_mult_less_top)
qed
lemma lebesgue_integrable_real_affine_iff:
fixes f :: "real \<Rightarrow> 'a :: euclidean_space"
shows "c \<noteq> 0 \<Longrightarrow> integrable lebesgue (\<lambda>x. f(t + c * x)) \<longleftrightarrow> integrable lebesgue f"
using lebesgue_integrable_real_affine[of f c t]
lebesgue_integrable_real_affine[of "\<lambda>x. f(t + c * x)" "1/c" "-t/c"]
by (auto simp: field_simps)
lemma\<^marker>\<open>tag important\<close> lebesgue_integral_real_affine:
fixes f :: "real \<Rightarrow> 'a :: euclidean_space" and c :: real
assumes c: "c \<noteq> 0" shows "(\<integral>x. f x \<partial> lebesgue) = \<bar>c\<bar> *\<^sub>R (\<integral>x. f(t + c * x) \<partial>lebesgue)"
proof cases
have "(\<lambda>x. t + c * x) \<in> lebesgue \<rightarrow>\<^sub>M lebesgue"
using lebesgue_affine_measurable[where c= "\<lambda>x::real. c"] \<open>c \<noteq> 0\<close> by simp
moreover
assume "integrable lebesgue f"
ultimately show ?thesis
by (subst lebesgue_real_affine[OF c, of t]) (auto simp: integral_density integral_distr)
next
assume "\<not> integrable lebesgue f" with c show ?thesis
by (simp add: lebesgue_integrable_real_affine_iff not_integrable_integral_eq)
qed
lemma has_bochner_integral_lebesgue_real_affine_iff:
fixes i :: "'a :: euclidean_space"
shows "c \<noteq> 0 \<Longrightarrow>
has_bochner_integral lebesgue f i \<longleftrightarrow>
has_bochner_integral lebesgue (\<lambda>x. f(t + c * x)) (i /\<^sub>R \<bar>c\<bar>)"
unfolding has_bochner_integral_iff lebesgue_integrable_real_affine_iff
by (simp_all add: lebesgue_integral_real_affine[symmetric] divideR_right cong: conj_cong)
lemma has_bochner_integral_reflect_real_lemma[intro]:
fixes f :: "real \<Rightarrow> 'a::euclidean_space"
assumes "has_bochner_integral (lebesgue_on {a..b}) f i"
shows "has_bochner_integral (lebesgue_on {-b..-a}) (\<lambda>x. f(-x)) i"
proof -
have eq: "indicat_real {a..b} (- x) *\<^sub>R f(- x) = indicat_real {- b..- a} x *\<^sub>R f(- x)" for x
by (auto simp: indicator_def)
have i: "has_bochner_integral lebesgue (\<lambda>x. indicator {a..b} x *\<^sub>R f x) i"
using assms by (auto simp: has_bochner_integral_restrict_space)
then have "has_bochner_integral lebesgue (\<lambda>x. indicator {-b..-a} x *\<^sub>R f(-x)) i"
using has_bochner_integral_lebesgue_real_affine_iff [of "-1" "(\<lambda>x. indicator {a..b} x *\<^sub>R f x)" i 0]
by (auto simp: eq)
then show ?thesis
by (auto simp: has_bochner_integral_restrict_space)
qed
lemma has_bochner_integral_reflect_real[simp]:
fixes f :: "real \<Rightarrow> 'a::euclidean_space"
shows "has_bochner_integral (lebesgue_on {-b..-a}) (\<lambda>x. f(-x)) i \<longleftrightarrow> has_bochner_integral (lebesgue_on {a..b}) f i"
by (auto simp: dest: has_bochner_integral_reflect_real_lemma)
lemma integrable_reflect_real[simp]:
fixes f :: "real \<Rightarrow> 'a::euclidean_space"
shows "integrable (lebesgue_on {-b..-a}) (\<lambda>x. f(-x)) \<longleftrightarrow> integrable (lebesgue_on {a..b}) f"
by (metis has_bochner_integral_iff has_bochner_integral_reflect_real)
lemma integral_reflect_real[simp]:
fixes f :: "real \<Rightarrow> 'a::euclidean_space"
shows "integral\<^sup>L (lebesgue_on {-b .. -a}) (\<lambda>x. f(-x)) = integral\<^sup>L (lebesgue_on {a..b::real}) f"
using has_bochner_integral_reflect_real [of b a f]
by (metis has_bochner_integral_iff not_integrable_integral_eq)
subsection\<open>More results on integrability\<close>
lemma integrable_on_all_intervals_UNIV:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
assumes intf: "\<And>a b. f integrable_on cbox a b"
and normf: "\<And>x. norm(f x) \<le> g x" and g: "g integrable_on UNIV"
shows "f integrable_on UNIV"
proof -
have intg: "(\<forall>a b. g integrable_on cbox a b)"
and gle_e: "\<forall>e>0. \<exists>B>0. \<forall>a b c d.
ball 0 B \<subseteq> cbox a b \<and> cbox a b \<subseteq> cbox c d \<longrightarrow>
\<bar>integral (cbox a b) g - integral (cbox c d) g\<bar>
< e"
using g
by (auto simp: integrable_alt_subset [of _ UNIV] intf)
have le: "norm (integral (cbox a b) f - integral (cbox c d) f) \<le> \<bar>integral (cbox a b) g - integral (cbox c d) g\<bar>"
if "cbox a b \<subseteq> cbox c d" for a b c d
proof -
have "norm (integral (cbox a b) f - integral (cbox c d) f) = norm (integral (cbox c d - cbox a b) f)"
using intf that by (simp add: norm_minus_commute integral_setdiff)
also have "\<dots> \<le> integral (cbox c d - cbox a b) g"
proof (rule integral_norm_bound_integral [OF _ _ normf])
show "f integrable_on cbox c d - cbox a b" "g integrable_on cbox c d - cbox a b"
by (meson integrable_integral integrable_setdiff intf intg negligible_setdiff that)+
qed
also have "\<dots> = integral (cbox c d) g - integral (cbox a b) g"
using intg that by (simp add: integral_setdiff)
also have "\<dots> \<le> \<bar>integral (cbox a b) g - integral (cbox c d) g\<bar>"
by simp
finally show ?thesis .
qed
show ?thesis
using gle_e
apply (simp add: integrable_alt_subset [of _ UNIV] intf)
apply (erule imp_forward all_forward ex_forward asm_rl)+
by (meson not_less order_trans le)
qed
lemma integrable_on_all_intervals_integrable_bound:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
assumes intf: "\<And>a b. (\<lambda>x. if x \<in> S then f x else 0) integrable_on cbox a b"
and normf: "\<And>x. x \<in> S \<Longrightarrow> norm(f x) \<le> g x" and g: "g integrable_on S"
shows "f integrable_on S"
using integrable_on_all_intervals_UNIV [OF intf, of "(\<lambda>x. if x \<in> S then g x else 0)"]
by (simp add: g integrable_restrict_UNIV normf)
lemma measurable_bounded_lemma:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes f: "f \<in> borel_measurable lebesgue" and g: "g integrable_on cbox a b"
and normf: "\<And>x. x \<in> cbox a b \<Longrightarrow> norm(f x) \<le> g x"
shows "f integrable_on cbox a b"
proof -
have "g absolutely_integrable_on cbox a b"
by (metis (full_types) add_increasing g le_add_same_cancel1 nonnegative_absolutely_integrable_1 norm_ge_zero normf)
then have "integrable (lebesgue_on (cbox a b)) g"
by (simp add: integrable_restrict_space set_integrable_def)
then have "integrable (lebesgue_on (cbox a b)) f"
proof (rule Bochner_Integration.integrable_bound)
show "AE x in lebesgue_on (cbox a b). norm (f x) \<le> norm (g x)"
by (rule AE_I2) (auto intro: normf order_trans)
qed (simp add: f measurable_restrict_space1)
then show ?thesis
by (simp add: integrable_on_lebesgue_on)
qed
proposition measurable_bounded_by_integrable_imp_integrable:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes f: "f \<in> borel_measurable (lebesgue_on S)" and g: "g integrable_on S"
and normf: "\<And>x. x \<in> S \<Longrightarrow> norm(f x) \<le> g x" and S: "S \<in> sets lebesgue"
shows "f integrable_on S"
proof (rule integrable_on_all_intervals_integrable_bound [OF _ normf g])
show "(\<lambda>x. if x \<in> S then f x else 0) integrable_on cbox a b" for a b
proof (rule measurable_bounded_lemma)
show "(\<lambda>x. if x \<in> S then f x else 0) \<in> borel_measurable lebesgue"
by (simp add: S borel_measurable_if f)
show "(\<lambda>x. if x \<in> S then g x else 0) integrable_on cbox a b"
by (simp add: g integrable_altD(1))
show "norm (if x \<in> S then f x else 0) \<le> (if x \<in> S then g x else 0)" for x
using normf by simp
qed
qed
lemma measurable_bounded_by_integrable_imp_lebesgue_integrable:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes f: "f \<in> borel_measurable (lebesgue_on S)" and g: "integrable (lebesgue_on S) g"
and normf: "\<And>x. x \<in> S \<Longrightarrow> norm(f x) \<le> g x" and S: "S \<in> sets lebesgue"
shows "integrable (lebesgue_on S) f"
proof -
have "f absolutely_integrable_on S"
by (metis (no_types) S absolutely_integrable_integrable_bound f g integrable_on_lebesgue_on measurable_bounded_by_integrable_imp_integrable normf)
then show ?thesis
by (simp add: S integrable_restrict_space set_integrable_def)
qed
lemma measurable_bounded_by_integrable_imp_integrable_real:
fixes f :: "'a::euclidean_space \<Rightarrow> real"
assumes "f \<in> borel_measurable (lebesgue_on S)" "g integrable_on S" "\<And>x. x \<in> S \<Longrightarrow> abs(f x) \<le> g x" "S \<in> sets lebesgue"
shows "f integrable_on S"
using measurable_bounded_by_integrable_imp_integrable [of f S g] assms by simp
subsection\<open> Relation between Borel measurability and integrability.\<close>
lemma integrable_imp_measurable_weak:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "S \<in> sets lebesgue" "f integrable_on S"
shows "f \<in> borel_measurable (lebesgue_on S)"
by (metis (mono_tags, lifting) assms has_integral_implies_lebesgue_measurable borel_measurable_restrict_space_iff integrable_on_def sets.Int_space_eq2)
lemma integrable_imp_measurable:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "f integrable_on S"
shows "f \<in> borel_measurable (lebesgue_on S)"
proof -
have "(UNIV::'a set) \<in> sets lborel"
by simp
then show ?thesis
by (metis (mono_tags, lifting) assms borel_measurable_if_D integrable_imp_measurable_weak integrable_restrict_UNIV lebesgue_on_UNIV_eq sets_lebesgue_on_refl)
qed
lemma integrable_iff_integrable_on:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "S \<in> sets lebesgue" "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>lebesgue_on S) < \<infinity>"
shows "integrable (lebesgue_on S) f \<longleftrightarrow> f integrable_on S"
using assms integrable_iff_bounded integrable_imp_measurable integrable_on_lebesgue_on by blast
lemma absolutely_integrable_measurable:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "S \<in> sets lebesgue"
shows "f absolutely_integrable_on S \<longleftrightarrow> f \<in> borel_measurable (lebesgue_on S) \<and> integrable (lebesgue_on S) (norm \<circ> f)"
(is "?lhs = ?rhs")
proof
assume L: ?lhs
then have "f \<in> borel_measurable (lebesgue_on S)"
by (simp add: absolutely_integrable_on_def integrable_imp_measurable)
then show ?rhs
using assms set_integrable_norm [of lebesgue S f] L
by (simp add: integrable_restrict_space set_integrable_def)
next
assume ?rhs then show ?lhs
using assms integrable_on_lebesgue_on
by (metis absolutely_integrable_integrable_bound comp_def eq_iff measurable_bounded_by_integrable_imp_integrable)
qed
lemma absolutely_integrable_measurable_real:
fixes f :: "'a::euclidean_space \<Rightarrow> real"
assumes "S \<in> sets lebesgue"
shows "f absolutely_integrable_on S \<longleftrightarrow>
f \<in> borel_measurable (lebesgue_on S) \<and> integrable (lebesgue_on S) (\<lambda>x. \<bar>f x\<bar>)"
by (simp add: absolutely_integrable_measurable assms o_def)
lemma absolutely_integrable_measurable_real':
fixes f :: "'a::euclidean_space \<Rightarrow> real"
assumes "S \<in> sets lebesgue"
shows "f absolutely_integrable_on S \<longleftrightarrow> f \<in> borel_measurable (lebesgue_on S) \<and> (\<lambda>x. \<bar>f x\<bar>) integrable_on S"
using assms
apply (auto simp: absolutely_integrable_measurable integrable_on_lebesgue_on)
apply (simp add: integrable_on_lebesgue_on measurable_bounded_by_integrable_imp_lebesgue_integrable)
using abs_absolutely_integrableI_1 absolutely_integrable_measurable measurable_bounded_by_integrable_imp_integrable_real by blast
lemma absolutely_integrable_imp_borel_measurable:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "f absolutely_integrable_on S" "S \<in> sets lebesgue"
shows "f \<in> borel_measurable (lebesgue_on S)"
using absolutely_integrable_measurable assms by blast
lemma measurable_bounded_by_integrable_imp_absolutely_integrable:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "f \<in> borel_measurable (lebesgue_on S)" "S \<in> sets lebesgue"
and "g integrable_on S" and "\<And>x. x \<in> S \<Longrightarrow> norm(f x) \<le> (g x)"
shows "f absolutely_integrable_on S"
using assms absolutely_integrable_integrable_bound measurable_bounded_by_integrable_imp_integrable by blast
proposition negligible_differentiable_vimage:
fixes f :: "'a \<Rightarrow> 'a::euclidean_space"
assumes "negligible T"
and f': "\<And>x. x \<in> S \<Longrightarrow> inj(f' x)"
and derf: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
shows "negligible {x \<in> S. f x \<in> T}"
proof -
define U where
"U \<equiv> \<lambda>n::nat. {x \<in> S. \<forall>y. y \<in> S \<and> norm(y - x) < 1/n
\<longrightarrow> norm(y - x) \<le> n * norm(f y - f x)}"
have "negligible {x \<in> U n. f x \<in> T}" if "n > 0" for n
proof (subst locally_negligible_alt, clarify)
fix a
assume a: "a \<in> U n" and fa: "f a \<in> T"
define V where "V \<equiv> {x. x \<in> U n \<and> f x \<in> T} \<inter> ball a (1 / n / 2)"
show "\<exists>V. openin (top_of_set {x \<in> U n. f x \<in> T}) V \<and> a \<in> V \<and> negligible V"
proof (intro exI conjI)
have noxy: "norm(x - y) \<le> n * norm(f x - f y)" if "x \<in> V" "y \<in> V" for x y
using that unfolding U_def V_def mem_Collect_eq Int_iff mem_ball dist_norm
by (meson norm_triangle_half_r)
then have "inj_on f V"
by (force simp: inj_on_def)
then obtain g where g: "\<And>x. x \<in> V \<Longrightarrow> g(f x) = x"
by (metis inv_into_f_f)
have "\<exists>T' B. open T' \<and> f x \<in> T' \<and>
(\<forall>y\<in>f ` V \<inter> T \<inter> T'. norm (g y - g (f x)) \<le> B * norm (y - f x))"
if "f x \<in> T" "x \<in> V" for x
apply (rule_tac x = "ball (f x) 1" in exI)
using that noxy by (force simp: g)
then have "negligible (g ` (f ` V \<inter> T))"
by (force simp: \<open>negligible T\<close> negligible_Int intro!: negligible_locally_Lipschitz_image)
moreover have "V \<subseteq> g ` (f ` V \<inter> T)"
by (force simp: g image_iff V_def)
ultimately show "negligible V"
by (rule negligible_subset)
qed (use a fa V_def that in auto)
qed
with negligible_countable_Union have "negligible (\<Union>n \<in> {0<..}. {x. x \<in> U n \<and> f x \<in> T})"
by auto
moreover have "{x \<in> S. f x \<in> T} \<subseteq> (\<Union>n \<in> {0<..}. {x. x \<in> U n \<and> f x \<in> T})"
proof clarsimp
fix x
assume "x \<in> S" and "f x \<in> T"
then obtain inj: "inj(f' x)" and der: "(f has_derivative f' x) (at x within S)"
using assms by metis
moreover have "linear(f' x)"
and eps: "\<And>\<epsilon>. \<epsilon> > 0 \<Longrightarrow> \<exists>\<delta>>0. \<forall>y\<in>S. norm (y - x) < \<delta> \<longrightarrow>
norm (f y - f x - f' x (y - x)) \<le> \<epsilon> * norm (y - x)"
using der by (auto simp: has_derivative_within_alt linear_linear)
ultimately obtain g where "linear g" and g: "g \<circ> f' x = id"
using linear_injective_left_inverse by metis
then obtain B where "B > 0" and B: "\<And>z. B * norm z \<le> norm(f' x z)"
using linear_invertible_bounded_below_pos \<open>linear (f' x)\<close> by blast
then obtain i where "i \<noteq> 0" and i: "1 / real i < B"
by (metis (full_types) inverse_eq_divide real_arch_invD)
then obtain \<delta> where "\<delta> > 0"
and \<delta>: "\<And>y. \<lbrakk>y\<in>S; norm (y - x) < \<delta>\<rbrakk> \<Longrightarrow>
norm (f y - f x - f' x (y - x)) \<le> (B - 1 / real i) * norm (y - x)"
using eps [of "B - 1/i"] by auto
then obtain j where "j \<noteq> 0" and j: "inverse (real j) < \<delta>"
using real_arch_inverse by blast
have "norm (y - x)/(max i j) \<le> norm (f y - f x)"
if "y \<in> S" and less: "norm (y - x) < 1 / (max i j)" for y
proof -
have "1 / real (max i j) < \<delta>"
using j \<open>j \<noteq> 0\<close> \<open>0 < \<delta>\<close>
by (auto simp: field_split_simps max_mult_distrib_left of_nat_max)
then have "norm (y - x) < \<delta>"
using less by linarith
with \<delta> \<open>y \<in> S\<close> have le: "norm (f y - f x - f' x (y - x)) \<le> B * norm (y - x) - norm (y - x)/i"
by (auto simp: algebra_simps)
have *: "\<lbrakk>norm(f - f' - y) \<le> b - c; b \<le> norm y; d \<le> c\<rbrakk> \<Longrightarrow> d \<le> norm(f - f')"
for b c d and y f f'::'a
using norm_triangle_ineq3 [of "f - f'" y] by simp
show ?thesis
apply (rule * [OF le B])
using \<open>i \<noteq> 0\<close> \<open>j \<noteq> 0\<close> by (simp add: field_split_simps max_mult_distrib_left of_nat_max less_max_iff_disj)
qed
with \<open>x \<in> S\<close> \<open>i \<noteq> 0\<close> \<open>j \<noteq> 0\<close> show "\<exists>n\<in>{0<..}. x \<in> U n"
by (rule_tac x="max i j" in bexI) (auto simp: field_simps U_def less_max_iff_disj)
qed
ultimately show ?thesis
by (rule negligible_subset)
qed
lemma absolutely_integrable_Un:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes S: "f absolutely_integrable_on S" and T: "f absolutely_integrable_on T"
shows "f absolutely_integrable_on (S \<union> T)"
proof -
have [simp]: "{x. (if x \<in> A then f x else 0) \<noteq> 0} = {x \<in> A. f x \<noteq> 0}" for A
by auto
let ?ST = "{x \<in> S. f x \<noteq> 0} \<inter> {x \<in> T. f x \<noteq> 0}"
have "?ST \<in> sets lebesgue"
proof (rule Sigma_Algebra.sets.Int)
have "f integrable_on S"
using S absolutely_integrable_on_def by blast
then have "(\<lambda>x. if x \<in> S then f x else 0) integrable_on UNIV"
by (simp add: integrable_restrict_UNIV)
then have borel: "(\<lambda>x. if x \<in> S then f x else 0) \<in> borel_measurable (lebesgue_on UNIV)"
using integrable_imp_measurable lebesgue_on_UNIV_eq by blast
then show "{x \<in> S. f x \<noteq> 0} \<in> sets lebesgue"
unfolding borel_measurable_vimage_open
by (rule allE [where x = "-{0}"]) auto
next
have "f integrable_on T"
using T absolutely_integrable_on_def by blast
then have "(\<lambda>x. if x \<in> T then f x else 0) integrable_on UNIV"
by (simp add: integrable_restrict_UNIV)
then have borel: "(\<lambda>x. if x \<in> T then f x else 0) \<in> borel_measurable (lebesgue_on UNIV)"
using integrable_imp_measurable lebesgue_on_UNIV_eq by blast
then show "{x \<in> T. f x \<noteq> 0} \<in> sets lebesgue"
unfolding borel_measurable_vimage_open
by (rule allE [where x = "-{0}"]) auto
qed
then have "f absolutely_integrable_on ?ST"
by (rule set_integrable_subset [OF S]) auto
then have Int: "(\<lambda>x. if x \<in> ?ST then f x else 0) absolutely_integrable_on UNIV"
using absolutely_integrable_restrict_UNIV by blast
have "(\<lambda>x. if x \<in> S then f x else 0) absolutely_integrable_on UNIV"
"(\<lambda>x. if x \<in> T then f x else 0) absolutely_integrable_on UNIV"
using S T absolutely_integrable_restrict_UNIV by blast+
then have "(\<lambda>x. (if x \<in> S then f x else 0) + (if x \<in> T then f x else 0)) absolutely_integrable_on UNIV"
by (rule set_integral_add)
then have "(\<lambda>x. ((if x \<in> S then f x else 0) + (if x \<in> T then f x else 0)) - (if x \<in> ?ST then f x else 0)) absolutely_integrable_on UNIV"
using Int by (rule set_integral_diff)
then have "(\<lambda>x. if x \<in> S \<union> T then f x else 0) absolutely_integrable_on UNIV"
by (rule absolutely_integrable_spike) (auto intro: empty_imp_negligible)
then show ?thesis
unfolding absolutely_integrable_restrict_UNIV .
qed
lemma absolutely_integrable_on_combine:
fixes f :: "real \<Rightarrow> 'a::euclidean_space"
assumes "f absolutely_integrable_on {a..c}"
and "f absolutely_integrable_on {c..b}"
and "a \<le> c"
and "c \<le> b"
shows "f absolutely_integrable_on {a..b}"
by (metis absolutely_integrable_Un assms ivl_disj_un_two_touch(4))
lemma uniform_limit_set_lebesgue_integral_at_top:
fixes f :: "'a \<Rightarrow> real \<Rightarrow> 'b::{banach, second_countable_topology}"
and g :: "real \<Rightarrow> real"
assumes bound: "\<And>x y. x \<in> A \<Longrightarrow> y \<ge> a \<Longrightarrow> norm (f x y) \<le> g y"
assumes integrable: "set_integrable M {a..} g"
assumes measurable: "\<And>x. x \<in> A \<Longrightarrow> set_borel_measurable M {a..} (f x)"
assumes "sets borel \<subseteq> sets M"
shows "uniform_limit A (\<lambda>b x. LINT y:{a..b}|M. f x y) (\<lambda>x. LINT y:{a..}|M. f x y) at_top"
proof (cases "A = {}")
case False
then obtain x where x: "x \<in> A" by auto
have g_nonneg: "g y \<ge> 0" if "y \<ge> a" for y
proof -
have "0 \<le> norm (f x y)" by simp
also have "\<dots> \<le> g y" using bound[OF x that] by simp
finally show ?thesis .
qed
have integrable': "set_integrable M {a..} (\<lambda>y. f x y)" if "x \<in> A" for x
unfolding set_integrable_def
proof (rule Bochner_Integration.integrable_bound)
show "integrable M (\<lambda>x. indicator {a..} x * g x)"
using integrable by (simp add: set_integrable_def)
show "(\<lambda>y. indicat_real {a..} y *\<^sub>R f x y) \<in> borel_measurable M" using measurable[OF that]
by (simp add: set_borel_measurable_def)
show "AE y in M. norm (indicat_real {a..} y *\<^sub>R f x y) \<le> norm (indicat_real {a..} y * g y)"
using bound[OF that] by (intro AE_I2) (auto simp: indicator_def g_nonneg)
qed
show ?thesis
proof (rule uniform_limitI)
fix e :: real assume e: "e > 0"
have sets [intro]: "A \<in> sets M" if "A \<in> sets borel" for A
using that assms by blast
have "((\<lambda>b. LINT y:{a..b}|M. g y) \<longlongrightarrow> (LINT y:{a..}|M. g y)) at_top"
by (intro tendsto_set_lebesgue_integral_at_top assms sets) auto
with e obtain b0 :: real where b0: "\<forall>b\<ge>b0. \<bar>(LINT y:{a..}|M. g y) - (LINT y:{a..b}|M. g y)\<bar> < e"
by (auto simp: tendsto_iff eventually_at_top_linorder dist_real_def abs_minus_commute)
define b where "b = max a b0"
have "a \<le> b" by (simp add: b_def)
from b0 have "\<bar>(LINT y:{a..}|M. g y) - (LINT y:{a..b}|M. g y)\<bar> < e"
by (auto simp: b_def)
also have "{a..} = {a..b} \<union> {b<..}" by (auto simp: b_def)
also have "\<bar>(LINT y:\<dots>|M. g y) - (LINT y:{a..b}|M. g y)\<bar> = \<bar>(LINT y:{b<..}|M. g y)\<bar>"
using \<open>a \<le> b\<close> by (subst set_integral_Un) (auto intro!: set_integrable_subset[OF integrable])
also have "(LINT y:{b<..}|M. g y) \<ge> 0"
using g_nonneg \<open>a \<le> b\<close> unfolding set_lebesgue_integral_def
by (intro Bochner_Integration.integral_nonneg) (auto simp: indicator_def)
hence "\<bar>(LINT y:{b<..}|M. g y)\<bar> = (LINT y:{b<..}|M. g y)" by simp
finally have less: "(LINT y:{b<..}|M. g y) < e" .
have "eventually (\<lambda>b. b \<ge> b0) at_top" by (rule eventually_ge_at_top)
moreover have "eventually (\<lambda>b. b \<ge> a) at_top" by (rule eventually_ge_at_top)
ultimately show "eventually (\<lambda>b. \<forall>x\<in>A.
dist (LINT y:{a..b}|M. f x y) (LINT y:{a..}|M. f x y) < e) at_top"
proof eventually_elim
case (elim b)
show ?case
proof
fix x assume x: "x \<in> A"
have "dist (LINT y:{a..b}|M. f x y) (LINT y:{a..}|M. f x y) =
norm ((LINT y:{a..}|M. f x y) - (LINT y:{a..b}|M. f x y))"
by (simp add: dist_norm norm_minus_commute)
also have "{a..} = {a..b} \<union> {b<..}" using elim by auto
also have "(LINT y:\<dots>|M. f x y) - (LINT y:{a..b}|M. f x y) = (LINT y:{b<..}|M. f x y)"
using elim x
by (subst set_integral_Un) (auto intro!: set_integrable_subset[OF integrable'])
also have "norm \<dots> \<le> (LINT y:{b<..}|M. norm (f x y))" using elim x
by (intro set_integral_norm_bound set_integrable_subset[OF integrable']) auto
also have "\<dots> \<le> (LINT y:{b<..}|M. g y)" using elim x bound g_nonneg
by (intro set_integral_mono set_integrable_norm set_integrable_subset[OF integrable']
set_integrable_subset[OF integrable]) auto
also have "(LINT y:{b<..}|M. g y) \<ge> 0"
using g_nonneg \<open>a \<le> b\<close> unfolding set_lebesgue_integral_def
by (intro Bochner_Integration.integral_nonneg) (auto simp: indicator_def)
hence "(LINT y:{b<..}|M. g y) = \<bar>(LINT y:{b<..}|M. g y)\<bar>" by simp
also have "\<dots> = \<bar>(LINT y:{a..b} \<union> {b<..}|M. g y) - (LINT y:{a..b}|M. g y)\<bar>"
using elim by (subst set_integral_Un) (auto intro!: set_integrable_subset[OF integrable])
also have "{a..b} \<union> {b<..} = {a..}" using elim by auto
also have "\<bar>(LINT y:{a..}|M. g y) - (LINT y:{a..b}|M. g y)\<bar> < e"
using b0 elim by blast
finally show "dist (LINT y:{a..b}|M. f x y) (LINT y:{a..}|M. f x y) < e" .
qed
qed
qed
qed auto
subsubsection\<open>Differentiability of inverse function (most basic form)\<close>
proposition has_derivative_inverse_within:
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space"
assumes der_f: "(f has_derivative f') (at a within S)"
and cont_g: "continuous (at (f a) within f ` S) g"
and "a \<in> S" "linear g'" and id: "g' \<circ> f' = id"
and gf: "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x"
shows "(g has_derivative g') (at (f a) within f ` S)"
proof -
have [simp]: "g' (f' x) = x" for x
by (simp add: local.id pointfree_idE)
have "bounded_linear f'"
and f': "\<And>e. e>0 \<Longrightarrow> \<exists>d>0. \<forall>y\<in>S. norm (y - a) < d \<longrightarrow>
norm (f y - f a - f' (y - a)) \<le> e * norm (y - a)"
using der_f by (auto simp: has_derivative_within_alt)
obtain C where "C > 0" and C: "\<And>x. norm (g' x) \<le> C * norm x"
using linear_bounded_pos [OF \<open>linear g'\<close>] by metis
obtain B k where "B > 0" "k > 0"
and Bk: "\<And>x. \<lbrakk>x \<in> S; norm(f x - f a) < k\<rbrakk> \<Longrightarrow> norm(x - a) \<le> B * norm(f x - f a)"
proof -
obtain B where "B > 0" and B: "\<And>x. B * norm x \<le> norm (f' x)"
using linear_inj_bounded_below_pos [of f'] \<open>linear g'\<close> id der_f has_derivative_linear
linear_invertible_bounded_below_pos by blast
then obtain d where "d>0"
and d: "\<And>y. \<lbrakk>y \<in> S; norm (y - a) < d\<rbrakk> \<Longrightarrow>
norm (f y - f a - f' (y - a)) \<le> B / 2 * norm (y - a)"
using f' [of "B/2"] by auto
then obtain e where "e > 0"
and e: "\<And>x. \<lbrakk>x \<in> S; norm (f x - f a) < e\<rbrakk> \<Longrightarrow> norm (g (f x) - g (f a)) < d"
using cont_g by (auto simp: continuous_within_eps_delta dist_norm)
show thesis
proof
show "2/B > 0"
using \<open>B > 0\<close> by simp
show "norm (x - a) \<le> 2 / B * norm (f x - f a)"
if "x \<in> S" "norm (f x - f a) < e" for x
proof -
have xa: "norm (x - a) < d"
using e [OF that] gf by (simp add: \<open>a \<in> S\<close> that)
have *: "\<lbrakk>norm(y - f') \<le> B / 2 * norm x; B * norm x \<le> norm f'\<rbrakk>
\<Longrightarrow> norm y \<ge> B / 2 * norm x" for y f'::'b and x::'a
using norm_triangle_ineq3 [of y f'] by linarith
show ?thesis
using * [OF d [OF \<open>x \<in> S\<close> xa] B] \<open>B > 0\<close> by (simp add: field_simps)
qed
qed (use \<open>e > 0\<close> in auto)
qed
show ?thesis
unfolding has_derivative_within_alt
proof (intro conjI impI allI)
show "bounded_linear g'"
using \<open>linear g'\<close> by (simp add: linear_linear)
next
fix e :: "real"
assume "e > 0"
then obtain d where "d>0"
and d: "\<And>y. \<lbrakk>y \<in> S; norm (y - a) < d\<rbrakk> \<Longrightarrow>
norm (f y - f a - f' (y - a)) \<le> e / (B * C) * norm (y - a)"
using f' [of "e / (B * C)"] \<open>B > 0\<close> \<open>C > 0\<close> by auto
have "norm (x - a - g' (f x - f a)) \<le> e * norm (f x - f a)"
if "x \<in> S" and lt_k: "norm (f x - f a) < k" and lt_dB: "norm (f x - f a) < d/B" for x
proof -
have "norm (x - a) \<le> B * norm(f x - f a)"
using Bk lt_k \<open>x \<in> S\<close> by blast
also have "\<dots> < d"
by (metis \<open>0 < B\<close> lt_dB mult.commute pos_less_divide_eq)
finally have lt_d: "norm (x - a) < d" .
have "norm (x - a - g' (f x - f a)) \<le> norm(g'(f x - f a - (f' (x - a))))"
by (simp add: linear_diff [OF \<open>linear g'\<close>] norm_minus_commute)
also have "\<dots> \<le> C * norm (f x - f a - f' (x - a))"
using C by blast
also have "\<dots> \<le> e * norm (f x - f a)"
proof -
have "norm (f x - f a - f' (x - a)) \<le> e / (B * C) * norm (x - a)"
using d [OF \<open>x \<in> S\<close> lt_d] .
also have "\<dots> \<le> (norm (f x - f a) * e) / C"
using \<open>B > 0\<close> \<open>C > 0\<close> \<open>e > 0\<close> by (simp add: field_simps Bk lt_k \<open>x \<in> S\<close>)
finally show ?thesis
using \<open>C > 0\<close> by (simp add: field_simps)
qed
finally show ?thesis .
qed
then show "\<exists>d>0. \<forall>y\<in>f ` S.
norm (y - f a) < d \<longrightarrow>
norm (g y - g (f a) - g' (y - f a)) \<le> e * norm (y - f a)"
apply (rule_tac x="min k (d / B)" in exI)
using \<open>k > 0\<close> \<open>B > 0\<close> \<open>d > 0\<close> \<open>a \<in> S\<close> by (auto simp: gf)
qed
qed
end
diff --git a/src/HOL/Analysis/Equivalence_Measurable_On_Borel.thy b/src/HOL/Analysis/Equivalence_Measurable_On_Borel.thy
--- a/src/HOL/Analysis/Equivalence_Measurable_On_Borel.thy
+++ b/src/HOL/Analysis/Equivalence_Measurable_On_Borel.thy
@@ -1,1743 +1,1743 @@
(* Title: HOL/Analysis/Equivalence_Measurable_On_Borel
Author: LC Paulson (some material ported from HOL Light)
*)
section\<open>Equivalence Between Classical Borel Measurability and HOL Light's\<close>
theory Equivalence_Measurable_On_Borel
imports Equivalence_Lebesgue_Henstock_Integration Improper_Integral Continuous_Extension
begin
(*Borrowed from Ergodic_Theory/SG_Library_Complement*)
abbreviation sym_diff :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" where
"sym_diff A B \<equiv> ((A - B) \<union> (B-A))"
subsection\<open>Austin's Lemma\<close>
lemma Austin_Lemma:
fixes \<D> :: "'a::euclidean_space set set"
assumes "finite \<D>" and \<D>: "\<And>D. D \<in> \<D> \<Longrightarrow> \<exists>k a b. D = cbox a b \<and> (\<forall>i \<in> Basis. b\<bullet>i - a\<bullet>i = k)"
obtains \<C> where "\<C> \<subseteq> \<D>" "pairwise disjnt \<C>"
"measure lebesgue (\<Union>\<C>) \<ge> measure lebesgue (\<Union>\<D>) / 3 ^ (DIM('a))"
using assms
proof (induction "card \<D>" arbitrary: \<D> thesis rule: less_induct)
case less
show ?case
proof (cases "\<D> = {}")
case True
then show thesis
using less by auto
next
case False
then have "Max (Sigma_Algebra.measure lebesgue ` \<D>) \<in> Sigma_Algebra.measure lebesgue ` \<D>"
using Max_in finite_imageI \<open>finite \<D>\<close> by blast
then obtain D where "D \<in> \<D>" and "measure lebesgue D = Max (measure lebesgue ` \<D>)"
by auto
then have D: "\<And>C. C \<in> \<D> \<Longrightarrow> measure lebesgue C \<le> measure lebesgue D"
by (simp add: \<open>finite \<D>\<close>)
let ?\<E> = "{C. C \<in> \<D> - {D} \<and> disjnt C D}"
obtain \<D>' where \<D>'sub: "\<D>' \<subseteq> ?\<E>" and \<D>'dis: "pairwise disjnt \<D>'"
and \<D>'m: "measure lebesgue (\<Union>\<D>') \<ge> measure lebesgue (\<Union>?\<E>) / 3 ^ (DIM('a))"
proof (rule less.hyps)
have *: "?\<E> \<subset> \<D>"
using \<open>D \<in> \<D>\<close> by auto
then show "card ?\<E> < card \<D>" "finite ?\<E>"
by (auto simp: \<open>finite \<D>\<close> psubset_card_mono)
show "\<exists>k a b. D = cbox a b \<and> (\<forall>i\<in>Basis. b \<bullet> i - a \<bullet> i = k)" if "D \<in> ?\<E>" for D
using less.prems(3) that by auto
qed
then have [simp]: "\<Union>\<D>' - D = \<Union>\<D>'"
by (auto simp: disjnt_iff)
show ?thesis
proof (rule less.prems)
show "insert D \<D>' \<subseteq> \<D>"
using \<D>'sub \<open>D \<in> \<D>\<close> by blast
show "disjoint (insert D \<D>')"
using \<D>'dis \<D>'sub by (fastforce simp add: pairwise_def disjnt_sym)
obtain a3 b3 where m3: "content (cbox a3 b3) = 3 ^ DIM('a) * measure lebesgue D"
and sub3: "\<And>C. \<lbrakk>C \<in> \<D>; \<not> disjnt C D\<rbrakk> \<Longrightarrow> C \<subseteq> cbox a3 b3"
proof -
obtain k a b where ab: "D = cbox a b" and k: "\<And>i. i \<in> Basis \<Longrightarrow> b\<bullet>i - a\<bullet>i = k"
using less.prems \<open>D \<in> \<D>\<close> by meson
then have eqk: "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i \<longleftrightarrow> k \<ge> 0"
by force
show thesis
proof
let ?a = "(a + b) /\<^sub>R 2 - (3/2) *\<^sub>R (b - a)"
let ?b = "(a + b) /\<^sub>R 2 + (3/2) *\<^sub>R (b - a)"
have eq: "(\<Prod>i\<in>Basis. b \<bullet> i * 3 - a \<bullet> i * 3) = (\<Prod>i\<in>Basis. b \<bullet> i - a \<bullet> i) * 3 ^ DIM('a)"
by (simp add: comm_monoid_mult_class.prod.distrib flip: left_diff_distrib inner_diff_left)
show "content (cbox ?a ?b) = 3 ^ DIM('a) * measure lebesgue D"
by (simp add: content_cbox_if box_eq_empty algebra_simps eq ab k)
show "C \<subseteq> cbox ?a ?b" if "C \<in> \<D>" and CD: "\<not> disjnt C D" for C
proof -
obtain k' a' b' where ab': "C = cbox a' b'" and k': "\<And>i. i \<in> Basis \<Longrightarrow> b'\<bullet>i - a'\<bullet>i = k'"
using less.prems \<open>C \<in> \<D>\<close> by meson
then have eqk': "\<And>i. i \<in> Basis \<Longrightarrow> a' \<bullet> i \<le> b' \<bullet> i \<longleftrightarrow> k' \<ge> 0"
by force
show ?thesis
proof (clarsimp simp add: disjoint_interval disjnt_def ab ab' not_less subset_box algebra_simps)
show "a \<bullet> i * 2 \<le> a' \<bullet> i + b \<bullet> i \<and> a \<bullet> i + b' \<bullet> i \<le> b \<bullet> i * 2"
if * [rule_format]: "\<forall>j\<in>Basis. a' \<bullet> j \<le> b' \<bullet> j" and "i \<in> Basis" for i
proof -
have "a' \<bullet> i \<le> b' \<bullet> i \<and> a \<bullet> i \<le> b \<bullet> i \<and> a \<bullet> i \<le> b' \<bullet> i \<and> a' \<bullet> i \<le> b \<bullet> i"
using \<open>i \<in> Basis\<close> CD by (simp_all add: disjoint_interval disjnt_def ab ab' not_less)
then show ?thesis
using D [OF \<open>C \<in> \<D>\<close>] \<open>i \<in> Basis\<close>
apply (simp add: ab ab' k k' eqk eqk' content_cbox_cases)
using k k' by fastforce
qed
qed
qed
qed
qed
have \<D>lm: "\<And>D. D \<in> \<D> \<Longrightarrow> D \<in> lmeasurable"
using less.prems(3) by blast
have "measure lebesgue (\<Union>\<D>) \<le> measure lebesgue (cbox a3 b3 \<union> (\<Union>\<D> - cbox a3 b3))"
proof (rule measure_mono_fmeasurable)
show "\<Union>\<D> \<in> sets lebesgue"
using \<D>lm \<open>finite \<D>\<close> by blast
show "cbox a3 b3 \<union> (\<Union>\<D> - cbox a3 b3) \<in> lmeasurable"
by (simp add: \<D>lm fmeasurable.Un fmeasurable.finite_Union less.prems(2) subset_eq)
qed auto
also have "\<dots> = content (cbox a3 b3) + measure lebesgue (\<Union>\<D> - cbox a3 b3)"
by (simp add: \<D>lm fmeasurable.finite_Union less.prems(2) measure_Un2 subsetI)
also have "\<dots> \<le> (measure lebesgue D + measure lebesgue (\<Union>\<D>')) * 3 ^ DIM('a)"
proof -
have "(\<Union>\<D> - cbox a3 b3) \<subseteq> \<Union>?\<E>"
using sub3 by fastforce
then have "measure lebesgue (\<Union>\<D> - cbox a3 b3) \<le> measure lebesgue (\<Union>?\<E>)"
proof (rule measure_mono_fmeasurable)
show "\<Union> \<D> - cbox a3 b3 \<in> sets lebesgue"
by (simp add: \<D>lm fmeasurableD less.prems(2) sets.Diff sets.finite_Union subsetI)
show "\<Union> {C \<in> \<D> - {D}. disjnt C D} \<in> lmeasurable"
using \<D>lm less.prems(2) by auto
qed
then have "measure lebesgue (\<Union>\<D> - cbox a3 b3) / 3 ^ DIM('a) \<le> measure lebesgue (\<Union> \<D>')"
using \<D>'m by (simp add: field_split_simps)
then show ?thesis
by (simp add: m3 field_simps)
qed
also have "\<dots> \<le> measure lebesgue (\<Union>(insert D \<D>')) * 3 ^ DIM('a)"
proof (simp add: \<D>lm \<open>D \<in> \<D>\<close>)
show "measure lebesgue D + measure lebesgue (\<Union>\<D>') \<le> measure lebesgue (D \<union> \<Union> \<D>')"
proof (subst measure_Un2)
show "\<Union> \<D>' \<in> lmeasurable"
by (meson \<D>lm \<open>insert D \<D>' \<subseteq> \<D>\<close> fmeasurable.finite_Union less.prems(2) finite_subset subset_eq subset_insertI)
show "measure lebesgue D + measure lebesgue (\<Union> \<D>') \<le> measure lebesgue D + measure lebesgue (\<Union> \<D>' - D)"
using \<open>insert D \<D>' \<subseteq> \<D>\<close> infinite_super less.prems(2) by force
qed (simp add: \<D>lm \<open>D \<in> \<D>\<close>)
qed
finally show "measure lebesgue (\<Union>\<D>) / 3 ^ DIM('a) \<le> measure lebesgue (\<Union>(insert D \<D>'))"
by (simp add: field_split_simps)
qed
qed
qed
subsection\<open>A differentiability-like property of the indefinite integral. \<close>
proposition integrable_ccontinuous_explicit:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "\<And>a b::'a. f integrable_on cbox a b"
obtains N where
"negligible N"
"\<And>x e. \<lbrakk>x \<notin> N; 0 < e\<rbrakk> \<Longrightarrow>
\<exists>d>0. \<forall>h. 0 < h \<and> h < d \<longrightarrow>
norm(integral (cbox x (x + h *\<^sub>R One)) f /\<^sub>R h ^ DIM('a) - f x) < e"
proof -
define BOX where "BOX \<equiv> \<lambda>h. \<lambda>x::'a. cbox x (x + h *\<^sub>R One)"
define BOX2 where "BOX2 \<equiv> \<lambda>h. \<lambda>x::'a. cbox (x - h *\<^sub>R One) (x + h *\<^sub>R One)"
define i where "i \<equiv> \<lambda>h x. integral (BOX h x) f /\<^sub>R h ^ DIM('a)"
define \<Psi> where "\<Psi> \<equiv> \<lambda>x r. \<forall>d>0. \<exists>h. 0 < h \<and> h < d \<and> r \<le> norm(i h x - f x)"
let ?N = "{x. \<exists>e>0. \<Psi> x e}"
have "\<exists>N. negligible N \<and> (\<forall>x e. x \<notin> N \<and> 0 < e \<longrightarrow> \<not> \<Psi> x e)"
proof (rule exI ; intro conjI allI impI)
let ?M = "\<Union>n. {x. \<Psi> x (inverse(real n + 1))}"
have "negligible ({x. \<Psi> x \<mu>} \<inter> cbox a b)"
if "\<mu> > 0" for a b \<mu>
proof (cases "negligible(cbox a b)")
case True
then show ?thesis
by (simp add: negligible_Int)
next
case False
then have "box a b \<noteq> {}"
by (simp add: negligible_interval)
then have ab: "\<And>i. i \<in> Basis \<Longrightarrow> a\<bullet>i < b\<bullet>i"
by (simp add: box_ne_empty)
show ?thesis
unfolding negligible_outer_le
proof (intro allI impI)
fix e::real
let ?ee = "(e * \<mu>) / 2 / 6 ^ (DIM('a))"
assume "e > 0"
then have gt0: "?ee > 0"
using \<open>\<mu> > 0\<close> by auto
have f': "f integrable_on cbox (a - One) (b + One)"
using assms by blast
obtain \<gamma> where "gauge \<gamma>"
and \<gamma>: "\<And>p. \<lbrakk>p tagged_partial_division_of (cbox (a - One) (b + One)); \<gamma> fine p\<rbrakk>
\<Longrightarrow> (\<Sum>(x, k)\<in>p. norm (content k *\<^sub>R f x - integral k f)) < ?ee"
using Henstock_lemma [OF f' gt0] that by auto
let ?E = "{x. x \<in> cbox a b \<and> \<Psi> x \<mu>}"
have "\<exists>h>0. BOX h x \<subseteq> \<gamma> x \<and>
BOX h x \<subseteq> cbox (a - One) (b + One) \<and> \<mu> \<le> norm (i h x - f x)"
if "x \<in> cbox a b" "\<Psi> x \<mu>" for x
proof -
obtain d where "d > 0" and d: "ball x d \<subseteq> \<gamma> x"
using gaugeD [OF \<open>gauge \<gamma>\<close>, of x] openE by blast
then obtain h where "0 < h" "h < 1" and hless: "h < d / real DIM('a)"
and mule: "\<mu> \<le> norm (i h x - f x)"
using \<open>\<Psi> x \<mu>\<close> [unfolded \<Psi>_def, rule_format, of "min 1 (d / DIM('a))"]
by auto
show ?thesis
proof (intro exI conjI)
show "0 < h" "\<mu> \<le> norm (i h x - f x)" by fact+
have "BOX h x \<subseteq> ball x d"
proof (clarsimp simp: BOX_def mem_box dist_norm algebra_simps)
fix y
assume "\<forall>i\<in>Basis. x \<bullet> i \<le> y \<bullet> i \<and> y \<bullet> i \<le> h + x \<bullet> i"
then have lt: "\<bar>(x - y) \<bullet> i\<bar> < d / real DIM('a)" if "i \<in> Basis" for i
using hless that by (force simp: inner_diff_left)
have "norm (x - y) \<le> (\<Sum>i\<in>Basis. \<bar>(x - y) \<bullet> i\<bar>)"
using norm_le_l1 by blast
also have "\<dots> < d"
using sum_bounded_above_strict [of Basis "\<lambda>i. \<bar>(x - y) \<bullet> i\<bar>" "d / DIM('a)", OF lt]
by auto
finally show "norm (x - y) < d" .
qed
with d show "BOX h x \<subseteq> \<gamma> x"
by blast
show "BOX h x \<subseteq> cbox (a - One) (b + One)"
using that \<open>h < 1\<close>
by (force simp: BOX_def mem_box algebra_simps intro: subset_box_imp)
qed
qed
then obtain \<eta> where h0: "\<And>x. x \<in> ?E \<Longrightarrow> \<eta> x > 0"
and BOX_\<gamma>: "\<And>x. x \<in> ?E \<Longrightarrow> BOX (\<eta> x) x \<subseteq> \<gamma> x"
and "\<And>x. x \<in> ?E \<Longrightarrow> BOX (\<eta> x) x \<subseteq> cbox (a - One) (b + One) \<and> \<mu> \<le> norm (i (\<eta> x) x - f x)"
by simp metis
then have BOX_cbox: "\<And>x. x \<in> ?E \<Longrightarrow> BOX (\<eta> x) x \<subseteq> cbox (a - One) (b + One)"
and \<mu>_le: "\<And>x. x \<in> ?E \<Longrightarrow> \<mu> \<le> norm (i (\<eta> x) x - f x)"
by blast+
define \<gamma>' where "\<gamma>' \<equiv> \<lambda>x. if x \<in> cbox a b \<and> \<Psi> x \<mu> then ball x (\<eta> x) else \<gamma> x"
have "gauge \<gamma>'"
using \<open>gauge \<gamma>\<close> by (auto simp: h0 gauge_def \<gamma>'_def)
obtain \<D> where "countable \<D>"
and \<D>: "\<Union>\<D> \<subseteq> cbox a b"
"\<And>K. K \<in> \<D> \<Longrightarrow> interior K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
and Dcovered: "\<And>K. K \<in> \<D> \<Longrightarrow> \<exists>x. x \<in> cbox a b \<and> \<Psi> x \<mu> \<and> x \<in> K \<and> K \<subseteq> \<gamma>' x"
and subUD: "?E \<subseteq> \<Union>\<D>"
by (rule covering_lemma [of ?E a b \<gamma>']) (simp_all add: Bex_def \<open>box a b \<noteq> {}\<close> \<open>gauge \<gamma>'\<close>)
then have "\<D> \<subseteq> sets lebesgue"
by fastforce
show "\<exists>T. {x. \<Psi> x \<mu>} \<inter> cbox a b \<subseteq> T \<and> T \<in> lmeasurable \<and> measure lebesgue T \<le> e"
proof (intro exI conjI)
show "{x. \<Psi> x \<mu>} \<inter> cbox a b \<subseteq> \<Union>\<D>"
apply auto
using subUD by auto
have mUE: "measure lebesgue (\<Union> \<E>) \<le> measure lebesgue (cbox a b)"
if "\<E> \<subseteq> \<D>" "finite \<E>" for \<E>
proof (rule measure_mono_fmeasurable)
show "\<Union> \<E> \<subseteq> cbox a b"
using \<D>(1) that(1) by blast
show "\<Union> \<E> \<in> sets lebesgue"
by (metis \<D>(2) fmeasurable.finite_Union fmeasurableD lmeasurable_cbox subset_eq that)
qed auto
then show "\<Union>\<D> \<in> lmeasurable"
by (metis \<D>(2) \<open>countable \<D>\<close> fmeasurable_Union_bound lmeasurable_cbox)
then have leab: "measure lebesgue (\<Union>\<D>) \<le> measure lebesgue (cbox a b)"
by (meson \<D>(1) fmeasurableD lmeasurable_cbox measure_mono_fmeasurable)
obtain \<F> where "\<F> \<subseteq> \<D>" "finite \<F>"
and \<F>: "measure lebesgue (\<Union>\<D>) \<le> 2 * measure lebesgue (\<Union>\<F>)"
proof (cases "measure lebesgue (\<Union>\<D>) = 0")
case True
then show ?thesis
by (force intro: that [where \<F> = "{}"])
next
case False
obtain \<F> where "\<F> \<subseteq> \<D>" "finite \<F>"
and \<F>: "measure lebesgue (\<Union>\<D>)/2 < measure lebesgue (\<Union>\<F>)"
proof (rule measure_countable_Union_approachable [of \<D> "measure lebesgue (\<Union>\<D>) / 2" "content (cbox a b)"])
show "countable \<D>"
by fact
show "0 < measure lebesgue (\<Union> \<D>) / 2"
using False by (simp add: zero_less_measure_iff)
show Dlm: "D \<in> lmeasurable" if "D \<in> \<D>" for D
using \<D>(2) that by blast
show "measure lebesgue (\<Union> \<F>) \<le> content (cbox a b)"
if "\<F> \<subseteq> \<D>" "finite \<F>" for \<F>
proof -
have "measure lebesgue (\<Union> \<F>) \<le> measure lebesgue (\<Union>\<D>)"
proof (rule measure_mono_fmeasurable)
show "\<Union> \<F> \<subseteq> \<Union> \<D>"
by (simp add: Sup_subset_mono \<open>\<F> \<subseteq> \<D>\<close>)
show "\<Union> \<F> \<in> sets lebesgue"
by (meson Dlm fmeasurableD sets.finite_Union subset_eq that)
show "\<Union> \<D> \<in> lmeasurable"
by fact
qed
also have "\<dots> \<le> measure lebesgue (cbox a b)"
proof (rule measure_mono_fmeasurable)
show "\<Union> \<D> \<in> sets lebesgue"
by (simp add: \<open>\<Union> \<D> \<in> lmeasurable\<close> fmeasurableD)
qed (auto simp:\<D>(1))
finally show ?thesis
by simp
qed
qed auto
then show ?thesis
using that by auto
qed
obtain tag where tag_in_E: "\<And>D. D \<in> \<D> \<Longrightarrow> tag D \<in> ?E"
and tag_in_self: "\<And>D. D \<in> \<D> \<Longrightarrow> tag D \<in> D"
and tag_sub: "\<And>D. D \<in> \<D> \<Longrightarrow> D \<subseteq> \<gamma>' (tag D)"
using Dcovered by simp metis
then have sub_ball_tag: "\<And>D. D \<in> \<D> \<Longrightarrow> D \<subseteq> ball (tag D) (\<eta> (tag D))"
by (simp add: \<gamma>'_def)
define \<Phi> where "\<Phi> \<equiv> \<lambda>D. BOX (\<eta>(tag D)) (tag D)"
define \<Phi>2 where "\<Phi>2 \<equiv> \<lambda>D. BOX2 (\<eta>(tag D)) (tag D)"
obtain \<C> where "\<C> \<subseteq> \<Phi>2 ` \<F>" "pairwise disjnt \<C>"
"measure lebesgue (\<Union>\<C>) \<ge> measure lebesgue (\<Union>(\<Phi>2`\<F>)) / 3 ^ (DIM('a))"
proof (rule Austin_Lemma)
show "finite (\<Phi>2`\<F>)"
using \<open>finite \<F>\<close> by blast
have "\<exists>k a b. \<Phi>2 D = cbox a b \<and> (\<forall>i\<in>Basis. b \<bullet> i - a \<bullet> i = k)" if "D \<in> \<F>" for D
apply (rule_tac x="2 * \<eta>(tag D)" in exI)
apply (rule_tac x="tag D - \<eta>(tag D) *\<^sub>R One" in exI)
apply (rule_tac x="tag D + \<eta>(tag D) *\<^sub>R One" in exI)
using that
apply (auto simp: \<Phi>2_def BOX2_def algebra_simps)
done
then show "\<And>D. D \<in> \<Phi>2 ` \<F> \<Longrightarrow> \<exists>k a b. D = cbox a b \<and> (\<forall>i\<in>Basis. b \<bullet> i - a \<bullet> i = k)"
by blast
qed auto
then obtain \<G> where "\<G> \<subseteq> \<F>" and disj: "pairwise disjnt (\<Phi>2 ` \<G>)"
and "measure lebesgue (\<Union>(\<Phi>2 ` \<G>)) \<ge> measure lebesgue (\<Union>(\<Phi>2`\<F>)) / 3 ^ (DIM('a))"
unfolding \<Phi>2_def subset_image_iff
by (meson empty_subsetI equals0D pairwise_imageI)
moreover
have "measure lebesgue (\<Union>(\<Phi>2 ` \<G>)) * 3 ^ DIM('a) \<le> e/2"
proof -
have "finite \<G>"
using \<open>finite \<F>\<close> \<open>\<G> \<subseteq> \<F>\<close> infinite_super by blast
have BOX2_m: "\<And>x. x \<in> tag ` \<G> \<Longrightarrow> BOX2 (\<eta> x) x \<in> lmeasurable"
by (auto simp: BOX2_def)
have BOX_m: "\<And>x. x \<in> tag ` \<G> \<Longrightarrow> BOX (\<eta> x) x \<in> lmeasurable"
by (auto simp: BOX_def)
have BOX_sub: "BOX (\<eta> x) x \<subseteq> BOX2 (\<eta> x) x" for x
by (auto simp: BOX_def BOX2_def subset_box algebra_simps)
have DISJ2: "BOX2 (\<eta> (tag X)) (tag X) \<inter> BOX2 (\<eta> (tag Y)) (tag Y) = {}"
if "X \<in> \<G>" "Y \<in> \<G>" "tag X \<noteq> tag Y" for X Y
proof -
obtain i where i: "i \<in> Basis" "tag X \<bullet> i \<noteq> tag Y \<bullet> i"
using \<open>tag X \<noteq> tag Y\<close> by (auto simp: euclidean_eq_iff [of "tag X"])
have XY: "X \<in> \<D>" "Y \<in> \<D>"
using \<open>\<F> \<subseteq> \<D>\<close> \<open>\<G> \<subseteq> \<F>\<close> that by auto
then have "0 \<le> \<eta> (tag X)" "0 \<le> \<eta> (tag Y)"
by (meson h0 le_cases not_le tag_in_E)+
with XY i have "BOX2 (\<eta> (tag X)) (tag X) \<noteq> BOX2 (\<eta> (tag Y)) (tag Y)"
unfolding eq_iff
by (fastforce simp add: BOX2_def subset_box algebra_simps)
then show ?thesis
using disj that by (auto simp: pairwise_def disjnt_def \<Phi>2_def)
qed
then have BOX2_disj: "pairwise (\<lambda>x y. negligible (BOX2 (\<eta> x) x \<inter> BOX2 (\<eta> y) y)) (tag ` \<G>)"
by (simp add: pairwise_imageI)
then have BOX_disj: "pairwise (\<lambda>x y. negligible (BOX (\<eta> x) x \<inter> BOX (\<eta> y) y)) (tag ` \<G>)"
proof (rule pairwise_mono)
show "negligible (BOX (\<eta> x) x \<inter> BOX (\<eta> y) y)"
if "negligible (BOX2 (\<eta> x) x \<inter> BOX2 (\<eta> y) y)" for x y
by (metis (no_types, hide_lams) that Int_mono negligible_subset BOX_sub)
qed auto
have eq: "\<And>box. (\<lambda>D. box (\<eta> (tag D)) (tag D)) ` \<G> = (\<lambda>t. box (\<eta> t) t) ` tag ` \<G>"
by (simp add: image_comp)
have "measure lebesgue (BOX2 (\<eta> t) t) * 3 ^ DIM('a)
= measure lebesgue (BOX (\<eta> t) t) * (2*3) ^ DIM('a)"
if "t \<in> tag ` \<G>" for t
proof -
have "content (cbox (t - \<eta> t *\<^sub>R One) (t + \<eta> t *\<^sub>R One))
= content (cbox t (t + \<eta> t *\<^sub>R One)) * 2 ^ DIM('a)"
using that by (simp add: algebra_simps content_cbox_if box_eq_empty)
then show ?thesis
by (simp add: BOX2_def BOX_def flip: power_mult_distrib)
qed
then have "measure lebesgue (\<Union>(\<Phi>2 ` \<G>)) * 3 ^ DIM('a) = measure lebesgue (\<Union>(\<Phi> ` \<G>)) * 6 ^ DIM('a)"
unfolding \<Phi>_def \<Phi>2_def eq
by (simp add: measure_negligible_finite_Union_image
\<open>finite \<G>\<close> BOX2_m BOX_m BOX2_disj BOX_disj sum_distrib_right
del: UN_simps)
also have "\<dots> \<le> e/2"
proof -
have "\<mu> * measure lebesgue (\<Union>D\<in>\<G>. \<Phi> D) \<le> \<mu> * (\<Sum>D \<in> \<Phi>`\<G>. measure lebesgue D)"
using \<open>\<mu> > 0\<close> \<open>finite \<G>\<close> by (force simp: BOX_m \<Phi>_def fmeasurableD intro: measure_Union_le)
also have "\<dots> = (\<Sum>D \<in> \<Phi>`\<G>. measure lebesgue D * \<mu>)"
by (metis mult.commute sum_distrib_right)
also have "\<dots> \<le> (\<Sum>(x, K) \<in> (\<lambda>D. (tag D, \<Phi> D)) ` \<G>. norm (content K *\<^sub>R f x - integral K f))"
proof (rule sum_le_included; clarify?)
fix D
assume "D \<in> \<G>"
then have "\<eta> (tag D) > 0"
using \<open>\<F> \<subseteq> \<D>\<close> \<open>\<G> \<subseteq> \<F>\<close> h0 tag_in_E by auto
then have m_\<Phi>: "measure lebesgue (\<Phi> D) > 0"
by (simp add: \<Phi>_def BOX_def algebra_simps)
have "\<mu> \<le> norm (i (\<eta>(tag D)) (tag D) - f(tag D))"
using \<mu>_le \<open>D \<in> \<G>\<close> \<open>\<F> \<subseteq> \<D>\<close> \<open>\<G> \<subseteq> \<F>\<close> tag_in_E by auto
also have "\<dots> = norm ((content (\<Phi> D) *\<^sub>R f(tag D) - integral (\<Phi> D) f) /\<^sub>R measure lebesgue (\<Phi> D))"
using m_\<Phi>
unfolding i_def \<Phi>_def BOX_def
by (simp add: algebra_simps content_cbox_plus norm_minus_commute)
finally have "measure lebesgue (\<Phi> D) * \<mu> \<le> norm (content (\<Phi> D) *\<^sub>R f(tag D) - integral (\<Phi> D) f)"
using m_\<Phi> by simp (simp add: field_simps)
then show "\<exists>y\<in>(\<lambda>D. (tag D, \<Phi> D)) ` \<G>.
snd y = \<Phi> D \<and> measure lebesgue (\<Phi> D) * \<mu> \<le> (case y of (x, k) \<Rightarrow> norm (content k *\<^sub>R f x - integral k f))"
using \<open>D \<in> \<G>\<close> by auto
qed (use \<open>finite \<G>\<close> in auto)
also have "\<dots> < ?ee"
proof (rule \<gamma>)
show "(\<lambda>D. (tag D, \<Phi> D)) ` \<G> tagged_partial_division_of cbox (a - One) (b + One)"
unfolding tagged_partial_division_of_def
proof (intro conjI allI impI ; clarify ?)
show "tag D \<in> \<Phi> D"
if "D \<in> \<G>" for D
using that \<open>\<F> \<subseteq> \<D>\<close> \<open>\<G> \<subseteq> \<F>\<close> h0 tag_in_E
by (auto simp: \<Phi>_def BOX_def mem_box algebra_simps eucl_less_le_not_le in_mono)
show "y \<in> cbox (a - One) (b + One)" if "D \<in> \<G>" "y \<in> \<Phi> D" for D y
using that BOX_cbox \<Phi>_def \<open>\<F> \<subseteq> \<D>\<close> \<open>\<G> \<subseteq> \<F>\<close> tag_in_E by blast
show "tag D = tag E \<and> \<Phi> D = \<Phi> E"
if "D \<in> \<G>" "E \<in> \<G>" and ne: "interior (\<Phi> D) \<inter> interior (\<Phi> E) \<noteq> {}" for D E
proof -
have "BOX2 (\<eta> (tag D)) (tag D) \<inter> BOX2 (\<eta> (tag E)) (tag E) = {} \<or> tag E = tag D"
using DISJ2 \<open>D \<in> \<G>\<close> \<open>E \<in> \<G>\<close> by force
then have "BOX (\<eta> (tag D)) (tag D) \<inter> BOX (\<eta> (tag E)) (tag E) = {} \<or> tag E = tag D"
using BOX_sub by blast
then show "tag D = tag E \<and> \<Phi> D = \<Phi> E"
by (metis \<Phi>_def interior_Int interior_empty ne)
qed
qed (use \<open>finite \<G>\<close> \<Phi>_def BOX_def in auto)
show "\<gamma> fine (\<lambda>D. (tag D, \<Phi> D)) ` \<G>"
unfolding fine_def \<Phi>_def using BOX_\<gamma> \<open>\<F> \<subseteq> \<D>\<close> \<open>\<G> \<subseteq> \<F>\<close> tag_in_E by blast
qed
finally show ?thesis
using \<open>\<mu> > 0\<close> by (auto simp: field_split_simps)
qed
finally show ?thesis .
qed
moreover
have "measure lebesgue (\<Union>\<F>) \<le> measure lebesgue (\<Union>(\<Phi>2`\<F>))"
proof (rule measure_mono_fmeasurable)
have "D \<subseteq> ball (tag D) (\<eta>(tag D))" if "D \<in> \<F>" for D
using \<open>\<F> \<subseteq> \<D>\<close> sub_ball_tag that by blast
moreover have "ball (tag D) (\<eta>(tag D)) \<subseteq> BOX2 (\<eta> (tag D)) (tag D)" if "D \<in> \<F>" for D
proof (clarsimp simp: \<Phi>2_def BOX2_def mem_box algebra_simps dist_norm)
fix x and i::'a
assume "norm (tag D - x) < \<eta> (tag D)" and "i \<in> Basis"
then have "\<bar>tag D \<bullet> i - x \<bullet> i\<bar> \<le> \<eta> (tag D)"
by (metis eucl_less_le_not_le inner_commute inner_diff_right norm_bound_Basis_le)
then show "tag D \<bullet> i \<le> x \<bullet> i + \<eta> (tag D) \<and> x \<bullet> i \<le> \<eta> (tag D) + tag D \<bullet> i"
by (simp add: abs_diff_le_iff)
qed
ultimately show "\<Union>\<F> \<subseteq> \<Union>(\<Phi>2`\<F>)"
by (force simp: \<Phi>2_def)
show "\<Union>\<F> \<in> sets lebesgue"
using \<open>finite \<F>\<close> \<open>\<D> \<subseteq> sets lebesgue\<close> \<open>\<F> \<subseteq> \<D>\<close> by blast
show "\<Union>(\<Phi>2`\<F>) \<in> lmeasurable"
unfolding \<Phi>2_def BOX2_def using \<open>finite \<F>\<close> by blast
qed
ultimately
have "measure lebesgue (\<Union>\<F>) \<le> e/2"
by (auto simp: field_split_simps)
then show "measure lebesgue (\<Union>\<D>) \<le> e"
using \<F> by linarith
qed
qed
qed
then have "\<And>j. negligible {x. \<Psi> x (inverse(real j + 1))}"
using negligible_on_intervals
by (metis (full_types) inverse_positive_iff_positive le_add_same_cancel1 linorder_not_le nat_le_real_less not_add_less1 of_nat_0)
then have "negligible ?M"
by auto
moreover have "?N \<subseteq> ?M"
proof (clarsimp simp: dist_norm)
fix y e
assume "0 < e"
and ye [rule_format]: "\<Psi> y e"
then obtain k where k: "0 < k" "inverse (real k + 1) < e"
by (metis One_nat_def add.commute less_add_same_cancel2 less_imp_inverse_less less_trans neq0_conv of_nat_1 of_nat_Suc reals_Archimedean zero_less_one)
with ye show "\<exists>n. \<Psi> y (inverse (real n + 1))"
apply (rule_tac x=k in exI)
unfolding \<Psi>_def
by (force intro: less_le_trans)
qed
ultimately show "negligible ?N"
by (blast intro: negligible_subset)
show "\<not> \<Psi> x e" if "x \<notin> ?N \<and> 0 < e" for x e
using that by blast
qed
with that show ?thesis
unfolding i_def BOX_def \<Psi>_def by (fastforce simp add: not_le)
qed
subsection\<open>HOL Light measurability\<close>
definition measurable_on :: "('a::euclidean_space \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a set \<Rightarrow> bool"
(infixr "measurable'_on" 46)
where "f measurable_on S \<equiv>
\<exists>N g. negligible N \<and>
(\<forall>n. continuous_on UNIV (g n)) \<and>
(\<forall>x. x \<notin> N \<longrightarrow> (\<lambda>n. g n x) \<longlonglongrightarrow> (if x \<in> S then f x else 0))"
lemma measurable_on_UNIV:
"(\<lambda>x. if x \<in> S then f x else 0) measurable_on UNIV \<longleftrightarrow> f measurable_on S"
by (auto simp: measurable_on_def)
lemma measurable_on_spike_set:
assumes f: "f measurable_on S" and neg: "negligible ((S - T) \<union> (T - S))"
shows "f measurable_on T"
proof -
obtain N and F
where N: "negligible N"
and conF: "\<And>n. continuous_on UNIV (F n)"
and tendsF: "\<And>x. x \<notin> N \<Longrightarrow> (\<lambda>n. F n x) \<longlonglongrightarrow> (if x \<in> S then f x else 0)"
using f by (auto simp: measurable_on_def)
show ?thesis
unfolding measurable_on_def
proof (intro exI conjI allI impI)
show "continuous_on UNIV (\<lambda>x. F n x)" for n
by (intro conF continuous_intros)
show "negligible (N \<union> (S - T) \<union> (T - S))"
by (metis (full_types) N neg negligible_Un_eq)
show "(\<lambda>n. F n x) \<longlonglongrightarrow> (if x \<in> T then f x else 0)"
if "x \<notin> (N \<union> (S - T) \<union> (T - S))" for x
using that tendsF [of x] by auto
qed
qed
text\<open> Various common equivalent forms of function measurability. \<close>
lemma measurable_on_0 [simp]: "(\<lambda>x. 0) measurable_on S"
unfolding measurable_on_def
proof (intro exI conjI allI impI)
show "(\<lambda>n. 0) \<longlonglongrightarrow> (if x \<in> S then 0::'b else 0)" for x
by force
qed auto
lemma measurable_on_scaleR_const:
assumes f: "f measurable_on S"
shows "(\<lambda>x. c *\<^sub>R f x) measurable_on S"
proof -
obtain NF and F
where NF: "negligible NF"
and conF: "\<And>n. continuous_on UNIV (F n)"
and tendsF: "\<And>x. x \<notin> NF \<Longrightarrow> (\<lambda>n. F n x) \<longlonglongrightarrow> (if x \<in> S then f x else 0)"
using f by (auto simp: measurable_on_def)
show ?thesis
unfolding measurable_on_def
proof (intro exI conjI allI impI)
show "continuous_on UNIV (\<lambda>x. c *\<^sub>R F n x)" for n
by (intro conF continuous_intros)
show "(\<lambda>n. c *\<^sub>R F n x) \<longlonglongrightarrow> (if x \<in> S then c *\<^sub>R f x else 0)"
if "x \<notin> NF" for x
using tendsto_scaleR [OF tendsto_const tendsF, of x] that by auto
qed (auto simp: NF)
qed
lemma measurable_on_cmul:
fixes c :: real
assumes "f measurable_on S"
shows "(\<lambda>x. c * f x) measurable_on S"
using measurable_on_scaleR_const [OF assms] by simp
lemma measurable_on_cdivide:
fixes c :: real
assumes "f measurable_on S"
shows "(\<lambda>x. f x / c) measurable_on S"
proof (cases "c=0")
case False
then show ?thesis
using measurable_on_cmul [of f S "1/c"]
by (simp add: assms)
qed auto
lemma measurable_on_minus:
"f measurable_on S \<Longrightarrow> (\<lambda>x. -(f x)) measurable_on S"
using measurable_on_scaleR_const [of f S "-1"] by auto
lemma continuous_imp_measurable_on:
"continuous_on UNIV f \<Longrightarrow> f measurable_on UNIV"
unfolding measurable_on_def
apply (rule_tac x="{}" in exI)
apply (rule_tac x="\<lambda>n. f" in exI, auto)
done
proposition integrable_subintervals_imp_measurable:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "\<And>a b. f integrable_on cbox a b"
shows "f measurable_on UNIV"
proof -
define BOX where "BOX \<equiv> \<lambda>h. \<lambda>x::'a. cbox x (x + h *\<^sub>R One)"
define i where "i \<equiv> \<lambda>h x. integral (BOX h x) f /\<^sub>R h ^ DIM('a)"
obtain N where "negligible N"
and k: "\<And>x e. \<lbrakk>x \<notin> N; 0 < e\<rbrakk>
\<Longrightarrow> \<exists>d>0. \<forall>h. 0 < h \<and> h < d \<longrightarrow>
norm (integral (cbox x (x + h *\<^sub>R One)) f /\<^sub>R h ^ DIM('a) - f x) < e"
using integrable_ccontinuous_explicit assms by blast
show ?thesis
unfolding measurable_on_def
proof (intro exI conjI allI impI)
show "continuous_on UNIV ((\<lambda>n x. i (inverse(Suc n)) x) n)" for n
proof (clarsimp simp: continuous_on_iff)
show "\<exists>d>0. \<forall>x'. dist x' x < d \<longrightarrow> dist (i (inverse (1 + real n)) x') (i (inverse (1 + real n)) x) < e"
if "0 < e"
for x e
proof -
let ?e = "e / (1 + real n) ^ DIM('a)"
have "?e > 0"
using \<open>e > 0\<close> by auto
moreover have "x \<in> cbox (x - 2 *\<^sub>R One) (x + 2 *\<^sub>R One)"
by (simp add: mem_box inner_diff_left inner_left_distrib)
moreover have "x + One /\<^sub>R real (Suc n) \<in> cbox (x - 2 *\<^sub>R One) (x + 2 *\<^sub>R One)"
by (auto simp: mem_box inner_diff_left inner_left_distrib field_simps)
ultimately obtain \<delta> where "\<delta> > 0"
and \<delta>: "\<And>c' d'. \<lbrakk>c' \<in> cbox (x - 2 *\<^sub>R One) (x + 2 *\<^sub>R One);
d' \<in> cbox (x - 2 *\<^sub>R One) (x + 2 *\<^sub>R One);
norm(c' - x) \<le> \<delta>; norm(d' - (x + One /\<^sub>R Suc n)) \<le> \<delta>\<rbrakk>
\<Longrightarrow> norm(integral(cbox c' d') f - integral(cbox x (x + One /\<^sub>R Suc n)) f) < ?e"
by (blast intro: indefinite_integral_continuous [of f _ _ x] assms)
show ?thesis
proof (intro exI impI conjI allI)
show "min \<delta> 1 > 0"
using \<open>\<delta> > 0\<close> by auto
show "dist (i (inverse (1 + real n)) y) (i (inverse (1 + real n)) x) < e"
if "dist y x < min \<delta> 1" for y
proof -
have no: "norm (y - x) < 1"
using that by (auto simp: dist_norm)
have le1: "inverse (1 + real n) \<le> 1"
by (auto simp: field_split_simps)
have "norm (integral (cbox y (y + One /\<^sub>R real (Suc n))) f
- integral (cbox x (x + One /\<^sub>R real (Suc n))) f)
< e / (1 + real n) ^ DIM('a)"
proof (rule \<delta>)
show "y \<in> cbox (x - 2 *\<^sub>R One) (x + 2 *\<^sub>R One)"
using no by (auto simp: mem_box algebra_simps dest: Basis_le_norm [of _ "y-x"])
show "y + One /\<^sub>R real (Suc n) \<in> cbox (x - 2 *\<^sub>R One) (x + 2 *\<^sub>R One)"
proof (simp add: dist_norm mem_box algebra_simps, intro ballI conjI)
fix i::'a
assume "i \<in> Basis"
then have 1: "\<bar>y \<bullet> i - x \<bullet> i\<bar> < 1"
by (metis inner_commute inner_diff_right no norm_bound_Basis_lt)
moreover have "\<dots> < (2 + inverse (1 + real n))" "1 \<le> 2 - inverse (1 + real n)"
by (auto simp: field_simps)
ultimately show "x \<bullet> i \<le> y \<bullet> i + (2 + inverse (1 + real n))"
"y \<bullet> i + inverse (1 + real n) \<le> x \<bullet> i + 2"
by linarith+
qed
show "norm (y - x) \<le> \<delta>" "norm (y + One /\<^sub>R real (Suc n) - (x + One /\<^sub>R real (Suc n))) \<le> \<delta>"
using that by (auto simp: dist_norm)
qed
then show ?thesis
using that by (simp add: dist_norm i_def BOX_def flip: scaleR_diff_right) (simp add: field_simps)
qed
qed
qed
qed
show "negligible N"
by (simp add: \<open>negligible N\<close>)
show "(\<lambda>n. i (inverse (Suc n)) x) \<longlonglongrightarrow> (if x \<in> UNIV then f x else 0)"
if "x \<notin> N" for x
unfolding lim_sequentially
proof clarsimp
show "\<exists>no. \<forall>n\<ge>no. dist (i (inverse (1 + real n)) x) (f x) < e"
if "0 < e" for e
proof -
obtain d where "d > 0"
and d: "\<And>h. \<lbrakk>0 < h; h < d\<rbrakk> \<Longrightarrow>
norm (integral (cbox x (x + h *\<^sub>R One)) f /\<^sub>R h ^ DIM('a) - f x) < e"
using k [of x e] \<open>x \<notin> N\<close> \<open>0 < e\<close> by blast
then obtain M where M: "M \<noteq> 0" "0 < inverse (real M)" "inverse (real M) < d"
using real_arch_invD by auto
show ?thesis
proof (intro exI allI impI)
show "dist (i (inverse (1 + real n)) x) (f x) < e"
if "M \<le> n" for n
proof -
have *: "0 < inverse (1 + real n)" "inverse (1 + real n) \<le> inverse M"
using that \<open>M \<noteq> 0\<close> by auto
show ?thesis
using that M
apply (simp add: i_def BOX_def dist_norm)
apply (blast intro: le_less_trans * d)
done
qed
qed
qed
qed
qed
qed
subsection\<open>Composing continuous and measurable functions; a few variants\<close>
lemma measurable_on_compose_continuous:
assumes f: "f measurable_on UNIV" and g: "continuous_on UNIV g"
shows "(g \<circ> f) measurable_on UNIV"
proof -
obtain N and F
where "negligible N"
and conF: "\<And>n. continuous_on UNIV (F n)"
and tendsF: "\<And>x. x \<notin> N \<Longrightarrow> (\<lambda>n. F n x) \<longlonglongrightarrow> f x"
using f by (auto simp: measurable_on_def)
show ?thesis
unfolding measurable_on_def
proof (intro exI conjI allI impI)
show "negligible N"
by fact
show "continuous_on UNIV (g \<circ> (F n))" for n
using conF continuous_on_compose continuous_on_subset g by blast
show "(\<lambda>n. (g \<circ> F n) x) \<longlonglongrightarrow> (if x \<in> UNIV then (g \<circ> f) x else 0)"
if "x \<notin> N" for x :: 'a
using that g tendsF by (auto simp: continuous_on_def intro: tendsto_compose)
qed
qed
lemma measurable_on_compose_continuous_0:
assumes f: "f measurable_on S" and g: "continuous_on UNIV g" and "g 0 = 0"
shows "(g \<circ> f) measurable_on S"
proof -
have f': "(\<lambda>x. if x \<in> S then f x else 0) measurable_on UNIV"
using f measurable_on_UNIV by blast
show ?thesis
using measurable_on_compose_continuous [OF f' g]
by (simp add: measurable_on_UNIV o_def if_distrib \<open>g 0 = 0\<close> cong: if_cong)
qed
lemma measurable_on_compose_continuous_box:
assumes fm: "f measurable_on UNIV" and fab: "\<And>x. f x \<in> box a b"
and contg: "continuous_on (box a b) g"
shows "(g \<circ> f) measurable_on UNIV"
proof -
have "\<exists>\<gamma>. (\<forall>n. continuous_on UNIV (\<gamma> n)) \<and> (\<forall>x. x \<notin> N \<longrightarrow> (\<lambda>n. \<gamma> n x) \<longlonglongrightarrow> g (f x))"
if "negligible N"
and conth [rule_format]: "\<forall>n. continuous_on UNIV (\<lambda>x. h n x)"
and tends [rule_format]: "\<forall>x. x \<notin> N \<longrightarrow> (\<lambda>n. h n x) \<longlonglongrightarrow> f x"
for N and h :: "nat \<Rightarrow> 'a \<Rightarrow> 'b"
proof -
define \<theta> where "\<theta> \<equiv> \<lambda>n x. (\<Sum>i\<in>Basis. (max (a\<bullet>i + (b\<bullet>i - a\<bullet>i) / real (n+2))
(min ((h n x)\<bullet>i)
(b\<bullet>i - (b\<bullet>i - a\<bullet>i) / real (n+2)))) *\<^sub>R i)"
have aibi: "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i < b \<bullet> i"
using box_ne_empty(2) fab by auto
then have *: "\<And>i n. i \<in> Basis \<Longrightarrow> a \<bullet> i + real n * (a \<bullet> i) < b \<bullet> i + real n * (b \<bullet> i)"
by (meson add_mono_thms_linordered_field(3) less_eq_real_def mult_left_mono of_nat_0_le_iff)
show ?thesis
proof (intro exI conjI allI impI)
show "continuous_on UNIV (g \<circ> (\<theta> n))" for n :: nat
unfolding \<theta>_def
apply (intro continuous_on_compose2 [OF contg] continuous_intros conth)
- apply (auto simp: aibi * mem_box less_max_iff_disj min_less_iff_disj algebra_simps field_split_simps)
+ apply (auto simp: aibi * mem_box less_max_iff_disj min_less_iff_disj field_split_simps)
done
show "(\<lambda>n. (g \<circ> \<theta> n) x) \<longlonglongrightarrow> g (f x)"
if "x \<notin> N" for x
unfolding o_def
proof (rule isCont_tendsto_compose [where g=g])
show "isCont g (f x)"
using contg fab continuous_on_eq_continuous_at by blast
have "(\<lambda>n. \<theta> n x) \<longlonglongrightarrow> (\<Sum>i\<in>Basis. max (a \<bullet> i) (min (f x \<bullet> i) (b \<bullet> i)) *\<^sub>R i)"
unfolding \<theta>_def
proof (intro tendsto_intros \<open>x \<notin> N\<close> tends)
fix i::'b
assume "i \<in> Basis"
have a: "(\<lambda>n. a \<bullet> i + (b \<bullet> i - a \<bullet> i) / real n) \<longlonglongrightarrow> a\<bullet>i + 0"
by (intro tendsto_add lim_const_over_n tendsto_const)
show "(\<lambda>n. a \<bullet> i + (b \<bullet> i - a \<bullet> i) / real (n + 2)) \<longlonglongrightarrow> a \<bullet> i"
using LIMSEQ_ignore_initial_segment [where k=2, OF a] by simp
have b: "(\<lambda>n. b\<bullet>i - (b \<bullet> i - a \<bullet> i) / (real n)) \<longlonglongrightarrow> b\<bullet>i - 0"
by (intro tendsto_diff lim_const_over_n tendsto_const)
show "(\<lambda>n. b \<bullet> i - (b \<bullet> i - a \<bullet> i) / real (n + 2)) \<longlonglongrightarrow> b \<bullet> i"
using LIMSEQ_ignore_initial_segment [where k=2, OF b] by simp
qed
also have "(\<Sum>i\<in>Basis. max (a \<bullet> i) (min (f x \<bullet> i) (b \<bullet> i)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (f x \<bullet> i) *\<^sub>R i)"
apply (rule sum.cong)
using fab
apply auto
apply (intro order_antisym)
apply (auto simp: mem_box)
using less_imp_le apply blast
by (metis (full_types) linear max_less_iff_conj min.bounded_iff not_le)
also have "\<dots> = f x"
using euclidean_representation by blast
finally show "(\<lambda>n. \<theta> n x) \<longlonglongrightarrow> f x" .
qed
qed
qed
then show ?thesis
using fm by (auto simp: measurable_on_def)
qed
lemma measurable_on_Pair:
assumes f: "f measurable_on S" and g: "g measurable_on S"
shows "(\<lambda>x. (f x, g x)) measurable_on S"
proof -
obtain NF and F
where NF: "negligible NF"
and conF: "\<And>n. continuous_on UNIV (F n)"
and tendsF: "\<And>x. x \<notin> NF \<Longrightarrow> (\<lambda>n. F n x) \<longlonglongrightarrow> (if x \<in> S then f x else 0)"
using f by (auto simp: measurable_on_def)
obtain NG and G
where NG: "negligible NG"
and conG: "\<And>n. continuous_on UNIV (G n)"
and tendsG: "\<And>x. x \<notin> NG \<Longrightarrow> (\<lambda>n. G n x) \<longlonglongrightarrow> (if x \<in> S then g x else 0)"
using g by (auto simp: measurable_on_def)
show ?thesis
unfolding measurable_on_def
proof (intro exI conjI allI impI)
show "negligible (NF \<union> NG)"
by (simp add: NF NG)
show "continuous_on UNIV (\<lambda>x. (F n x, G n x))" for n
using conF conG continuous_on_Pair by blast
show "(\<lambda>n. (F n x, G n x)) \<longlonglongrightarrow> (if x \<in> S then (f x, g x) else 0)"
if "x \<notin> NF \<union> NG" for x
using tendsto_Pair [OF tendsF tendsG, of x x] that unfolding zero_prod_def
by (simp add: split: if_split_asm)
qed
qed
lemma measurable_on_combine:
assumes f: "f measurable_on S" and g: "g measurable_on S"
and h: "continuous_on UNIV (\<lambda>x. h (fst x) (snd x))" and "h 0 0 = 0"
shows "(\<lambda>x. h (f x) (g x)) measurable_on S"
proof -
have *: "(\<lambda>x. h (f x) (g x)) = (\<lambda>x. h (fst x) (snd x)) \<circ> (\<lambda>x. (f x, g x))"
by auto
show ?thesis
unfolding * by (auto simp: measurable_on_compose_continuous_0 measurable_on_Pair assms)
qed
lemma measurable_on_add:
assumes f: "f measurable_on S" and g: "g measurable_on S"
shows "(\<lambda>x. f x + g x) measurable_on S"
by (intro continuous_intros measurable_on_combine [OF assms]) auto
lemma measurable_on_diff:
assumes f: "f measurable_on S" and g: "g measurable_on S"
shows "(\<lambda>x. f x - g x) measurable_on S"
by (intro continuous_intros measurable_on_combine [OF assms]) auto
lemma measurable_on_scaleR:
assumes f: "f measurable_on S" and g: "g measurable_on S"
shows "(\<lambda>x. f x *\<^sub>R g x) measurable_on S"
by (intro continuous_intros measurable_on_combine [OF assms]) auto
lemma measurable_on_sum:
assumes "finite I" "\<And>i. i \<in> I \<Longrightarrow> f i measurable_on S"
shows "(\<lambda>x. sum (\<lambda>i. f i x) I) measurable_on S"
using assms by (induction I) (auto simp: measurable_on_add)
lemma measurable_on_spike:
assumes f: "f measurable_on T" and "negligible S" and gf: "\<And>x. x \<in> T - S \<Longrightarrow> g x = f x"
shows "g measurable_on T"
proof -
obtain NF and F
where NF: "negligible NF"
and conF: "\<And>n. continuous_on UNIV (F n)"
and tendsF: "\<And>x. x \<notin> NF \<Longrightarrow> (\<lambda>n. F n x) \<longlonglongrightarrow> (if x \<in> T then f x else 0)"
using f by (auto simp: measurable_on_def)
show ?thesis
unfolding measurable_on_def
proof (intro exI conjI allI impI)
show "negligible (NF \<union> S)"
by (simp add: NF \<open>negligible S\<close>)
show "\<And>x. x \<notin> NF \<union> S \<Longrightarrow> (\<lambda>n. F n x) \<longlonglongrightarrow> (if x \<in> T then g x else 0)"
by (metis (full_types) Diff_iff Un_iff gf tendsF)
qed (auto simp: conF)
qed
proposition indicator_measurable_on:
assumes "S \<in> sets lebesgue"
shows "indicat_real S measurable_on UNIV"
proof -
{ fix n::nat
let ?\<epsilon> = "(1::real) / (2 * 2^n)"
have \<epsilon>: "?\<epsilon> > 0"
by auto
obtain T where "closed T" "T \<subseteq> S" "S-T \<in> lmeasurable" and ST: "emeasure lebesgue (S - T) < ?\<epsilon>"
by (meson \<epsilon> assms sets_lebesgue_inner_closed)
obtain U where "open U" "S \<subseteq> U" "(U - S) \<in> lmeasurable" and US: "emeasure lebesgue (U - S) < ?\<epsilon>"
by (meson \<epsilon> assms sets_lebesgue_outer_open)
have eq: "-T \<inter> U = (S-T) \<union> (U - S)"
using \<open>T \<subseteq> S\<close> \<open>S \<subseteq> U\<close> by auto
have "emeasure lebesgue ((S-T) \<union> (U - S)) \<le> emeasure lebesgue (S - T) + emeasure lebesgue (U - S)"
using \<open>S - T \<in> lmeasurable\<close> \<open>U - S \<in> lmeasurable\<close> emeasure_subadditive by blast
also have "\<dots> < ?\<epsilon> + ?\<epsilon>"
using ST US add_mono_ennreal by metis
finally have le: "emeasure lebesgue (-T \<inter> U) < ennreal (1 / 2^n)"
by (simp add: eq)
have 1: "continuous_on (T \<union> -U) (indicat_real S)"
unfolding indicator_def
proof (rule continuous_on_cases [OF \<open>closed T\<close>])
show "closed (- U)"
using \<open>open U\<close> by blast
show "continuous_on T (\<lambda>x. 1::real)" "continuous_on (- U) (\<lambda>x. 0::real)"
by (auto simp: continuous_on)
show "\<forall>x. x \<in> T \<and> x \<notin> S \<or> x \<in> - U \<and> x \<in> S \<longrightarrow> (1::real) = 0"
using \<open>T \<subseteq> S\<close> \<open>S \<subseteq> U\<close> by auto
qed
have 2: "closedin (top_of_set UNIV) (T \<union> -U)"
using \<open>closed T\<close> \<open>open U\<close> by auto
obtain g where "continuous_on UNIV g" "\<And>x. x \<in> T \<union> -U \<Longrightarrow> g x = indicat_real S x" "\<And>x. norm(g x) \<le> 1"
by (rule Tietze [OF 1 2, of 1]) auto
with le have "\<exists>g E. continuous_on UNIV g \<and> (\<forall>x \<in> -E. g x = indicat_real S x) \<and>
(\<forall>x. norm(g x) \<le> 1) \<and> E \<in> sets lebesgue \<and> emeasure lebesgue E < ennreal (1 / 2^n)"
apply (rule_tac x=g in exI)
apply (rule_tac x="-T \<inter> U" in exI)
using \<open>S - T \<in> lmeasurable\<close> \<open>U - S \<in> lmeasurable\<close> eq by auto
}
then obtain g E where cont: "\<And>n. continuous_on UNIV (g n)"
and geq: "\<And>n x. x \<in> - E n \<Longrightarrow> g n x = indicat_real S x"
and ng1: "\<And>n x. norm(g n x) \<le> 1"
and Eset: "\<And>n. E n \<in> sets lebesgue"
and Em: "\<And>n. emeasure lebesgue (E n) < ennreal (1 / 2^n)"
by metis
have null: "limsup E \<in> null_sets lebesgue"
proof (rule borel_cantelli_limsup1 [OF Eset])
show "emeasure lebesgue (E n) < \<infinity>" for n
by (metis Em infinity_ennreal_def order.asym top.not_eq_extremum)
show "summable (\<lambda>n. measure lebesgue (E n))"
proof (rule summable_comparison_test' [OF summable_geometric, of "1/2" 0])
show "norm (measure lebesgue (E n)) \<le> (1/2) ^ n" for n
using Em [of n] by (simp add: measure_def enn2real_leI power_one_over)
qed auto
qed
have tends: "(\<lambda>n. g n x) \<longlonglongrightarrow> indicat_real S x" if "x \<notin> limsup E" for x
proof -
have "\<forall>\<^sub>F n in sequentially. x \<in> - E n"
using that by (simp add: mem_limsup_iff not_frequently)
then show ?thesis
unfolding tendsto_iff dist_real_def
by (simp add: eventually_mono geq)
qed
show ?thesis
unfolding measurable_on_def
proof (intro exI conjI allI impI)
show "negligible (limsup E)"
using negligible_iff_null_sets null by blast
show "continuous_on UNIV (g n)" for n
using cont by blast
qed (use tends in auto)
qed
lemma measurable_on_restrict:
assumes f: "f measurable_on UNIV" and S: "S \<in> sets lebesgue"
shows "(\<lambda>x. if x \<in> S then f x else 0) measurable_on UNIV"
proof -
have "indicat_real S measurable_on UNIV"
by (simp add: S indicator_measurable_on)
then show ?thesis
using measurable_on_scaleR [OF _ f, of "indicat_real S"]
by (simp add: indicator_scaleR_eq_if)
qed
lemma measurable_on_const_UNIV: "(\<lambda>x. k) measurable_on UNIV"
by (simp add: continuous_imp_measurable_on)
lemma measurable_on_const [simp]: "S \<in> sets lebesgue \<Longrightarrow> (\<lambda>x. k) measurable_on S"
using measurable_on_UNIV measurable_on_const_UNIV measurable_on_restrict by blast
lemma simple_function_indicator_representation_real:
fixes f ::"'a \<Rightarrow> real"
assumes f: "simple_function M f" and x: "x \<in> space M" and nn: "\<And>x. f x \<ge> 0"
shows "f x = (\<Sum>y \<in> f ` space M. y * indicator (f -` {y} \<inter> space M) x)"
proof -
have f': "simple_function M (ennreal \<circ> f)"
by (simp add: f)
have *: "f x =
enn2real
(\<Sum>y\<in> ennreal ` f ` space M.
y * indicator ((ennreal \<circ> f) -` {y} \<inter> space M) x)"
using arg_cong [OF simple_function_indicator_representation [OF f' x], of enn2real, simplified nn o_def] nn
unfolding o_def image_comp
by (metis enn2real_ennreal)
have "enn2real (\<Sum>y\<in>ennreal ` f ` space M. if ennreal (f x) = y \<and> x \<in> space M then y else 0)
= sum (enn2real \<circ> (\<lambda>y. if ennreal (f x) = y \<and> x \<in> space M then y else 0))
(ennreal ` f ` space M)"
by (rule enn2real_sum) auto
also have "\<dots> = sum (enn2real \<circ> (\<lambda>y. if ennreal (f x) = y \<and> x \<in> space M then y else 0) \<circ> ennreal)
(f ` space M)"
by (rule sum.reindex) (use nn in \<open>auto simp: inj_on_def intro: sum.cong\<close>)
also have "\<dots> = (\<Sum>y\<in>f ` space M. if f x = y \<and> x \<in> space M then y else 0)"
using nn
by (auto simp: inj_on_def intro: sum.cong)
finally show ?thesis
by (subst *) (simp add: enn2real_sum indicator_def if_distrib cong: if_cong)
qed
lemma\<^marker>\<open>tag important\<close> simple_function_induct_real
[consumes 1, case_names cong set mult add, induct set: simple_function]:
fixes u :: "'a \<Rightarrow> real"
assumes u: "simple_function M u"
assumes cong: "\<And>f g. simple_function M f \<Longrightarrow> simple_function M g \<Longrightarrow> (AE x in M. f x = g x) \<Longrightarrow> P f \<Longrightarrow> P g"
assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)"
assumes mult: "\<And>u c. P u \<Longrightarrow> P (\<lambda>x. c * u x)"
assumes add: "\<And>u v. P u \<Longrightarrow> P v \<Longrightarrow> P (\<lambda>x. u x + v x)"
and nn: "\<And>x. u x \<ge> 0"
shows "P u"
proof (rule cong)
from AE_space show "AE x in M. (\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x) = u x"
proof eventually_elim
fix x assume x: "x \<in> space M"
from simple_function_indicator_representation_real[OF u x] nn
show "(\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x) = u x"
by metis
qed
next
from u have "finite (u ` space M)"
unfolding simple_function_def by auto
then show "P (\<lambda>x. \<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x)"
proof induct
case empty
then show ?case
using set[of "{}"] by (simp add: indicator_def[abs_def])
next
case (insert a F)
have eq: "\<Sum> {y. u x = y \<and> (y = a \<or> y \<in> F) \<and> x \<in> space M}
= (if u x = a \<and> x \<in> space M then a else 0) + \<Sum> {y. u x = y \<and> y \<in> F \<and> x \<in> space M}" for x
proof (cases "x \<in> space M")
case True
have *: "{y. u x = y \<and> (y = a \<or> y \<in> F)} = {y. u x = a \<and> y = a} \<union> {y. u x = y \<and> y \<in> F}"
by auto
show ?thesis
using insert by (simp add: * True)
qed auto
have a: "P (\<lambda>x. a * indicator (u -` {a} \<inter> space M) x)"
proof (intro mult set)
show "u -` {a} \<inter> space M \<in> sets M"
using u by auto
qed
show ?case
using nn insert a
by (simp add: eq indicator_times_eq_if [where f = "\<lambda>x. a"] add)
qed
next
show "simple_function M (\<lambda>x. (\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x))"
apply (subst simple_function_cong)
apply (rule simple_function_indicator_representation_real[symmetric])
apply (auto intro: u nn)
done
qed fact
proposition simple_function_measurable_on_UNIV:
fixes f :: "'a::euclidean_space \<Rightarrow> real"
assumes f: "simple_function lebesgue f" and nn: "\<And>x. f x \<ge> 0"
shows "f measurable_on UNIV"
using f
proof (induction f)
case (cong f g)
then obtain N where "negligible N" "{x. g x \<noteq> f x} \<subseteq> N"
by (auto simp: eventually_ae_filter_negligible eq_commute)
then show ?case
by (blast intro: measurable_on_spike cong)
next
case (set S)
then show ?case
by (simp add: indicator_measurable_on)
next
case (mult u c)
then show ?case
by (simp add: measurable_on_cmul)
case (add u v)
then show ?case
by (simp add: measurable_on_add)
qed (auto simp: nn)
lemma simple_function_lebesgue_if:
fixes f :: "'a::euclidean_space \<Rightarrow> real"
assumes f: "simple_function lebesgue f" and S: "S \<in> sets lebesgue"
shows "simple_function lebesgue (\<lambda>x. if x \<in> S then f x else 0)"
proof -
have ffin: "finite (range f)" and fsets: "\<forall>x. f -` {f x} \<in> sets lebesgue"
using f by (auto simp: simple_function_def)
have "finite (f ` S)"
by (meson finite_subset subset_image_iff ffin top_greatest)
moreover have "finite ((\<lambda>x. 0::real) ` T)" for T :: "'a set"
by (auto simp: image_def)
moreover have if_sets: "(\<lambda>x. if x \<in> S then f x else 0) -` {f a} \<in> sets lebesgue" for a
proof -
have *: "(\<lambda>x. if x \<in> S then f x else 0) -` {f a}
= (if f a = 0 then -S \<union> f -` {f a} else (f -` {f a}) \<inter> S)"
by (auto simp: split: if_split_asm)
show ?thesis
unfolding * by (metis Compl_in_sets_lebesgue S sets.Int sets.Un fsets)
qed
moreover have "(\<lambda>x. if x \<in> S then f x else 0) -` {0} \<in> sets lebesgue"
proof (cases "0 \<in> range f")
case True
then show ?thesis
by (metis (no_types, lifting) if_sets rangeE)
next
case False
then have "(\<lambda>x. if x \<in> S then f x else 0) -` {0} = -S"
by auto
then show ?thesis
by (simp add: Compl_in_sets_lebesgue S)
qed
ultimately show ?thesis
by (auto simp: simple_function_def)
qed
corollary simple_function_measurable_on:
fixes f :: "'a::euclidean_space \<Rightarrow> real"
assumes f: "simple_function lebesgue f" and nn: "\<And>x. f x \<ge> 0" and S: "S \<in> sets lebesgue"
shows "f measurable_on S"
by (simp add: measurable_on_UNIV [symmetric, of f] S f simple_function_lebesgue_if nn simple_function_measurable_on_UNIV)
lemma
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::ordered_euclidean_space"
assumes f: "f measurable_on S" and g: "g measurable_on S"
shows measurable_on_sup: "(\<lambda>x. sup (f x) (g x)) measurable_on S"
and measurable_on_inf: "(\<lambda>x. inf (f x) (g x)) measurable_on S"
proof -
obtain NF and F
where NF: "negligible NF"
and conF: "\<And>n. continuous_on UNIV (F n)"
and tendsF: "\<And>x. x \<notin> NF \<Longrightarrow> (\<lambda>n. F n x) \<longlonglongrightarrow> (if x \<in> S then f x else 0)"
using f by (auto simp: measurable_on_def)
obtain NG and G
where NG: "negligible NG"
and conG: "\<And>n. continuous_on UNIV (G n)"
and tendsG: "\<And>x. x \<notin> NG \<Longrightarrow> (\<lambda>n. G n x) \<longlonglongrightarrow> (if x \<in> S then g x else 0)"
using g by (auto simp: measurable_on_def)
show "(\<lambda>x. sup (f x) (g x)) measurable_on S"
unfolding measurable_on_def
proof (intro exI conjI allI impI)
show "continuous_on UNIV (\<lambda>x. sup (F n x) (G n x))" for n
unfolding sup_max eucl_sup by (intro conF conG continuous_intros)
show "(\<lambda>n. sup (F n x) (G n x)) \<longlonglongrightarrow> (if x \<in> S then sup (f x) (g x) else 0)"
if "x \<notin> NF \<union> NG" for x
using tendsto_sup [OF tendsF tendsG, of x x] that by auto
qed (simp add: NF NG)
show "(\<lambda>x. inf (f x) (g x)) measurable_on S"
unfolding measurable_on_def
proof (intro exI conjI allI impI)
show "continuous_on UNIV (\<lambda>x. inf (F n x) (G n x))" for n
unfolding inf_min eucl_inf by (intro conF conG continuous_intros)
show "(\<lambda>n. inf (F n x) (G n x)) \<longlonglongrightarrow> (if x \<in> S then inf (f x) (g x) else 0)"
if "x \<notin> NF \<union> NG" for x
using tendsto_inf [OF tendsF tendsG, of x x] that by auto
qed (simp add: NF NG)
qed
proposition measurable_on_componentwise_UNIV:
"f measurable_on UNIV \<longleftrightarrow> (\<forall>i\<in>Basis. (\<lambda>x. (f x \<bullet> i) *\<^sub>R i) measurable_on UNIV)"
(is "?lhs = ?rhs")
proof
assume L: ?lhs
show ?rhs
proof
fix i::'b
assume "i \<in> Basis"
have cont: "continuous_on UNIV (\<lambda>x. (x \<bullet> i) *\<^sub>R i)"
by (intro continuous_intros)
show "(\<lambda>x. (f x \<bullet> i) *\<^sub>R i) measurable_on UNIV"
using measurable_on_compose_continuous [OF L cont]
by (simp add: o_def)
qed
next
assume ?rhs
then have "\<exists>N g. negligible N \<and>
(\<forall>n. continuous_on UNIV (g n)) \<and>
(\<forall>x. x \<notin> N \<longrightarrow> (\<lambda>n. g n x) \<longlonglongrightarrow> (f x \<bullet> i) *\<^sub>R i)"
if "i \<in> Basis" for i
by (simp add: measurable_on_def that)
then obtain N g where N: "\<And>i. i \<in> Basis \<Longrightarrow> negligible (N i)"
and cont: "\<And>i n. i \<in> Basis \<Longrightarrow> continuous_on UNIV (g i n)"
and tends: "\<And>i x. \<lbrakk>i \<in> Basis; x \<notin> N i\<rbrakk> \<Longrightarrow> (\<lambda>n. g i n x) \<longlonglongrightarrow> (f x \<bullet> i) *\<^sub>R i"
by metis
show ?lhs
unfolding measurable_on_def
proof (intro exI conjI allI impI)
show "negligible (\<Union>i \<in> Basis. N i)"
using N eucl.finite_Basis by blast
show "continuous_on UNIV (\<lambda>x. (\<Sum>i\<in>Basis. g i n x))" for n
by (intro continuous_intros cont)
next
fix x
assume "x \<notin> (\<Union>i \<in> Basis. N i)"
then have "\<And>i. i \<in> Basis \<Longrightarrow> x \<notin> N i"
by auto
then have "(\<lambda>n. (\<Sum>i\<in>Basis. g i n x)) \<longlonglongrightarrow> (\<Sum>i\<in>Basis. (f x \<bullet> i) *\<^sub>R i)"
by (intro tends tendsto_intros)
then show "(\<lambda>n. (\<Sum>i\<in>Basis. g i n x)) \<longlonglongrightarrow> (if x \<in> UNIV then f x else 0)"
by (simp add: euclidean_representation)
qed
qed
corollary measurable_on_componentwise:
"f measurable_on S \<longleftrightarrow> (\<forall>i\<in>Basis. (\<lambda>x. (f x \<bullet> i) *\<^sub>R i) measurable_on S)"
apply (subst measurable_on_UNIV [symmetric])
apply (subst measurable_on_componentwise_UNIV)
apply (simp add: measurable_on_UNIV if_distrib [of "\<lambda>x. inner x _"] if_distrib [of "\<lambda>x. scaleR x _"] cong: if_cong)
done
(*FIXME: avoid duplication of proofs WRT borel_measurable_implies_simple_function_sequence*)
lemma\<^marker>\<open>tag important\<close> borel_measurable_implies_simple_function_sequence_real:
fixes u :: "'a \<Rightarrow> real"
assumes u[measurable]: "u \<in> borel_measurable M" and nn: "\<And>x. u x \<ge> 0"
shows "\<exists>f. incseq f \<and> (\<forall>i. simple_function M (f i)) \<and> (\<forall>x. bdd_above (range (\<lambda>i. f i x))) \<and>
(\<forall>i x. 0 \<le> f i x) \<and> u = (SUP i. f i)"
proof -
define f where [abs_def]:
"f i x = real_of_int (floor ((min i (u x)) * 2^i)) / 2^i" for i x
have [simp]: "0 \<le> f i x" for i x
by (auto simp: f_def intro!: divide_nonneg_nonneg mult_nonneg_nonneg nn)
have *: "2^n * real_of_int x = real_of_int (2^n * x)" for n x
by simp
have "real_of_int \<lfloor>real i * 2 ^ i\<rfloor> = real_of_int \<lfloor>i * 2 ^ i\<rfloor>" for i
by (intro arg_cong[where f=real_of_int]) simp
then have [simp]: "real_of_int \<lfloor>real i * 2 ^ i\<rfloor> = i * 2 ^ i" for i
unfolding floor_of_nat by simp
have bdd: "bdd_above (range (\<lambda>i. f i x))" for x
by (rule bdd_aboveI [where M = "u x"]) (auto simp: f_def field_simps min_def)
have "incseq f"
proof (intro monoI le_funI)
fix m n :: nat and x assume "m \<le> n"
moreover
{ fix d :: nat
have "\<lfloor>2^d::real\<rfloor> * \<lfloor>2^m * (min (of_nat m) (u x))\<rfloor> \<le> \<lfloor>2^d * (2^m * (min (of_nat m) (u x)))\<rfloor>"
by (rule le_mult_floor) (auto simp: nn)
also have "\<dots> \<le> \<lfloor>2^d * (2^m * (min (of_nat d + of_nat m) (u x)))\<rfloor>"
by (intro floor_mono mult_mono min.mono)
(auto simp: nn min_less_iff_disj of_nat_less_top)
finally have "f m x \<le> f(m + d) x"
unfolding f_def
by (auto simp: field_simps power_add * simp del: of_int_mult) }
ultimately show "f m x \<le> f n x"
by (auto simp: le_iff_add)
qed
then have inc_f: "incseq (\<lambda>i. f i x)" for x
by (auto simp: incseq_def le_fun_def)
moreover
have "simple_function M (f i)" for i
proof (rule simple_function_borel_measurable)
have "\<lfloor>(min (of_nat i) (u x)) * 2 ^ i\<rfloor> \<le> \<lfloor>int i * 2 ^ i\<rfloor>" for x
by (auto split: split_min intro!: floor_mono)
then have "f i ` space M \<subseteq> (\<lambda>n. real_of_int n / 2^i) ` {0 .. of_nat i * 2^i}"
unfolding floor_of_int by (auto simp: f_def nn intro!: imageI)
then show "finite (f i ` space M)"
by (rule finite_subset) auto
show "f i \<in> borel_measurable M"
unfolding f_def enn2real_def by measurable
qed
moreover
{ fix x
have "(SUP i. (f i x)) = u x"
proof -
obtain n where "u x \<le> of_nat n" using real_arch_simple by auto
then have min_eq_r: "\<forall>\<^sub>F i in sequentially. min (real i) (u x) = u x"
by (auto simp: eventually_sequentially intro!: exI[of _ n] split: split_min)
have "(\<lambda>i. real_of_int \<lfloor>min (real i) (u x) * 2^i\<rfloor> / 2^i) \<longlonglongrightarrow> u x"
proof (rule tendsto_sandwich)
show "(\<lambda>n. u x - (1/2)^n) \<longlonglongrightarrow> u x"
by (auto intro!: tendsto_eq_intros LIMSEQ_power_zero)
show "\<forall>\<^sub>F n in sequentially. real_of_int \<lfloor>min (real n) (u x) * 2 ^ n\<rfloor> / 2 ^ n \<le> u x"
using min_eq_r by eventually_elim (auto simp: field_simps)
have *: "u x * (2 ^ n * 2 ^ n) \<le> 2^n + 2^n * real_of_int \<lfloor>u x * 2 ^ n\<rfloor>" for n
using real_of_int_floor_ge_diff_one[of "u x * 2^n", THEN mult_left_mono, of "2^n"]
by (auto simp: field_simps)
show "\<forall>\<^sub>F n in sequentially. u x - (1/2)^n \<le> real_of_int \<lfloor>min (real n) (u x) * 2 ^ n\<rfloor> / 2 ^ n"
using min_eq_r by eventually_elim (insert *, auto simp: field_simps)
qed auto
then have "(\<lambda>i. (f i x)) \<longlonglongrightarrow> u x"
by (simp add: f_def)
from LIMSEQ_unique LIMSEQ_incseq_SUP [OF bdd inc_f] this
show ?thesis
by blast
qed }
ultimately show ?thesis
by (intro exI [of _ "\<lambda>i x. f i x"]) (auto simp: \<open>incseq f\<close> bdd image_comp)
qed
lemma homeomorphic_open_interval_UNIV:
fixes a b:: real
assumes "a < b"
shows "{a<..<b} homeomorphic (UNIV::real set)"
proof -
have "{a<..<b} = ball ((b+a) / 2) ((b-a) / 2)"
using assms
by (auto simp: dist_real_def abs_if field_split_simps split: if_split_asm)
then show ?thesis
by (simp add: homeomorphic_ball_UNIV assms)
qed
proposition homeomorphic_box_UNIV:
fixes a b:: "'a::euclidean_space"
assumes "box a b \<noteq> {}"
shows "box a b homeomorphic (UNIV::'a set)"
proof -
have "{a \<bullet> i <..<b \<bullet> i} homeomorphic (UNIV::real set)" if "i \<in> Basis" for i
using assms box_ne_empty that by (blast intro: homeomorphic_open_interval_UNIV)
then have "\<exists>f g. (\<forall>x. a \<bullet> i < x \<and> x < b \<bullet> i \<longrightarrow> g (f x) = x) \<and>
(\<forall>y. a \<bullet> i < g y \<and> g y < b \<bullet> i \<and> f(g y) = y) \<and>
continuous_on {a \<bullet> i<..<b \<bullet> i} f \<and>
continuous_on (UNIV::real set) g"
if "i \<in> Basis" for i
using that by (auto simp: homeomorphic_minimal mem_box Ball_def)
then obtain f g where gf: "\<And>i x. \<lbrakk>i \<in> Basis; a \<bullet> i < x; x < b \<bullet> i\<rbrakk> \<Longrightarrow> g i (f i x) = x"
and fg: "\<And>i y. i \<in> Basis \<Longrightarrow> a \<bullet> i < g i y \<and> g i y < b \<bullet> i \<and> f i (g i y) = y"
and contf: "\<And>i. i \<in> Basis \<Longrightarrow> continuous_on {a \<bullet> i<..<b \<bullet> i} (f i)"
and contg: "\<And>i. i \<in> Basis \<Longrightarrow> continuous_on (UNIV::real set) (g i)"
by metis
define F where "F \<equiv> \<lambda>x. \<Sum>i\<in>Basis. (f i (x \<bullet> i)) *\<^sub>R i"
define G where "G \<equiv> \<lambda>x. \<Sum>i\<in>Basis. (g i (x \<bullet> i)) *\<^sub>R i"
show ?thesis
unfolding homeomorphic_minimal
proof (intro exI conjI ballI)
show "G y \<in> box a b" for y
using fg by (simp add: G_def mem_box)
show "G (F x) = x" if "x \<in> box a b" for x
using that by (simp add: F_def G_def gf mem_box euclidean_representation)
show "F (G y) = y" for y
by (simp add: F_def G_def fg mem_box euclidean_representation)
show "continuous_on (box a b) F"
unfolding F_def
proof (intro continuous_intros continuous_on_compose2 [OF contf continuous_on_inner])
show "(\<lambda>x. x \<bullet> i) ` box a b \<subseteq> {a \<bullet> i<..<b \<bullet> i}" if "i \<in> Basis" for i
using that by (auto simp: mem_box)
qed
show "continuous_on UNIV G"
unfolding G_def
by (intro continuous_intros continuous_on_compose2 [OF contg continuous_on_inner]) auto
qed auto
qed
lemma diff_null_sets_lebesgue: "\<lbrakk>N \<in> null_sets (lebesgue_on S); X-N \<in> sets (lebesgue_on S); N \<subseteq> X\<rbrakk>
\<Longrightarrow> X \<in> sets (lebesgue_on S)"
by (metis Int_Diff_Un inf.commute inf.orderE null_setsD2 sets.Un)
lemma borel_measurable_diff_null:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes N: "N \<in> null_sets (lebesgue_on S)" and S: "S \<in> sets lebesgue"
shows "f \<in> borel_measurable (lebesgue_on (S-N)) \<longleftrightarrow> f \<in> borel_measurable (lebesgue_on S)"
unfolding in_borel_measurable space_lebesgue_on sets_restrict_UNIV
proof (intro ball_cong iffI)
show "f -` T \<inter> S \<in> sets (lebesgue_on S)"
if "f -` T \<inter> (S-N) \<in> sets (lebesgue_on (S-N))" for T
proof -
have "N \<inter> S = N"
by (metis N S inf.orderE null_sets_restrict_space)
moreover have "N \<inter> S \<in> sets lebesgue"
by (metis N S inf.orderE null_setsD2 null_sets_restrict_space)
moreover have "f -` T \<inter> S \<inter> (f -` T \<inter> N) \<in> sets lebesgue"
by (metis N S completion.complete inf.absorb2 inf_le2 inf_mono null_sets_restrict_space)
ultimately show ?thesis
by (metis Diff_Int_distrib Int_Diff_Un S inf_le2 sets.Diff sets.Un sets_restrict_space_iff space_lebesgue_on space_restrict_space that)
qed
show "f -` T \<inter> (S-N) \<in> sets (lebesgue_on (S-N))"
if "f -` T \<inter> S \<in> sets (lebesgue_on S)" for T
proof -
have "(S - N) \<inter> f -` T = (S - N) \<inter> (f -` T \<inter> S)"
by blast
then have "(S - N) \<inter> f -` T \<in> sets.restricted_space lebesgue (S - N)"
by (metis S image_iff sets.Int_space_eq2 sets_restrict_space_iff that)
then show ?thesis
by (simp add: inf.commute sets_restrict_space)
qed
qed auto
lemma lebesgue_measurable_diff_null:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "N \<in> null_sets lebesgue"
shows "f \<in> borel_measurable (lebesgue_on (-N)) \<longleftrightarrow> f \<in> borel_measurable lebesgue"
by (simp add: Compl_eq_Diff_UNIV assms borel_measurable_diff_null lebesgue_on_UNIV_eq)
proposition measurable_on_imp_borel_measurable_lebesgue_UNIV:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "f measurable_on UNIV"
shows "f \<in> borel_measurable lebesgue"
proof -
obtain N and F
where NF: "negligible N"
and conF: "\<And>n. continuous_on UNIV (F n)"
and tendsF: "\<And>x. x \<notin> N \<Longrightarrow> (\<lambda>n. F n x) \<longlonglongrightarrow> f x"
using assms by (auto simp: measurable_on_def)
obtain N where "N \<in> null_sets lebesgue" "f \<in> borel_measurable (lebesgue_on (-N))"
proof
show "f \<in> borel_measurable (lebesgue_on (- N))"
proof (rule borel_measurable_LIMSEQ_metric)
show "F i \<in> borel_measurable (lebesgue_on (- N))" for i
by (meson Compl_in_sets_lebesgue NF conF continuous_imp_measurable_on_sets_lebesgue continuous_on_subset negligible_imp_sets subset_UNIV)
show "(\<lambda>i. F i x) \<longlonglongrightarrow> f x" if "x \<in> space (lebesgue_on (- N))" for x
using that
by (simp add: tendsF)
qed
show "N \<in> null_sets lebesgue"
using NF negligible_iff_null_sets by blast
qed
then show ?thesis
using lebesgue_measurable_diff_null by blast
qed
corollary measurable_on_imp_borel_measurable_lebesgue:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "f measurable_on S" and S: "S \<in> sets lebesgue"
shows "f \<in> borel_measurable (lebesgue_on S)"
proof -
have "(\<lambda>x. if x \<in> S then f x else 0) measurable_on UNIV"
using assms(1) measurable_on_UNIV by blast
then show ?thesis
by (simp add: borel_measurable_if_D measurable_on_imp_borel_measurable_lebesgue_UNIV)
qed
proposition measurable_on_limit:
fixes f :: "nat \<Rightarrow> 'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes f: "\<And>n. f n measurable_on S" and N: "negligible N"
and lim: "\<And>x. x \<in> S - N \<Longrightarrow> (\<lambda>n. f n x) \<longlonglongrightarrow> g x"
shows "g measurable_on S"
proof -
have "box (0::'b) One homeomorphic (UNIV::'b set)"
by (simp add: homeomorphic_box_UNIV)
then obtain h h':: "'b\<Rightarrow>'b" where hh': "\<And>x. x \<in> box 0 One \<Longrightarrow> h (h' x) = x"
and h'im: "h' ` box 0 One = UNIV"
and conth: "continuous_on UNIV h"
and conth': "continuous_on (box 0 One) h'"
and h'h: "\<And>y. h' (h y) = y"
and rangeh: "range h = box 0 One"
by (auto simp: homeomorphic_def homeomorphism_def)
have "norm y \<le> DIM('b)" if y: "y \<in> box 0 One" for y::'b
proof -
have y01: "0 < y \<bullet> i" "y \<bullet> i < 1" if "i \<in> Basis" for i
using that y by (auto simp: mem_box)
have "norm y \<le> (\<Sum>i\<in>Basis. \<bar>y \<bullet> i\<bar>)"
using norm_le_l1 by blast
also have "\<dots> \<le> (\<Sum>i::'b\<in>Basis. 1)"
proof (rule sum_mono)
show "\<bar>y \<bullet> i\<bar> \<le> 1" if "i \<in> Basis" for i
using y01 that by fastforce
qed
also have "\<dots> \<le> DIM('b)"
by auto
finally show ?thesis .
qed
then have norm_le: "norm(h y) \<le> DIM('b)" for y
by (metis UNIV_I image_eqI rangeh)
have "(h' \<circ> (h \<circ> (\<lambda>x. if x \<in> S then g x else 0))) measurable_on UNIV"
proof (rule measurable_on_compose_continuous_box)
let ?\<chi> = "h \<circ> (\<lambda>x. if x \<in> S then g x else 0)"
let ?f = "\<lambda>n. h \<circ> (\<lambda>x. if x \<in> S then f n x else 0)"
show "?\<chi> measurable_on UNIV"
proof (rule integrable_subintervals_imp_measurable)
show "?\<chi> integrable_on cbox a b" for a b
proof (rule integrable_spike_set)
show "?\<chi> integrable_on (cbox a b - N)"
proof (rule dominated_convergence_integrable)
show const: "(\<lambda>x. DIM('b)) integrable_on cbox a b - N"
by (simp add: N has_integral_iff integrable_const integrable_negligible integrable_setdiff negligible_diff)
show "norm ((h \<circ> (\<lambda>x. if x \<in> S then g x else 0)) x) \<le> DIM('b)" if "x \<in> cbox a b - N" for x
using that norm_le by (simp add: o_def)
show "(\<lambda>k. ?f k x) \<longlonglongrightarrow> ?\<chi> x" if "x \<in> cbox a b - N" for x
using that lim [of x] conth
by (auto simp: continuous_on_def intro: tendsto_compose)
show "(?f n) absolutely_integrable_on cbox a b - N" for n
proof (rule measurable_bounded_by_integrable_imp_absolutely_integrable)
show "?f n \<in> borel_measurable (lebesgue_on (cbox a b - N))"
proof (rule measurable_on_imp_borel_measurable_lebesgue [OF measurable_on_spike_set])
show "?f n measurable_on cbox a b"
unfolding measurable_on_UNIV [symmetric, of _ "cbox a b"]
proof (rule measurable_on_restrict)
have f': "(\<lambda>x. if x \<in> S then f n x else 0) measurable_on UNIV"
by (simp add: f measurable_on_UNIV)
show "?f n measurable_on UNIV"
using measurable_on_compose_continuous [OF f' conth] by auto
qed auto
show "negligible (sym_diff (cbox a b) (cbox a b - N))"
by (auto intro: negligible_subset [OF N])
show "cbox a b - N \<in> sets lebesgue"
by (simp add: N negligible_imp_sets sets.Diff)
qed
show "cbox a b - N \<in> sets lebesgue"
by (simp add: N negligible_imp_sets sets.Diff)
show "norm (?f n x) \<le> DIM('b)"
if "x \<in> cbox a b - N" for x
using that local.norm_le by simp
qed (auto simp: const)
qed
show "negligible {x \<in> cbox a b - N - cbox a b. ?\<chi> x \<noteq> 0}"
by (auto simp: empty_imp_negligible)
have "{x \<in> cbox a b - (cbox a b - N). ?\<chi> x \<noteq> 0} \<subseteq> N"
by auto
then show "negligible {x \<in> cbox a b - (cbox a b - N). ?\<chi> x \<noteq> 0}"
using N negligible_subset by blast
qed
qed
show "?\<chi> x \<in> box 0 One" for x
using rangeh by auto
show "continuous_on (box 0 One) h'"
by (rule conth')
qed
then show ?thesis
by (simp add: o_def h'h measurable_on_UNIV)
qed
lemma measurable_on_if_simple_function_limit:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
shows "\<lbrakk>\<And>n. g n measurable_on UNIV; \<And>n. finite (range (g n)); \<And>x. (\<lambda>n. g n x) \<longlonglongrightarrow> f x\<rbrakk>
\<Longrightarrow> f measurable_on UNIV"
by (force intro: measurable_on_limit [where N="{}"])
lemma lebesgue_measurable_imp_measurable_on_nnreal_UNIV:
fixes u :: "'a::euclidean_space \<Rightarrow> real"
assumes u: "u \<in> borel_measurable lebesgue" and nn: "\<And>x. u x \<ge> 0"
shows "u measurable_on UNIV"
proof -
obtain f where "incseq f" and f: "\<forall>i. simple_function lebesgue (f i)"
and bdd: "\<And>x. bdd_above (range (\<lambda>i. f i x))"
and nnf: "\<And>i x. 0 \<le> f i x" and *: "u = (SUP i. f i)"
using borel_measurable_implies_simple_function_sequence_real nn u by metis
show ?thesis
unfolding *
proof (rule measurable_on_if_simple_function_limit [of concl: "Sup (range f)"])
show "(f i) measurable_on UNIV" for i
by (simp add: f nnf simple_function_measurable_on_UNIV)
show "finite (range (f i))" for i
by (metis f simple_function_def space_borel space_completion space_lborel)
show "(\<lambda>i. f i x) \<longlonglongrightarrow> Sup (range f) x" for x
proof -
have "incseq (\<lambda>i. f i x)"
using \<open>incseq f\<close> apply (auto simp: incseq_def)
by (simp add: le_funD)
then show ?thesis
by (metis SUP_apply bdd LIMSEQ_incseq_SUP)
qed
qed
qed
lemma lebesgue_measurable_imp_measurable_on_nnreal:
fixes u :: "'a::euclidean_space \<Rightarrow> real"
assumes "u \<in> borel_measurable lebesgue" "\<And>x. u x \<ge> 0""S \<in> sets lebesgue"
shows "u measurable_on S"
unfolding measurable_on_UNIV [symmetric, of u]
using assms
by (auto intro: lebesgue_measurable_imp_measurable_on_nnreal_UNIV)
lemma lebesgue_measurable_imp_measurable_on_real:
fixes u :: "'a::euclidean_space \<Rightarrow> real"
assumes u: "u \<in> borel_measurable lebesgue" and S: "S \<in> sets lebesgue"
shows "u measurable_on S"
proof -
let ?f = "\<lambda>x. \<bar>u x\<bar> + u x"
let ?g = "\<lambda>x. \<bar>u x\<bar> - u x"
have "?f measurable_on S" "?g measurable_on S"
using S u by (auto intro: lebesgue_measurable_imp_measurable_on_nnreal)
then have "(\<lambda>x. (?f x - ?g x) / 2) measurable_on S"
using measurable_on_cdivide measurable_on_diff by blast
then show ?thesis
by auto
qed
proposition lebesgue_measurable_imp_measurable_on:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes f: "f \<in> borel_measurable lebesgue" and S: "S \<in> sets lebesgue"
shows "f measurable_on S"
unfolding measurable_on_componentwise [of f]
proof
fix i::'b
assume "i \<in> Basis"
have "(\<lambda>x. (f x \<bullet> i)) \<in> borel_measurable lebesgue"
using \<open>i \<in> Basis\<close> borel_measurable_euclidean_space f by blast
then have "(\<lambda>x. (f x \<bullet> i)) measurable_on S"
using S lebesgue_measurable_imp_measurable_on_real by blast
then show "(\<lambda>x. (f x \<bullet> i) *\<^sub>R i) measurable_on S"
by (intro measurable_on_scaleR measurable_on_const S)
qed
proposition measurable_on_iff_borel_measurable:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "S \<in> sets lebesgue"
shows "f measurable_on S \<longleftrightarrow> f \<in> borel_measurable (lebesgue_on S)" (is "?lhs = ?rhs")
proof
show "f \<in> borel_measurable (lebesgue_on S)"
if "f measurable_on S"
using that by (simp add: assms measurable_on_imp_borel_measurable_lebesgue)
next
assume "f \<in> borel_measurable (lebesgue_on S)"
then have "(\<lambda>a. if a \<in> S then f a else 0) measurable_on UNIV"
by (simp add: assms borel_measurable_if lebesgue_measurable_imp_measurable_on)
then show "f measurable_on S"
using measurable_on_UNIV by blast
qed
subsection \<open>Measurability on generalisations of the binary product\<close>
lemma measurable_on_bilinear:
fixes h :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::euclidean_space"
assumes h: "bilinear h" and f: "f measurable_on S" and g: "g measurable_on S"
shows "(\<lambda>x. h (f x) (g x)) measurable_on S"
proof (rule measurable_on_combine [where h = h])
show "continuous_on UNIV (\<lambda>x. h (fst x) (snd x))"
by (simp add: bilinear_continuous_on_compose [OF continuous_on_fst continuous_on_snd h])
show "h 0 0 = 0"
by (simp add: bilinear_lzero h)
qed (auto intro: assms)
lemma borel_measurable_bilinear:
fixes h :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::euclidean_space"
assumes "bilinear h" "f \<in> borel_measurable (lebesgue_on S)" "g \<in> borel_measurable (lebesgue_on S)"
and S: "S \<in> sets lebesgue"
shows "(\<lambda>x. h (f x) (g x)) \<in> borel_measurable (lebesgue_on S)"
using assms measurable_on_bilinear [of h f S g]
by (simp flip: measurable_on_iff_borel_measurable)
lemma absolutely_integrable_bounded_measurable_product:
fixes h :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::euclidean_space"
assumes "bilinear h" and f: "f \<in> borel_measurable (lebesgue_on S)" "S \<in> sets lebesgue"
and bou: "bounded (f ` S)" and g: "g absolutely_integrable_on S"
shows "(\<lambda>x. h (f x) (g x)) absolutely_integrable_on S"
proof -
obtain B where "B > 0" and B: "\<And>x y. norm (h x y) \<le> B * norm x * norm y"
using bilinear_bounded_pos \<open>bilinear h\<close> by blast
obtain C where "C > 0" and C: "\<And>x. x \<in> S \<Longrightarrow> norm (f x) \<le> C"
using bounded_pos by (metis bou imageI)
show ?thesis
proof (rule measurable_bounded_by_integrable_imp_absolutely_integrable [OF _ \<open>S \<in> sets lebesgue\<close>])
show "norm (h (f x) (g x)) \<le> B * C * norm(g x)" if "x \<in> S" for x
by (meson less_le mult_left_mono mult_right_mono norm_ge_zero order_trans that \<open>B > 0\<close> B C)
show "(\<lambda>x. h (f x) (g x)) \<in> borel_measurable (lebesgue_on S)"
using \<open>bilinear h\<close> f g
by (blast intro: borel_measurable_bilinear dest: absolutely_integrable_measurable)
show "(\<lambda>x. B * C * norm(g x)) integrable_on S"
using \<open>0 < B\<close> \<open>0 < C\<close> absolutely_integrable_on_def g by auto
qed
qed
lemma absolutely_integrable_bounded_measurable_product_real:
fixes f :: "real \<Rightarrow> real"
assumes "f \<in> borel_measurable (lebesgue_on S)" "S \<in> sets lebesgue"
and "bounded (f ` S)" and "g absolutely_integrable_on S"
shows "(\<lambda>x. f x * g x) absolutely_integrable_on S"
using absolutely_integrable_bounded_measurable_product bilinear_times assms by blast
lemma borel_measurable_AE:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "f \<in> borel_measurable lebesgue" and ae: "AE x in lebesgue. f x = g x"
shows "g \<in> borel_measurable lebesgue"
proof -
obtain N where N: "N \<in> null_sets lebesgue" "\<And>x. x \<notin> N \<Longrightarrow> f x = g x"
using ae unfolding completion.AE_iff_null_sets by auto
have "f measurable_on UNIV"
by (simp add: assms lebesgue_measurable_imp_measurable_on)
then have "g measurable_on UNIV"
by (metis Diff_iff N measurable_on_spike negligible_iff_null_sets)
then show ?thesis
using measurable_on_imp_borel_measurable_lebesgue_UNIV by blast
qed
lemma has_bochner_integral_combine:
fixes f :: "real \<Rightarrow> 'a::euclidean_space"
assumes "a \<le> c" "c \<le> b"
and ac: "has_bochner_integral (lebesgue_on {a..c}) f i"
and cb: "has_bochner_integral (lebesgue_on {c..b}) f j"
shows "has_bochner_integral (lebesgue_on {a..b}) f(i + j)"
proof -
have i: "has_bochner_integral lebesgue (\<lambda>x. indicator {a..c} x *\<^sub>R f x) i"
and j: "has_bochner_integral lebesgue (\<lambda>x. indicator {c..b} x *\<^sub>R f x) j"
using assms by (auto simp: has_bochner_integral_restrict_space)
have AE: "AE x in lebesgue. indicat_real {a..c} x *\<^sub>R f x + indicat_real {c..b} x *\<^sub>R f x = indicat_real {a..b} x *\<^sub>R f x"
proof (rule AE_I')
have eq: "indicat_real {a..c} x *\<^sub>R f x + indicat_real {c..b} x *\<^sub>R f x = indicat_real {a..b} x *\<^sub>R f x" if "x \<noteq> c" for x
using assms that by (auto simp: indicator_def)
then show "{x \<in> space lebesgue. indicat_real {a..c} x *\<^sub>R f x + indicat_real {c..b} x *\<^sub>R f x \<noteq> indicat_real {a..b} x *\<^sub>R f x} \<subseteq> {c}"
by auto
qed auto
have "has_bochner_integral lebesgue (\<lambda>x. indicator {a..b} x *\<^sub>R f x) (i + j)"
proof (rule has_bochner_integralI_AE [OF has_bochner_integral_add [OF i j] _ AE])
have eq: "indicat_real {a..c} x *\<^sub>R f x + indicat_real {c..b} x *\<^sub>R f x = indicat_real {a..b} x *\<^sub>R f x" if "x \<noteq> c" for x
using assms that by (auto simp: indicator_def)
show "(\<lambda>x. indicat_real {a..b} x *\<^sub>R f x) \<in> borel_measurable lebesgue"
proof (rule borel_measurable_AE [OF borel_measurable_add AE])
show "(\<lambda>x. indicator {a..c} x *\<^sub>R f x) \<in> borel_measurable lebesgue"
"(\<lambda>x. indicator {c..b} x *\<^sub>R f x) \<in> borel_measurable lebesgue"
using i j by auto
qed
qed
then show ?thesis
by (simp add: has_bochner_integral_restrict_space)
qed
lemma integrable_combine:
fixes f :: "real \<Rightarrow> 'a::euclidean_space"
assumes "integrable (lebesgue_on {a..c}) f" "integrable (lebesgue_on {c..b}) f"
and "a \<le> c" "c \<le> b"
shows "integrable (lebesgue_on {a..b}) f"
using assms has_bochner_integral_combine has_bochner_integral_iff by blast
lemma integral_combine:
fixes f :: "real \<Rightarrow> 'a::euclidean_space"
assumes f: "integrable (lebesgue_on {a..b}) f" and "a \<le> c" "c \<le> b"
shows "integral\<^sup>L (lebesgue_on {a..b}) f = integral\<^sup>L (lebesgue_on {a..c}) f + integral\<^sup>L (lebesgue_on {c..b}) f"
proof -
have i: "has_bochner_integral (lebesgue_on {a..c}) f(integral\<^sup>L (lebesgue_on {a..c}) f)"
using integrable_subinterval \<open>c \<le> b\<close> f has_bochner_integral_iff by fastforce
have j: "has_bochner_integral (lebesgue_on {c..b}) f(integral\<^sup>L (lebesgue_on {c..b}) f)"
using integrable_subinterval \<open>a \<le> c\<close> f has_bochner_integral_iff by fastforce
show ?thesis
by (meson \<open>a \<le> c\<close> \<open>c \<le> b\<close> has_bochner_integral_combine has_bochner_integral_iff i j)
qed
lemma has_bochner_integral_null [intro]:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "N \<in> null_sets lebesgue"
shows "has_bochner_integral (lebesgue_on N) f 0"
unfolding has_bochner_integral_iff \<comment>\<open>strange that the proof's so long\<close>
proof
show "integrable (lebesgue_on N) f"
proof (subst integrable_restrict_space)
show "N \<inter> space lebesgue \<in> sets lebesgue"
using assms by force
show "integrable lebesgue (\<lambda>x. indicat_real N x *\<^sub>R f x)"
proof (rule integrable_cong_AE_imp)
show "integrable lebesgue (\<lambda>x. 0)"
by simp
show *: "AE x in lebesgue. 0 = indicat_real N x *\<^sub>R f x"
using assms
by (simp add: indicator_def completion.null_sets_iff_AE eventually_mono)
show "(\<lambda>x. indicat_real N x *\<^sub>R f x) \<in> borel_measurable lebesgue"
by (auto intro: borel_measurable_AE [OF _ *])
qed
qed
show "integral\<^sup>L (lebesgue_on N) f = 0"
proof (rule integral_eq_zero_AE)
show "AE x in lebesgue_on N. f x = 0"
by (rule AE_I' [where N=N]) (auto simp: assms null_setsD2 null_sets_restrict_space)
qed
qed
lemma has_bochner_integral_null_eq[simp]:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "N \<in> null_sets lebesgue"
shows "has_bochner_integral (lebesgue_on N) f i \<longleftrightarrow> i = 0"
using assms has_bochner_integral_eq by blast
end
diff --git a/src/HOL/Analysis/Extended_Real_Limits.thy b/src/HOL/Analysis/Extended_Real_Limits.thy
--- a/src/HOL/Analysis/Extended_Real_Limits.thy
+++ b/src/HOL/Analysis/Extended_Real_Limits.thy
@@ -1,1782 +1,1782 @@
(* 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> (* TO FIX: perhaps put all Nonstandard Analysis related
topics together? *)
theory Extended_Real_Limits
imports
Topology_Euclidean_Space
"HOL-Library.Extended_Real"
"HOL-Library.Extended_Nonnegative_Real"
"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\<^marker>\<open>tag unimportant\<close> enat :: second_countable_topology
proof
show "\<exists>B::enat set set. countable B \<and> open = generate_topology B"
proof (intro exI conjI)
show "countable (range lessThan \<union> range greaterThan::enat set set)"
by auto
qed (simp add: open_enat_def)
qed
instance\<^marker>\<open>tag unimportant\<close> 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 \<in> {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\<in>{0<..}. INF y\<in>(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)
+ by (auto simp: dist_commute)
then show "\<exists>r>0. Inf (f ` (Collect P)) \<le> Inf (f ` (S \<inter> ball x r - {x}))"
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>
Inf (f ` (S \<inter> ball x d - {x})) \<le> Inf (f ` (Collect P))"
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\<in>{0<..}. SUP y\<in>(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)
+ by (auto simp: dist_commute)
then show "\<exists>r>0. Sup (f ` (S \<inter> ball x r - {x})) \<le> Sup (f ` (Collect P))"
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>
Sup (f ` (Collect P)) \<le> Sup (f ` (S \<inter> ball x d - {x}))"
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\<in>{0<..}. INF y\<in>(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\<in>{0<..}. SUP y\<in>(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\<in>{0<..}. INF y\<in>ball x e. f y)"
apply (simp add: inf_min [symmetric] Liminf_at)
apply (subst inf_commute)
apply (subst SUP_inf)
apply auto
apply (metis (no_types, lifting) INF_insert centre_in_ball greaterThan_iff image_cong inf_commute insert_Diff)
done
subsection \<open>Extended-Real.thy\<close> (*FIX ME change title *)
lemma sum_constant_ereal:
fixes a::ereal
shows "(\<Sum>i\<in>I. a) = a * card I"
apply (cases "finite I", induct set: finite, simp_all)
apply (cases a, auto, metis (no_types, hide_lams) add.commute mult.commute semiring_normalization_rules(3))
done
lemma real_lim_then_eventually_real:
assumes "(u \<longlongrightarrow> ereal l) F"
shows "eventually (\<lambda>n. u n = ereal(real_of_ereal(u n))) F"
proof -
have "ereal l \<in> {-\<infinity><..<(\<infinity>::ereal)}" by simp
moreover have "open {-\<infinity><..<(\<infinity>::ereal)}" by simp
ultimately have "eventually (\<lambda>n. u n \<in> {-\<infinity><..<(\<infinity>::ereal)}) F" using assms tendsto_def by blast
moreover have "\<And>x. x \<in> {-\<infinity><..<(\<infinity>::ereal)} \<Longrightarrow> x = ereal(real_of_ereal x)" using ereal_real by auto
ultimately show ?thesis by (metis (mono_tags, lifting) eventually_mono)
qed
lemma ereal_Inf_cmult:
assumes "c>(0::real)"
shows "Inf {ereal c * x |x. P x} = ereal c * Inf {x. P x}"
proof -
have "(\<lambda>x::ereal. c * x) (Inf {x::ereal. P x}) = Inf ((\<lambda>x::ereal. c * x)`{x::ereal. P x})"
apply (rule mono_bij_Inf)
apply (simp add: assms ereal_mult_left_mono less_imp_le mono_def)
apply (rule bij_betw_byWitness[of _ "\<lambda>x. (x::ereal) / c"], auto simp add: assms ereal_mult_divide)
using assms ereal_divide_eq apply auto
done
then show ?thesis by (simp only: setcompr_eq_image[symmetric])
qed
subsubsection\<^marker>\<open>tag important\<close> \<open>Continuity of addition\<close>
text \<open>The next few lemmas remove an unnecessary assumption in \<open>tendsto_add_ereal\<close>, culminating
in \<open>tendsto_add_ereal_general\<close> which essentially says that the addition
is continuous on ereal times ereal, except at \<open>(-\<infinity>, \<infinity>)\<close> and \<open>(\<infinity>, -\<infinity>)\<close>.
It is much more convenient in many situations, see for instance the proof of
\<open>tendsto_sum_ereal\<close> below.\<close>
lemma tendsto_add_ereal_PInf:
fixes y :: ereal
assumes y: "y \<noteq> -\<infinity>"
assumes f: "(f \<longlongrightarrow> \<infinity>) F" and g: "(g \<longlongrightarrow> y) F"
shows "((\<lambda>x. f x + g x) \<longlongrightarrow> \<infinity>) F"
proof -
have "\<exists>C. eventually (\<lambda>x. g x > ereal C) F"
proof (cases y)
case (real r)
have "y > y-1" using y real by (simp add: ereal_between(1))
then have "eventually (\<lambda>x. g x > y - 1) F" using g y order_tendsto_iff by auto
moreover have "y-1 = ereal(real_of_ereal(y-1))"
by (metis real ereal_eq_1(1) ereal_minus(1) real_of_ereal.simps(1))
ultimately have "eventually (\<lambda>x. g x > ereal(real_of_ereal(y - 1))) F" by simp
then show ?thesis by auto
next
case (PInf)
have "eventually (\<lambda>x. g x > ereal 0) F" using g PInf by (simp add: tendsto_PInfty)
then show ?thesis by auto
qed (simp add: y)
then obtain C::real where ge: "eventually (\<lambda>x. g x > ereal C) F" by auto
{
fix M::real
have "eventually (\<lambda>x. f x > ereal(M - C)) F" using f by (simp add: tendsto_PInfty)
then have "eventually (\<lambda>x. (f x > ereal (M-C)) \<and> (g x > ereal C)) F"
by (auto simp add: ge eventually_conj_iff)
moreover have "\<And>x. ((f x > ereal (M-C)) \<and> (g x > ereal C)) \<Longrightarrow> (f x + g x > ereal M)"
using ereal_add_strict_mono2 by fastforce
ultimately have "eventually (\<lambda>x. f x + g x > ereal M) F" using eventually_mono by force
}
then show ?thesis by (simp add: tendsto_PInfty)
qed
text\<open>One would like to deduce the next lemma from the previous one, but the fact
that \<open>- (x + y)\<close> is in general different from \<open>(- x) + (- y)\<close> in ereal creates difficulties,
so it is more efficient to copy the previous proof.\<close>
lemma tendsto_add_ereal_MInf:
fixes y :: ereal
assumes y: "y \<noteq> \<infinity>"
assumes f: "(f \<longlongrightarrow> -\<infinity>) F" and g: "(g \<longlongrightarrow> y) F"
shows "((\<lambda>x. f x + g x) \<longlongrightarrow> -\<infinity>) F"
proof -
have "\<exists>C. eventually (\<lambda>x. g x < ereal C) F"
proof (cases y)
case (real r)
have "y < y+1" using y real by (simp add: ereal_between(1))
then have "eventually (\<lambda>x. g x < y + 1) F" using g y order_tendsto_iff by force
moreover have "y+1 = ereal(real_of_ereal (y+1))" by (simp add: real)
ultimately have "eventually (\<lambda>x. g x < ereal(real_of_ereal(y + 1))) F" by simp
then show ?thesis by auto
next
case (MInf)
have "eventually (\<lambda>x. g x < ereal 0) F" using g MInf by (simp add: tendsto_MInfty)
then show ?thesis by auto
qed (simp add: y)
then obtain C::real where ge: "eventually (\<lambda>x. g x < ereal C) F" by auto
{
fix M::real
have "eventually (\<lambda>x. f x < ereal(M - C)) F" using f by (simp add: tendsto_MInfty)
then have "eventually (\<lambda>x. (f x < ereal (M- C)) \<and> (g x < ereal C)) F"
by (auto simp add: ge eventually_conj_iff)
moreover have "\<And>x. ((f x < ereal (M-C)) \<and> (g x < ereal C)) \<Longrightarrow> (f x + g x < ereal M)"
using ereal_add_strict_mono2 by fastforce
ultimately have "eventually (\<lambda>x. f x + g x < ereal M) F" using eventually_mono by force
}
then show ?thesis by (simp add: tendsto_MInfty)
qed
lemma tendsto_add_ereal_general1:
fixes x y :: ereal
assumes y: "\<bar>y\<bar> \<noteq> \<infinity>"
assumes f: "(f \<longlongrightarrow> x) F" and g: "(g \<longlongrightarrow> y) F"
shows "((\<lambda>x. f x + g x) \<longlongrightarrow> x + y) F"
proof (cases x)
case (real r)
have a: "\<bar>x\<bar> \<noteq> \<infinity>" by (simp add: real)
show ?thesis by (rule tendsto_add_ereal[OF a, OF y, OF f, OF g])
next
case PInf
then show ?thesis using tendsto_add_ereal_PInf assms by force
next
case MInf
then show ?thesis using tendsto_add_ereal_MInf assms
by (metis abs_ereal.simps(3) ereal_MInfty_eq_plus)
qed
lemma tendsto_add_ereal_general2:
fixes x y :: ereal
assumes x: "\<bar>x\<bar> \<noteq> \<infinity>"
and f: "(f \<longlongrightarrow> x) F" and g: "(g \<longlongrightarrow> y) F"
shows "((\<lambda>x. f x + g x) \<longlongrightarrow> x + y) F"
proof -
have "((\<lambda>x. g x + f x) \<longlongrightarrow> x + y) F"
using tendsto_add_ereal_general1[OF x, OF g, OF f] add.commute[of "y", of "x"] by simp
moreover have "\<And>x. g x + f x = f x + g x" using add.commute by auto
ultimately show ?thesis by simp
qed
text \<open>The next lemma says that the addition is continuous on \<open>ereal\<close>, except at
the pairs \<open>(-\<infinity>, \<infinity>)\<close> and \<open>(\<infinity>, -\<infinity>)\<close>.\<close>
lemma tendsto_add_ereal_general [tendsto_intros]:
fixes x y :: ereal
assumes "\<not>((x=\<infinity> \<and> y=-\<infinity>) \<or> (x=-\<infinity> \<and> y=\<infinity>))"
and f: "(f \<longlongrightarrow> x) F" and g: "(g \<longlongrightarrow> y) F"
shows "((\<lambda>x. f x + g x) \<longlongrightarrow> x + y) F"
proof (cases x)
case (real r)
show ?thesis
apply (rule tendsto_add_ereal_general2) using real assms by auto
next
case (PInf)
then have "y \<noteq> -\<infinity>" using assms by simp
then show ?thesis using tendsto_add_ereal_PInf PInf assms by auto
next
case (MInf)
then have "y \<noteq> \<infinity>" using assms by simp
then show ?thesis using tendsto_add_ereal_MInf MInf f g by (metis ereal_MInfty_eq_plus)
qed
subsubsection\<^marker>\<open>tag important\<close> \<open>Continuity of multiplication\<close>
text \<open>In the same way as for addition, we prove that the multiplication is continuous on
ereal times ereal, except at \<open>(\<infinity>, 0)\<close> and \<open>(-\<infinity>, 0)\<close> and \<open>(0, \<infinity>)\<close> and \<open>(0, -\<infinity>)\<close>,
starting with specific situations.\<close>
lemma tendsto_mult_real_ereal:
assumes "(u \<longlongrightarrow> ereal l) F" "(v \<longlongrightarrow> ereal m) F"
shows "((\<lambda>n. u n * v n) \<longlongrightarrow> ereal l * ereal m) F"
proof -
have ureal: "eventually (\<lambda>n. u n = ereal(real_of_ereal(u n))) F" by (rule real_lim_then_eventually_real[OF assms(1)])
then have "((\<lambda>n. ereal(real_of_ereal(u n))) \<longlongrightarrow> ereal l) F" using assms by auto
then have limu: "((\<lambda>n. real_of_ereal(u n)) \<longlongrightarrow> l) F" by auto
have vreal: "eventually (\<lambda>n. v n = ereal(real_of_ereal(v n))) F" by (rule real_lim_then_eventually_real[OF assms(2)])
then have "((\<lambda>n. ereal(real_of_ereal(v n))) \<longlongrightarrow> ereal m) F" using assms by auto
then have limv: "((\<lambda>n. real_of_ereal(v n)) \<longlongrightarrow> m) F" by auto
{
fix n assume "u n = ereal(real_of_ereal(u n))" "v n = ereal(real_of_ereal(v n))"
then have "ereal(real_of_ereal(u n) * real_of_ereal(v n)) = u n * v n" by (metis times_ereal.simps(1))
}
then have *: "eventually (\<lambda>n. ereal(real_of_ereal(u n) * real_of_ereal(v n)) = u n * v n) F"
using eventually_elim2[OF ureal vreal] by auto
have "((\<lambda>n. real_of_ereal(u n) * real_of_ereal(v n)) \<longlongrightarrow> l * m) F" using tendsto_mult[OF limu limv] by auto
then have "((\<lambda>n. ereal(real_of_ereal(u n)) * real_of_ereal(v n)) \<longlongrightarrow> ereal(l * m)) F" by auto
then show ?thesis using * filterlim_cong by fastforce
qed
lemma tendsto_mult_ereal_PInf:
fixes f g::"_ \<Rightarrow> ereal"
assumes "(f \<longlongrightarrow> l) F" "l>0" "(g \<longlongrightarrow> \<infinity>) F"
shows "((\<lambda>x. f x * g x) \<longlongrightarrow> \<infinity>) F"
proof -
obtain a::real where "0 < ereal a" "a < l" using assms(2) using ereal_dense2 by blast
have *: "eventually (\<lambda>x. f x > a) F" using \<open>a < l\<close> assms(1) by (simp add: order_tendsto_iff)
{
fix K::real
define M where "M = max K 1"
then have "M > 0" by simp
then have "ereal(M/a) > 0" using \<open>ereal a > 0\<close> by simp
then have "\<And>x. ((f x > a) \<and> (g x > M/a)) \<Longrightarrow> (f x * g x > ereal a * ereal(M/a))"
using ereal_mult_mono_strict'[where ?c = "M/a", OF \<open>0 < ereal a\<close>] by auto
moreover have "ereal a * ereal(M/a) = M" using \<open>ereal a > 0\<close> by simp
ultimately have "\<And>x. ((f x > a) \<and> (g x > M/a)) \<Longrightarrow> (f x * g x > M)" by simp
moreover have "M \<ge> K" unfolding M_def by simp
ultimately have Imp: "\<And>x. ((f x > a) \<and> (g x > M/a)) \<Longrightarrow> (f x * g x > K)"
using ereal_less_eq(3) le_less_trans by blast
have "eventually (\<lambda>x. g x > M/a) F" using assms(3) by (simp add: tendsto_PInfty)
then have "eventually (\<lambda>x. (f x > a) \<and> (g x > M/a)) F"
using * by (auto simp add: eventually_conj_iff)
then have "eventually (\<lambda>x. f x * g x > K) F" using eventually_mono Imp by force
}
then show ?thesis by (auto simp add: tendsto_PInfty)
qed
lemma tendsto_mult_ereal_pos:
fixes f g::"_ \<Rightarrow> ereal"
assumes "(f \<longlongrightarrow> l) F" "(g \<longlongrightarrow> m) F" "l>0" "m>0"
shows "((\<lambda>x. f x * g x) \<longlongrightarrow> l * m) F"
proof (cases)
assume *: "l = \<infinity> \<or> m = \<infinity>"
then show ?thesis
proof (cases)
assume "m = \<infinity>"
then show ?thesis using tendsto_mult_ereal_PInf assms by auto
next
assume "\<not>(m = \<infinity>)"
then have "l = \<infinity>" using * by simp
then have "((\<lambda>x. g x * f x) \<longlongrightarrow> l * m) F" using tendsto_mult_ereal_PInf assms by auto
moreover have "\<And>x. g x * f x = f x * g x" using mult.commute by auto
ultimately show ?thesis by simp
qed
next
assume "\<not>(l = \<infinity> \<or> m = \<infinity>)"
then have "l < \<infinity>" "m < \<infinity>" by auto
then obtain lr mr where "l = ereal lr" "m = ereal mr"
using \<open>l>0\<close> \<open>m>0\<close> by (metis ereal_cases ereal_less(6) not_less_iff_gr_or_eq)
then show ?thesis using tendsto_mult_real_ereal assms by auto
qed
text \<open>We reduce the general situation to the positive case by multiplying by suitable signs.
Unfortunately, as ereal is not a ring, all the neat sign lemmas are not available there. We
give the bare minimum we need.\<close>
lemma ereal_sgn_abs:
fixes l::ereal
shows "sgn(l) * l = abs(l)"
apply (cases l) by (auto simp add: sgn_if ereal_less_uminus_reorder)
lemma sgn_squared_ereal:
assumes "l \<noteq> (0::ereal)"
shows "sgn(l) * sgn(l) = 1"
apply (cases l) using assms by (auto simp add: one_ereal_def sgn_if)
lemma tendsto_mult_ereal [tendsto_intros]:
fixes f g::"_ \<Rightarrow> ereal"
assumes "(f \<longlongrightarrow> l) F" "(g \<longlongrightarrow> m) F" "\<not>((l=0 \<and> abs(m) = \<infinity>) \<or> (m=0 \<and> abs(l) = \<infinity>))"
shows "((\<lambda>x. f x * g x) \<longlongrightarrow> l * m) F"
proof (cases)
assume "l=0 \<or> m=0"
then have "abs(l) \<noteq> \<infinity>" "abs(m) \<noteq> \<infinity>" using assms(3) by auto
then obtain lr mr where "l = ereal lr" "m = ereal mr" by auto
then show ?thesis using tendsto_mult_real_ereal assms by auto
next
have sgn_finite: "\<And>a::ereal. abs(sgn a) \<noteq> \<infinity>"
by (metis MInfty_neq_ereal(2) PInfty_neq_ereal(2) abs_eq_infinity_cases ereal_times(1) ereal_times(3) ereal_uminus_eq_reorder sgn_ereal.elims)
then have sgn_finite2: "\<And>a b::ereal. abs(sgn a * sgn b) \<noteq> \<infinity>"
by (metis abs_eq_infinity_cases abs_ereal.simps(2) abs_ereal.simps(3) ereal_mult_eq_MInfty ereal_mult_eq_PInfty)
assume "\<not>(l=0 \<or> m=0)"
then have "l \<noteq> 0" "m \<noteq> 0" by auto
then have "abs(l) > 0" "abs(m) > 0"
by (metis abs_ereal_ge0 abs_ereal_less0 abs_ereal_pos ereal_uminus_uminus ereal_uminus_zero less_le not_less)+
then have "sgn(l) * l > 0" "sgn(m) * m > 0" using ereal_sgn_abs by auto
moreover have "((\<lambda>x. sgn(l) * f x) \<longlongrightarrow> (sgn(l) * l)) F"
by (rule tendsto_cmult_ereal, auto simp add: sgn_finite assms(1))
moreover have "((\<lambda>x. sgn(m) * g x) \<longlongrightarrow> (sgn(m) * m)) F"
by (rule tendsto_cmult_ereal, auto simp add: sgn_finite assms(2))
ultimately have *: "((\<lambda>x. (sgn(l) * f x) * (sgn(m) * g x)) \<longlongrightarrow> (sgn(l) * l) * (sgn(m) * m)) F"
using tendsto_mult_ereal_pos by force
have "((\<lambda>x. (sgn(l) * sgn(m)) * ((sgn(l) * f x) * (sgn(m) * g x))) \<longlongrightarrow> (sgn(l) * sgn(m)) * ((sgn(l) * l) * (sgn(m) * m))) F"
by (rule tendsto_cmult_ereal, auto simp add: sgn_finite2 *)
moreover have "\<And>x. (sgn(l) * sgn(m)) * ((sgn(l) * f x) * (sgn(m) * g x)) = f x * g x"
by (metis mult.left_neutral sgn_squared_ereal[OF \<open>l \<noteq> 0\<close>] sgn_squared_ereal[OF \<open>m \<noteq> 0\<close>] mult.assoc mult.commute)
moreover have "(sgn(l) * sgn(m)) * ((sgn(l) * l) * (sgn(m) * m)) = l * m"
by (metis mult.left_neutral sgn_squared_ereal[OF \<open>l \<noteq> 0\<close>] sgn_squared_ereal[OF \<open>m \<noteq> 0\<close>] mult.assoc mult.commute)
ultimately show ?thesis by auto
qed
lemma tendsto_cmult_ereal_general [tendsto_intros]:
fixes f::"_ \<Rightarrow> ereal" and c::ereal
assumes "(f \<longlongrightarrow> l) F" "\<not> (l=0 \<and> abs(c) = \<infinity>)"
shows "((\<lambda>x. c * f x) \<longlongrightarrow> c * l) F"
by (cases "c = 0", auto simp add: assms tendsto_mult_ereal)
subsubsection\<^marker>\<open>tag important\<close> \<open>Continuity of division\<close>
lemma tendsto_inverse_ereal_PInf:
fixes u::"_ \<Rightarrow> ereal"
assumes "(u \<longlongrightarrow> \<infinity>) F"
shows "((\<lambda>x. 1/ u x) \<longlongrightarrow> 0) F"
proof -
{
fix e::real assume "e>0"
have "1/e < \<infinity>" by auto
then have "eventually (\<lambda>n. u n > 1/e) F" using assms(1) by (simp add: tendsto_PInfty)
moreover
{
fix z::ereal assume "z>1/e"
then have "z>0" using \<open>e>0\<close> using less_le_trans not_le by fastforce
then have "1/z \<ge> 0" by auto
moreover have "1/z < e" using \<open>e>0\<close> \<open>z>1/e\<close>
apply (cases z) apply auto
by (metis (mono_tags, hide_lams) less_ereal.simps(2) less_ereal.simps(4) divide_less_eq ereal_divide_less_pos ereal_less(4)
ereal_less_eq(4) less_le_trans mult_eq_0_iff not_le not_one_less_zero times_ereal.simps(1))
ultimately have "1/z \<ge> 0" "1/z < e" by auto
}
ultimately have "eventually (\<lambda>n. 1/u n<e) F" "eventually (\<lambda>n. 1/u n\<ge>0) F" by (auto simp add: eventually_mono)
} note * = this
show ?thesis
proof (subst order_tendsto_iff, auto)
fix a::ereal assume "a<0"
then show "eventually (\<lambda>n. 1/u n > a) F" using *(2) eventually_mono less_le_trans linordered_field_no_ub by fastforce
next
fix a::ereal assume "a>0"
then obtain e::real where "e>0" "a>e" using ereal_dense2 ereal_less(2) by blast
then have "eventually (\<lambda>n. 1/u n < e) F" using *(1) by auto
then show "eventually (\<lambda>n. 1/u n < a) F" using \<open>a>e\<close> by (metis (mono_tags, lifting) eventually_mono less_trans)
qed
qed
text \<open>The next lemma deserves to exist by itself, as it is so common and useful.\<close>
lemma tendsto_inverse_real [tendsto_intros]:
fixes u::"_ \<Rightarrow> real"
shows "(u \<longlongrightarrow> l) F \<Longrightarrow> l \<noteq> 0 \<Longrightarrow> ((\<lambda>x. 1/ u x) \<longlongrightarrow> 1/l) F"
using tendsto_inverse unfolding inverse_eq_divide .
lemma tendsto_inverse_ereal [tendsto_intros]:
fixes u::"_ \<Rightarrow> ereal"
assumes "(u \<longlongrightarrow> l) F" "l \<noteq> 0"
shows "((\<lambda>x. 1/ u x) \<longlongrightarrow> 1/l) F"
proof (cases l)
case (real r)
then have "r \<noteq> 0" using assms(2) by auto
then have "1/l = ereal(1/r)" using real by (simp add: one_ereal_def)
define v where "v = (\<lambda>n. real_of_ereal(u n))"
have ureal: "eventually (\<lambda>n. u n = ereal(v n)) F" unfolding v_def using real_lim_then_eventually_real assms(1) real by auto
then have "((\<lambda>n. ereal(v n)) \<longlongrightarrow> ereal r) F" using assms real v_def by auto
then have *: "((\<lambda>n. v n) \<longlongrightarrow> r) F" by auto
then have "((\<lambda>n. 1/v n) \<longlongrightarrow> 1/r) F" using \<open>r \<noteq> 0\<close> tendsto_inverse_real by auto
then have lim: "((\<lambda>n. ereal(1/v n)) \<longlongrightarrow> 1/l) F" using \<open>1/l = ereal(1/r)\<close> by auto
have "r \<in> -{0}" "open (-{(0::real)})" using \<open>r \<noteq> 0\<close> by auto
then have "eventually (\<lambda>n. v n \<in> -{0}) F" using * using topological_tendstoD by blast
then have "eventually (\<lambda>n. v n \<noteq> 0) F" by auto
moreover
{
fix n assume H: "v n \<noteq> 0" "u n = ereal(v n)"
then have "ereal(1/v n) = 1/ereal(v n)" by (simp add: one_ereal_def)
then have "ereal(1/v n) = 1/u n" using H(2) by simp
}
ultimately have "eventually (\<lambda>n. ereal(1/v n) = 1/u n) F" using ureal eventually_elim2 by force
with Lim_transform_eventually[OF lim this] show ?thesis by simp
next
case (PInf)
then have "1/l = 0" by auto
then show ?thesis using tendsto_inverse_ereal_PInf assms PInf by auto
next
case (MInf)
then have "1/l = 0" by auto
have "1/z = -1/ -z" if "z < 0" for z::ereal
apply (cases z) using divide_ereal_def \<open> z < 0 \<close> by auto
moreover have "eventually (\<lambda>n. u n < 0) F" by (metis (no_types) MInf assms(1) tendsto_MInfty zero_ereal_def)
ultimately have *: "eventually (\<lambda>n. -1/-u n = 1/u n) F" by (simp add: eventually_mono)
define v where "v = (\<lambda>n. - u n)"
have "(v \<longlongrightarrow> \<infinity>) F" unfolding v_def using MInf assms(1) tendsto_uminus_ereal by fastforce
then have "((\<lambda>n. 1/v n) \<longlongrightarrow> 0) F" using tendsto_inverse_ereal_PInf by auto
then have "((\<lambda>n. -1/v n) \<longlongrightarrow> 0) F" using tendsto_uminus_ereal by fastforce
then show ?thesis unfolding v_def using Lim_transform_eventually[OF _ *] \<open> 1/l = 0 \<close> by auto
qed
lemma tendsto_divide_ereal [tendsto_intros]:
fixes f g::"_ \<Rightarrow> ereal"
assumes "(f \<longlongrightarrow> l) F" "(g \<longlongrightarrow> m) F" "m \<noteq> 0" "\<not>(abs(l) = \<infinity> \<and> abs(m) = \<infinity>)"
shows "((\<lambda>x. f x / g x) \<longlongrightarrow> l / m) F"
proof -
define h where "h = (\<lambda>x. 1/ g x)"
have *: "(h \<longlongrightarrow> 1/m) F" unfolding h_def using assms(2) assms(3) tendsto_inverse_ereal by auto
have "((\<lambda>x. f x * h x) \<longlongrightarrow> l * (1/m)) F"
apply (rule tendsto_mult_ereal[OF assms(1) *]) using assms(3) assms(4) by (auto simp add: divide_ereal_def)
moreover have "f x * h x = f x / g x" for x unfolding h_def by (simp add: divide_ereal_def)
moreover have "l * (1/m) = l/m" by (simp add: divide_ereal_def)
ultimately show ?thesis unfolding h_def using Lim_transform_eventually by auto
qed
subsubsection \<open>Further limits\<close>
text \<open>The assumptions of @{thm tendsto_diff_ereal} are too strong, we weaken them here.\<close>
lemma tendsto_diff_ereal_general [tendsto_intros]:
fixes u v::"'a \<Rightarrow> ereal"
assumes "(u \<longlongrightarrow> l) F" "(v \<longlongrightarrow> m) F" "\<not>((l = \<infinity> \<and> m = \<infinity>) \<or> (l = -\<infinity> \<and> m = -\<infinity>))"
shows "((\<lambda>n. u n - v n) \<longlongrightarrow> l - m) F"
proof -
have "((\<lambda>n. u n + (-v n)) \<longlongrightarrow> l + (-m)) F"
apply (intro tendsto_intros assms) using assms by (auto simp add: ereal_uminus_eq_reorder)
then show ?thesis by (simp add: minus_ereal_def)
qed
lemma id_nat_ereal_tendsto_PInf [tendsto_intros]:
"(\<lambda> n::nat. real n) \<longlonglongrightarrow> \<infinity>"
by (simp add: filterlim_real_sequentially tendsto_PInfty_eq_at_top)
lemma tendsto_at_top_pseudo_inverse [tendsto_intros]:
fixes u::"nat \<Rightarrow> nat"
assumes "LIM n sequentially. u n :> at_top"
shows "LIM n sequentially. Inf {N. u N \<ge> n} :> at_top"
proof -
{
fix C::nat
define M where "M = Max {u n| n. n \<le> C}+1"
{
fix n assume "n \<ge> M"
have "eventually (\<lambda>N. u N \<ge> n) sequentially" using assms
by (simp add: filterlim_at_top)
then have *: "{N. u N \<ge> n} \<noteq> {}" by force
have "N > C" if "u N \<ge> n" for N
proof (rule ccontr)
assume "\<not>(N > C)"
have "u N \<le> Max {u n| n. n \<le> C}"
apply (rule Max_ge) using \<open>\<not>(N > C)\<close> by auto
then show False using \<open>u N \<ge> n\<close> \<open>n \<ge> M\<close> unfolding M_def by auto
qed
then have **: "{N. u N \<ge> n} \<subseteq> {C..}" by fastforce
have "Inf {N. u N \<ge> n} \<ge> C"
by (metis "*" "**" Inf_nat_def1 atLeast_iff subset_eq)
}
then have "eventually (\<lambda>n. Inf {N. u N \<ge> n} \<ge> C) sequentially"
using eventually_sequentially by auto
}
then show ?thesis using filterlim_at_top by auto
qed
lemma pseudo_inverse_finite_set:
fixes u::"nat \<Rightarrow> nat"
assumes "LIM n sequentially. u n :> at_top"
shows "finite {N. u N \<le> n}"
proof -
fix n
have "eventually (\<lambda>N. u N \<ge> n+1) sequentially" using assms
by (simp add: filterlim_at_top)
then obtain N1 where N1: "\<And>N. N \<ge> N1 \<Longrightarrow> u N \<ge> n + 1"
using eventually_sequentially by auto
have "{N. u N \<le> n} \<subseteq> {..<N1}"
apply auto using N1 by (metis Suc_eq_plus1 not_less not_less_eq_eq)
then show "finite {N. u N \<le> n}" by (simp add: finite_subset)
qed
lemma tendsto_at_top_pseudo_inverse2 [tendsto_intros]:
fixes u::"nat \<Rightarrow> nat"
assumes "LIM n sequentially. u n :> at_top"
shows "LIM n sequentially. Max {N. u N \<le> n} :> at_top"
proof -
{
fix N0::nat
have "N0 \<le> Max {N. u N \<le> n}" if "n \<ge> u N0" for n
apply (rule Max.coboundedI) using pseudo_inverse_finite_set[OF assms] that by auto
then have "eventually (\<lambda>n. N0 \<le> Max {N. u N \<le> n}) sequentially"
using eventually_sequentially by blast
}
then show ?thesis using filterlim_at_top by auto
qed
lemma ereal_truncation_top [tendsto_intros]:
fixes x::ereal
shows "(\<lambda>n::nat. min x n) \<longlonglongrightarrow> x"
proof (cases x)
case (real r)
then obtain K::nat where "K>0" "K > abs(r)" using reals_Archimedean2 gr0I by auto
then have "min x n = x" if "n \<ge> K" for n apply (subst real, subst real, auto) using that eq_iff by fastforce
then have "eventually (\<lambda>n. min x n = x) sequentially" using eventually_at_top_linorder by blast
then show ?thesis by (simp add: tendsto_eventually)
next
case (PInf)
then have "min x n = n" for n::nat by (auto simp add: min_def)
then show ?thesis using id_nat_ereal_tendsto_PInf PInf by auto
next
case (MInf)
then have "min x n = x" for n::nat by (auto simp add: min_def)
then show ?thesis by auto
qed
lemma ereal_truncation_real_top [tendsto_intros]:
fixes x::ereal
assumes "x \<noteq> - \<infinity>"
shows "(\<lambda>n::nat. real_of_ereal(min x n)) \<longlonglongrightarrow> x"
proof (cases x)
case (real r)
then obtain K::nat where "K>0" "K > abs(r)" using reals_Archimedean2 gr0I by auto
then have "min x n = x" if "n \<ge> K" for n apply (subst real, subst real, auto) using that eq_iff by fastforce
then have "real_of_ereal(min x n) = r" if "n \<ge> K" for n using real that by auto
then have "eventually (\<lambda>n. real_of_ereal(min x n) = r) sequentially" using eventually_at_top_linorder by blast
then have "(\<lambda>n. real_of_ereal(min x n)) \<longlonglongrightarrow> r" by (simp add: tendsto_eventually)
then show ?thesis using real by auto
next
case (PInf)
then have "real_of_ereal(min x n) = n" for n::nat by (auto simp add: min_def)
then show ?thesis using id_nat_ereal_tendsto_PInf PInf by auto
qed (simp add: assms)
lemma ereal_truncation_bottom [tendsto_intros]:
fixes x::ereal
shows "(\<lambda>n::nat. max x (- real n)) \<longlonglongrightarrow> x"
proof (cases x)
case (real r)
then obtain K::nat where "K>0" "K > abs(r)" using reals_Archimedean2 gr0I by auto
then have "max x (-real n) = x" if "n \<ge> K" for n apply (subst real, subst real, auto) using that eq_iff by fastforce
then have "eventually (\<lambda>n. max x (-real n) = x) sequentially" using eventually_at_top_linorder by blast
then show ?thesis by (simp add: tendsto_eventually)
next
case (MInf)
then have "max x (-real n) = (-1)* ereal(real n)" for n::nat by (auto simp add: max_def)
moreover have "(\<lambda>n. (-1)* ereal(real n)) \<longlonglongrightarrow> -\<infinity>"
using tendsto_cmult_ereal[of "-1", OF _ id_nat_ereal_tendsto_PInf] by (simp add: one_ereal_def)
ultimately show ?thesis using MInf by auto
next
case (PInf)
then have "max x (-real n) = x" for n::nat by (auto simp add: max_def)
then show ?thesis by auto
qed
lemma ereal_truncation_real_bottom [tendsto_intros]:
fixes x::ereal
assumes "x \<noteq> \<infinity>"
shows "(\<lambda>n::nat. real_of_ereal(max x (- real n))) \<longlonglongrightarrow> x"
proof (cases x)
case (real r)
then obtain K::nat where "K>0" "K > abs(r)" using reals_Archimedean2 gr0I by auto
then have "max x (-real n) = x" if "n \<ge> K" for n apply (subst real, subst real, auto) using that eq_iff by fastforce
then have "real_of_ereal(max x (-real n)) = r" if "n \<ge> K" for n using real that by auto
then have "eventually (\<lambda>n. real_of_ereal(max x (-real n)) = r) sequentially" using eventually_at_top_linorder by blast
then have "(\<lambda>n. real_of_ereal(max x (-real n))) \<longlonglongrightarrow> r" by (simp add: tendsto_eventually)
then show ?thesis using real by auto
next
case (MInf)
then have "real_of_ereal(max x (-real n)) = (-1)* ereal(real n)" for n::nat by (auto simp add: max_def)
moreover have "(\<lambda>n. (-1)* ereal(real n)) \<longlonglongrightarrow> -\<infinity>"
using tendsto_cmult_ereal[of "-1", OF _ id_nat_ereal_tendsto_PInf] by (simp add: one_ereal_def)
ultimately show ?thesis using MInf by auto
qed (simp add: assms)
text \<open>the next one is copied from \<open>tendsto_sum\<close>.\<close>
lemma tendsto_sum_ereal [tendsto_intros]:
fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> ereal"
assumes "\<And>i. i \<in> S \<Longrightarrow> (f i \<longlongrightarrow> a i) F"
"\<And>i. abs(a i) \<noteq> \<infinity>"
shows "((\<lambda>x. \<Sum>i\<in>S. f i x) \<longlongrightarrow> (\<Sum>i\<in>S. a i)) F"
proof (cases "finite S")
assume "finite S" then show ?thesis using assms
by (induct, simp, simp add: tendsto_add_ereal_general2 assms)
qed(simp)
lemma continuous_ereal_abs:
"continuous_on (UNIV::ereal set) abs"
proof -
have "continuous_on ({..0} \<union> {(0::ereal)..}) abs"
apply (rule continuous_on_closed_Un, auto)
apply (rule iffD1[OF continuous_on_cong, of "{..0}" _ "\<lambda>x. -x"])
using less_eq_ereal_def apply (auto simp add: continuous_uminus_ereal)
apply (rule iffD1[OF continuous_on_cong, of "{0..}" _ "\<lambda>x. x"])
apply (auto)
done
moreover have "(UNIV::ereal set) = {..0} \<union> {(0::ereal)..}" by auto
ultimately show ?thesis by auto
qed
lemmas continuous_on_compose_ereal_abs[continuous_intros] =
continuous_on_compose2[OF continuous_ereal_abs _ subset_UNIV]
lemma tendsto_abs_ereal [tendsto_intros]:
assumes "(u \<longlongrightarrow> (l::ereal)) F"
shows "((\<lambda>n. abs(u n)) \<longlongrightarrow> abs l) F"
using continuous_ereal_abs assms by (metis UNIV_I continuous_on tendsto_compose)
lemma ereal_minus_real_tendsto_MInf [tendsto_intros]:
"(\<lambda>x. ereal (- real x)) \<longlonglongrightarrow> - \<infinity>"
by (subst uminus_ereal.simps(1)[symmetric], intro tendsto_intros)
subsection \<open>Extended-Nonnegative-Real.thy\<close> (*FIX title *)
lemma tendsto_diff_ennreal_general [tendsto_intros]:
fixes u v::"'a \<Rightarrow> ennreal"
assumes "(u \<longlongrightarrow> l) F" "(v \<longlongrightarrow> m) F" "\<not>(l = \<infinity> \<and> m = \<infinity>)"
shows "((\<lambda>n. u n - v n) \<longlongrightarrow> l - m) F"
proof -
have "((\<lambda>n. e2ennreal(enn2ereal(u n) - enn2ereal(v n))) \<longlongrightarrow> e2ennreal(enn2ereal l - enn2ereal m)) F"
apply (intro tendsto_intros) using assms by auto
then show ?thesis by auto
qed
lemma tendsto_mult_ennreal [tendsto_intros]:
fixes l m::ennreal
assumes "(u \<longlongrightarrow> l) F" "(v \<longlongrightarrow> m) F" "\<not>((l = 0 \<and> m = \<infinity>) \<or> (l = \<infinity> \<and> m = 0))"
shows "((\<lambda>n. u n * v n) \<longlongrightarrow> l * m) F"
proof -
have "((\<lambda>n. e2ennreal(enn2ereal (u n) * enn2ereal (v n))) \<longlongrightarrow> e2ennreal(enn2ereal l * enn2ereal m)) F"
apply (intro tendsto_intros) using assms apply auto
using enn2ereal_inject zero_ennreal.rep_eq by fastforce+
moreover have "e2ennreal(enn2ereal (u n) * enn2ereal (v n)) = u n * v n" for n
by (subst times_ennreal.abs_eq[symmetric], auto simp add: eq_onp_same_args)
moreover have "e2ennreal(enn2ereal l * enn2ereal m) = l * m"
by (subst times_ennreal.abs_eq[symmetric], auto simp add: eq_onp_same_args)
ultimately show ?thesis
by auto
qed
subsection \<open>monoset\<close> (*FIX ME title *)
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" "Inf (f ` (Collect P)) \<in> S"
{
fix x
assume "P x"
then have "Inf (f ` (Collect P)) \<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> Inf (f ` (Collect P)) \<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 "Inf (f ` {x. B < f x}) \<le> y"
using P by auto
moreover have "B \<le> Inf (f ` {x. B < f x})"
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)" "Sup (f ` (Collect P)) \<in> S"
{
fix x
assume "P x"
then have "f x \<le> Sup (f ` (Collect P))"
by (intro complete_lattice_class.SUP_upper) simp
with S(1)[unfolded mono_set, rule_format, of "- Sup (f ` (Collect P))" "- 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> Sup (f ` (Collect P))"
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> Sup (f ` {x. f x < B})"
using P by auto
moreover have "Sup (f ` {x. f x < B}) \<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
lemma limsup_finite_then_bounded:
fixes u::"nat \<Rightarrow> real"
assumes "limsup u < \<infinity>"
shows "\<exists>C. \<forall>n. u n \<le> C"
proof -
obtain C where C: "limsup u < C" "C < \<infinity>" using assms ereal_dense2 by blast
then have "C = ereal(real_of_ereal C)" using ereal_real by force
have "eventually (\<lambda>n. u n < C) sequentially" using C(1) unfolding Limsup_def
apply (auto simp add: INF_less_iff)
using SUP_lessD eventually_mono by fastforce
then obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> u n < C" using eventually_sequentially by auto
define D where "D = max (real_of_ereal C) (Max {u n |n. n \<le> N})"
have "\<And>n. u n \<le> D"
proof -
fix n show "u n \<le> D"
proof (cases)
assume *: "n \<le> N"
have "u n \<le> Max {u n |n. n \<le> N}" by (rule Max_ge, auto simp add: *)
then show "u n \<le> D" unfolding D_def by linarith
next
assume "\<not>(n \<le> N)"
then have "n \<ge> N" by simp
then have "u n < C" using N by auto
then have "u n < real_of_ereal C" using \<open>C = ereal(real_of_ereal C)\<close> less_ereal.simps(1) by fastforce
then show "u n \<le> D" unfolding D_def by linarith
qed
qed
then show ?thesis by blast
qed
lemma liminf_finite_then_bounded_below:
fixes u::"nat \<Rightarrow> real"
assumes "liminf u > -\<infinity>"
shows "\<exists>C. \<forall>n. u n \<ge> C"
proof -
obtain C where C: "liminf u > C" "C > -\<infinity>" using assms using ereal_dense2 by blast
then have "C = ereal(real_of_ereal C)" using ereal_real by force
have "eventually (\<lambda>n. u n > C) sequentially" using C(1) unfolding Liminf_def
apply (auto simp add: less_SUP_iff)
using eventually_elim2 less_INF_D by fastforce
then obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> u n > C" using eventually_sequentially by auto
define D where "D = min (real_of_ereal C) (Min {u n |n. n \<le> N})"
have "\<And>n. u n \<ge> D"
proof -
fix n show "u n \<ge> D"
proof (cases)
assume *: "n \<le> N"
have "u n \<ge> Min {u n |n. n \<le> N}" by (rule Min_le, auto simp add: *)
then show "u n \<ge> D" unfolding D_def by linarith
next
assume "\<not>(n \<le> N)"
then have "n \<ge> N" by simp
then have "u n > C" using N by auto
then have "u n > real_of_ereal C" using \<open>C = ereal(real_of_ereal C)\<close> less_ereal.simps(1) by fastforce
then show "u n \<ge> D" unfolding D_def by linarith
qed
qed
then show ?thesis by blast
qed
lemma liminf_upper_bound:
fixes u:: "nat \<Rightarrow> ereal"
assumes "liminf u < l"
shows "\<exists>N>k. u N < l"
by (metis assms gt_ex less_le_trans liminf_bounded_iff not_less)
lemma limsup_shift:
"limsup (\<lambda>n. u (n+1)) = limsup u"
proof -
have "(SUP m\<in>{n+1..}. u m) = (SUP m\<in>{n..}. u (m + 1))" for n
apply (rule SUP_eq) using Suc_le_D by auto
then have a: "(INF n. SUP m\<in>{n..}. u (m + 1)) = (INF n. (SUP m\<in>{n+1..}. u m))" by auto
have b: "(INF n. (SUP m\<in>{n+1..}. u m)) = (INF n\<in>{1..}. (SUP m\<in>{n..}. u m))"
apply (rule INF_eq) using Suc_le_D by auto
have "(INF n\<in>{1..}. v n) = (INF n. v n)" if "decseq v" for v::"nat \<Rightarrow> 'a"
apply (rule INF_eq) using \<open>decseq v\<close> decseq_Suc_iff by auto
moreover have "decseq (\<lambda>n. (SUP m\<in>{n..}. u m))" by (simp add: SUP_subset_mono decseq_def)
ultimately have c: "(INF n\<in>{1..}. (SUP m\<in>{n..}. u m)) = (INF n. (SUP m\<in>{n..}. u m))" by simp
have "(INF n. Sup (u ` {n..})) = (INF n. SUP m\<in>{n..}. u (m + 1))" using a b c by simp
then show ?thesis by (auto cong: limsup_INF_SUP)
qed
lemma limsup_shift_k:
"limsup (\<lambda>n. u (n+k)) = limsup u"
proof (induction k)
case (Suc k)
have "limsup (\<lambda>n. u (n+k+1)) = limsup (\<lambda>n. u (n+k))" using limsup_shift[where ?u="\<lambda>n. u(n+k)"] by simp
then show ?case using Suc.IH by simp
qed (auto)
lemma liminf_shift:
"liminf (\<lambda>n. u (n+1)) = liminf u"
proof -
have "(INF m\<in>{n+1..}. u m) = (INF m\<in>{n..}. u (m + 1))" for n
apply (rule INF_eq) using Suc_le_D by (auto)
then have a: "(SUP n. INF m\<in>{n..}. u (m + 1)) = (SUP n. (INF m\<in>{n+1..}. u m))" by auto
have b: "(SUP n. (INF m\<in>{n+1..}. u m)) = (SUP n\<in>{1..}. (INF m\<in>{n..}. u m))"
apply (rule SUP_eq) using Suc_le_D by (auto)
have "(SUP n\<in>{1..}. v n) = (SUP n. v n)" if "incseq v" for v::"nat \<Rightarrow> 'a"
apply (rule SUP_eq) using \<open>incseq v\<close> incseq_Suc_iff by auto
moreover have "incseq (\<lambda>n. (INF m\<in>{n..}. u m))" by (simp add: INF_superset_mono mono_def)
ultimately have c: "(SUP n\<in>{1..}. (INF m\<in>{n..}. u m)) = (SUP n. (INF m\<in>{n..}. u m))" by simp
have "(SUP n. Inf (u ` {n..})) = (SUP n. INF m\<in>{n..}. u (m + 1))" using a b c by simp
then show ?thesis by (auto cong: liminf_SUP_INF)
qed
lemma liminf_shift_k:
"liminf (\<lambda>n. u (n+k)) = liminf u"
proof (induction k)
case (Suc k)
have "liminf (\<lambda>n. u (n+k+1)) = liminf (\<lambda>n. u (n+k))" using liminf_shift[where ?u="\<lambda>n. u(n+k)"] by simp
then show ?case using Suc.IH by simp
qed (auto)
lemma Limsup_obtain:
fixes u::"_ \<Rightarrow> 'a :: complete_linorder"
assumes "Limsup F u > c"
shows "\<exists>i. u i > c"
proof -
have "(INF P\<in>{P. eventually P F}. SUP x\<in>{x. P x}. u x) > c" using assms by (simp add: Limsup_def)
then show ?thesis by (metis eventually_True mem_Collect_eq less_INF_D less_SUP_iff)
qed
text \<open>The next lemma is extremely useful, as it often makes it possible to reduce statements
about limsups to statements about limits.\<close>
lemma limsup_subseq_lim:
fixes u::"nat \<Rightarrow> 'a :: {complete_linorder, linorder_topology}"
shows "\<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> (u o r) \<longlonglongrightarrow> limsup u"
proof (cases)
assume "\<forall>n. \<exists>p>n. \<forall>m\<ge>p. u m \<le> u p"
then have "\<exists>r. \<forall>n. (\<forall>m\<ge>r n. u m \<le> u (r n)) \<and> r n < r (Suc n)"
by (intro dependent_nat_choice) (auto simp: conj_commute)
then obtain r :: "nat \<Rightarrow> nat" where "strict_mono r" and mono: "\<And>n m. r n \<le> m \<Longrightarrow> u m \<le> u (r n)"
by (auto simp: strict_mono_Suc_iff)
define umax where "umax = (\<lambda>n. (SUP m\<in>{n..}. u m))"
have "decseq umax" unfolding umax_def by (simp add: SUP_subset_mono antimono_def)
then have "umax \<longlonglongrightarrow> limsup u" unfolding umax_def by (metis LIMSEQ_INF limsup_INF_SUP)
then have *: "(umax o r) \<longlonglongrightarrow> limsup u" by (simp add: LIMSEQ_subseq_LIMSEQ \<open>strict_mono r\<close>)
have "\<And>n. umax(r n) = u(r n)" unfolding umax_def using mono
by (metis SUP_le_iff antisym atLeast_def mem_Collect_eq order_refl)
then have "umax o r = u o r" unfolding o_def by simp
then have "(u o r) \<longlonglongrightarrow> limsup u" using * by simp
then show ?thesis using \<open>strict_mono r\<close> by blast
next
assume "\<not> (\<forall>n. \<exists>p>n. (\<forall>m\<ge>p. u m \<le> u p))"
then obtain N where N: "\<And>p. p > N \<Longrightarrow> \<exists>m>p. u p < u m" by (force simp: not_le le_less)
have "\<exists>r. \<forall>n. N < r n \<and> r n < r (Suc n) \<and> (\<forall>i\<in> {N<..r (Suc n)}. u i \<le> u (r (Suc n)))"
proof (rule dependent_nat_choice)
fix x assume "N < x"
then have a: "finite {N<..x}" "{N<..x} \<noteq> {}" by simp_all
have "Max {u i |i. i \<in> {N<..x}} \<in> {u i |i. i \<in> {N<..x}}" apply (rule Max_in) using a by (auto)
then obtain p where "p \<in> {N<..x}" and upmax: "u p = Max{u i |i. i \<in> {N<..x}}" by auto
define U where "U = {m. m > p \<and> u p < u m}"
have "U \<noteq> {}" unfolding U_def using N[of p] \<open>p \<in> {N<..x}\<close> by auto
define y where "y = Inf U"
then have "y \<in> U" using \<open>U \<noteq> {}\<close> by (simp add: Inf_nat_def1)
have a: "\<And>i. i \<in> {N<..x} \<Longrightarrow> u i \<le> u p"
proof -
fix i assume "i \<in> {N<..x}"
then have "u i \<in> {u i |i. i \<in> {N<..x}}" by blast
then show "u i \<le> u p" using upmax by simp
qed
moreover have "u p < u y" using \<open>y \<in> U\<close> U_def by auto
ultimately have "y \<notin> {N<..x}" using not_le by blast
moreover have "y > N" using \<open>y \<in> U\<close> U_def \<open>p \<in> {N<..x}\<close> by auto
ultimately have "y > x" by auto
have "\<And>i. i \<in> {N<..y} \<Longrightarrow> u i \<le> u y"
proof -
fix i assume "i \<in> {N<..y}" show "u i \<le> u y"
proof (cases)
assume "i = y"
then show ?thesis by simp
next
assume "\<not>(i=y)"
then have i:"i \<in> {N<..<y}" using \<open>i \<in> {N<..y}\<close> by simp
have "u i \<le> u p"
proof (cases)
assume "i \<le> x"
then have "i \<in> {N<..x}" using i by simp
then show ?thesis using a by simp
next
assume "\<not>(i \<le> x)"
then have "i > x" by simp
then have *: "i > p" using \<open>p \<in> {N<..x}\<close> by simp
have "i < Inf U" using i y_def by simp
then have "i \<notin> U" using Inf_nat_def not_less_Least by auto
then show ?thesis using U_def * by auto
qed
then show "u i \<le> u y" using \<open>u p < u y\<close> by auto
qed
qed
then have "N < y \<and> x < y \<and> (\<forall>i\<in>{N<..y}. u i \<le> u y)" using \<open>y > x\<close> \<open>y > N\<close> by auto
then show "\<exists>y>N. x < y \<and> (\<forall>i\<in>{N<..y}. u i \<le> u y)" by auto
qed (auto)
then obtain r where r: "\<forall>n. N < r n \<and> r n < r (Suc n) \<and> (\<forall>i\<in> {N<..r (Suc n)}. u i \<le> u (r (Suc n)))" by auto
have "strict_mono r" using r by (auto simp: strict_mono_Suc_iff)
have "incseq (u o r)" unfolding o_def using r by (simp add: incseq_SucI order.strict_implies_order)
then have "(u o r) \<longlonglongrightarrow> (SUP n. (u o r) n)" using LIMSEQ_SUP by blast
then have "limsup (u o r) = (SUP n. (u o r) n)" by (simp add: lim_imp_Limsup)
moreover have "limsup (u o r) \<le> limsup u" using \<open>strict_mono r\<close> by (simp add: limsup_subseq_mono)
ultimately have "(SUP n. (u o r) n) \<le> limsup u" by simp
{
fix i assume i: "i \<in> {N<..}"
obtain n where "i < r (Suc n)" using \<open>strict_mono r\<close> using Suc_le_eq seq_suble by blast
then have "i \<in> {N<..r(Suc n)}" using i by simp
then have "u i \<le> u (r(Suc n))" using r by simp
then have "u i \<le> (SUP n. (u o r) n)" unfolding o_def by (meson SUP_upper2 UNIV_I)
}
then have "(SUP i\<in>{N<..}. u i) \<le> (SUP n. (u o r) n)" using SUP_least by blast
then have "limsup u \<le> (SUP n. (u o r) n)" unfolding Limsup_def
by (metis (mono_tags, lifting) INF_lower2 atLeast_Suc_greaterThan atLeast_def eventually_ge_at_top mem_Collect_eq)
then have "limsup u = (SUP n. (u o r) n)" using \<open>(SUP n. (u o r) n) \<le> limsup u\<close> by simp
then have "(u o r) \<longlonglongrightarrow> limsup u" using \<open>(u o r) \<longlonglongrightarrow> (SUP n. (u o r) n)\<close> by simp
then show ?thesis using \<open>strict_mono r\<close> by auto
qed
lemma liminf_subseq_lim:
fixes u::"nat \<Rightarrow> 'a :: {complete_linorder, linorder_topology}"
shows "\<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> (u o r) \<longlonglongrightarrow> liminf u"
proof (cases)
assume "\<forall>n. \<exists>p>n. \<forall>m\<ge>p. u m \<ge> u p"
then have "\<exists>r. \<forall>n. (\<forall>m\<ge>r n. u m \<ge> u (r n)) \<and> r n < r (Suc n)"
by (intro dependent_nat_choice) (auto simp: conj_commute)
then obtain r :: "nat \<Rightarrow> nat" where "strict_mono r" and mono: "\<And>n m. r n \<le> m \<Longrightarrow> u m \<ge> u (r n)"
by (auto simp: strict_mono_Suc_iff)
define umin where "umin = (\<lambda>n. (INF m\<in>{n..}. u m))"
have "incseq umin" unfolding umin_def by (simp add: INF_superset_mono incseq_def)
then have "umin \<longlonglongrightarrow> liminf u" unfolding umin_def by (metis LIMSEQ_SUP liminf_SUP_INF)
then have *: "(umin o r) \<longlonglongrightarrow> liminf u" by (simp add: LIMSEQ_subseq_LIMSEQ \<open>strict_mono r\<close>)
have "\<And>n. umin(r n) = u(r n)" unfolding umin_def using mono
by (metis le_INF_iff antisym atLeast_def mem_Collect_eq order_refl)
then have "umin o r = u o r" unfolding o_def by simp
then have "(u o r) \<longlonglongrightarrow> liminf u" using * by simp
then show ?thesis using \<open>strict_mono r\<close> by blast
next
assume "\<not> (\<forall>n. \<exists>p>n. (\<forall>m\<ge>p. u m \<ge> u p))"
then obtain N where N: "\<And>p. p > N \<Longrightarrow> \<exists>m>p. u p > u m" by (force simp: not_le le_less)
have "\<exists>r. \<forall>n. N < r n \<and> r n < r (Suc n) \<and> (\<forall>i\<in> {N<..r (Suc n)}. u i \<ge> u (r (Suc n)))"
proof (rule dependent_nat_choice)
fix x assume "N < x"
then have a: "finite {N<..x}" "{N<..x} \<noteq> {}" by simp_all
have "Min {u i |i. i \<in> {N<..x}} \<in> {u i |i. i \<in> {N<..x}}" apply (rule Min_in) using a by (auto)
then obtain p where "p \<in> {N<..x}" and upmin: "u p = Min{u i |i. i \<in> {N<..x}}" by auto
define U where "U = {m. m > p \<and> u p > u m}"
have "U \<noteq> {}" unfolding U_def using N[of p] \<open>p \<in> {N<..x}\<close> by auto
define y where "y = Inf U"
then have "y \<in> U" using \<open>U \<noteq> {}\<close> by (simp add: Inf_nat_def1)
have a: "\<And>i. i \<in> {N<..x} \<Longrightarrow> u i \<ge> u p"
proof -
fix i assume "i \<in> {N<..x}"
then have "u i \<in> {u i |i. i \<in> {N<..x}}" by blast
then show "u i \<ge> u p" using upmin by simp
qed
moreover have "u p > u y" using \<open>y \<in> U\<close> U_def by auto
ultimately have "y \<notin> {N<..x}" using not_le by blast
moreover have "y > N" using \<open>y \<in> U\<close> U_def \<open>p \<in> {N<..x}\<close> by auto
ultimately have "y > x" by auto
have "\<And>i. i \<in> {N<..y} \<Longrightarrow> u i \<ge> u y"
proof -
fix i assume "i \<in> {N<..y}" show "u i \<ge> u y"
proof (cases)
assume "i = y"
then show ?thesis by simp
next
assume "\<not>(i=y)"
then have i:"i \<in> {N<..<y}" using \<open>i \<in> {N<..y}\<close> by simp
have "u i \<ge> u p"
proof (cases)
assume "i \<le> x"
then have "i \<in> {N<..x}" using i by simp
then show ?thesis using a by simp
next
assume "\<not>(i \<le> x)"
then have "i > x" by simp
then have *: "i > p" using \<open>p \<in> {N<..x}\<close> by simp
have "i < Inf U" using i y_def by simp
then have "i \<notin> U" using Inf_nat_def not_less_Least by auto
then show ?thesis using U_def * by auto
qed
then show "u i \<ge> u y" using \<open>u p > u y\<close> by auto
qed
qed
then have "N < y \<and> x < y \<and> (\<forall>i\<in>{N<..y}. u i \<ge> u y)" using \<open>y > x\<close> \<open>y > N\<close> by auto
then show "\<exists>y>N. x < y \<and> (\<forall>i\<in>{N<..y}. u i \<ge> u y)" by auto
qed (auto)
then obtain r :: "nat \<Rightarrow> nat"
where r: "\<forall>n. N < r n \<and> r n < r (Suc n) \<and> (\<forall>i\<in> {N<..r (Suc n)}. u i \<ge> u (r (Suc n)))" by auto
have "strict_mono r" using r by (auto simp: strict_mono_Suc_iff)
have "decseq (u o r)" unfolding o_def using r by (simp add: decseq_SucI order.strict_implies_order)
then have "(u o r) \<longlonglongrightarrow> (INF n. (u o r) n)" using LIMSEQ_INF by blast
then have "liminf (u o r) = (INF n. (u o r) n)" by (simp add: lim_imp_Liminf)
moreover have "liminf (u o r) \<ge> liminf u" using \<open>strict_mono r\<close> by (simp add: liminf_subseq_mono)
ultimately have "(INF n. (u o r) n) \<ge> liminf u" by simp
{
fix i assume i: "i \<in> {N<..}"
obtain n where "i < r (Suc n)" using \<open>strict_mono r\<close> using Suc_le_eq seq_suble by blast
then have "i \<in> {N<..r(Suc n)}" using i by simp
then have "u i \<ge> u (r(Suc n))" using r by simp
then have "u i \<ge> (INF n. (u o r) n)" unfolding o_def by (meson INF_lower2 UNIV_I)
}
then have "(INF i\<in>{N<..}. u i) \<ge> (INF n. (u o r) n)" using INF_greatest by blast
then have "liminf u \<ge> (INF n. (u o r) n)" unfolding Liminf_def
by (metis (mono_tags, lifting) SUP_upper2 atLeast_Suc_greaterThan atLeast_def eventually_ge_at_top mem_Collect_eq)
then have "liminf u = (INF n. (u o r) n)" using \<open>(INF n. (u o r) n) \<ge> liminf u\<close> by simp
then have "(u o r) \<longlonglongrightarrow> liminf u" using \<open>(u o r) \<longlonglongrightarrow> (INF n. (u o r) n)\<close> by simp
then show ?thesis using \<open>strict_mono r\<close> by auto
qed
text \<open>The following statement about limsups is reduced to a statement about limits using
subsequences thanks to \<open>limsup_subseq_lim\<close>. The statement for limits follows for instance from
\<open>tendsto_add_ereal_general\<close>.\<close>
lemma ereal_limsup_add_mono:
fixes u v::"nat \<Rightarrow> ereal"
shows "limsup (\<lambda>n. u n + v n) \<le> limsup u + limsup v"
proof (cases)
assume "(limsup u = \<infinity>) \<or> (limsup v = \<infinity>)"
then have "limsup u + limsup v = \<infinity>" by simp
then show ?thesis by auto
next
assume "\<not>((limsup u = \<infinity>) \<or> (limsup v = \<infinity>))"
then have "limsup u < \<infinity>" "limsup v < \<infinity>" by auto
define w where "w = (\<lambda>n. u n + v n)"
obtain r where r: "strict_mono r" "(w o r) \<longlonglongrightarrow> limsup w" using limsup_subseq_lim by auto
obtain s where s: "strict_mono s" "(u o r o s) \<longlonglongrightarrow> limsup (u o r)" using limsup_subseq_lim by auto
obtain t where t: "strict_mono t" "(v o r o s o t) \<longlonglongrightarrow> limsup (v o r o s)" using limsup_subseq_lim by auto
define a where "a = r o s o t"
have "strict_mono a" using r s t by (simp add: a_def strict_mono_o)
have l:"(w o a) \<longlonglongrightarrow> limsup w"
"(u o a) \<longlonglongrightarrow> limsup (u o r)"
"(v o a) \<longlonglongrightarrow> limsup (v o r o s)"
apply (metis (no_types, lifting) r(2) s(1) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc)
apply (metis (no_types, lifting) s(2) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc)
apply (metis (no_types, lifting) t(2) a_def comp_assoc)
done
have "limsup (u o r) \<le> limsup u" by (simp add: limsup_subseq_mono r(1))
then have a: "limsup (u o r) \<noteq> \<infinity>" using \<open>limsup u < \<infinity>\<close> by auto
have "limsup (v o r o s) \<le> limsup v"
by (simp add: comp_assoc limsup_subseq_mono r(1) s(1) strict_mono_o)
then have b: "limsup (v o r o s) \<noteq> \<infinity>" using \<open>limsup v < \<infinity>\<close> by auto
have "(\<lambda>n. (u o a) n + (v o a) n) \<longlonglongrightarrow> limsup (u o r) + limsup (v o r o s)"
using l tendsto_add_ereal_general a b by fastforce
moreover have "(\<lambda>n. (u o a) n + (v o a) n) = (w o a)" unfolding w_def by auto
ultimately have "(w o a) \<longlonglongrightarrow> limsup (u o r) + limsup (v o r o s)" by simp
then have "limsup w = limsup (u o r) + limsup (v o r o s)" using l(1) LIMSEQ_unique by blast
then have "limsup w \<le> limsup u + limsup v"
using \<open>limsup (u o r) \<le> limsup u\<close> \<open>limsup (v o r o s) \<le> limsup v\<close> add_mono by simp
then show ?thesis unfolding w_def by simp
qed
text \<open>There is an asymmetry between liminfs and limsups in \<open>ereal\<close>, as \<open>\<infinity> + (-\<infinity>) = \<infinity>\<close>.
This explains why there are more assumptions in the next lemma dealing with liminfs that in the
previous one about limsups.\<close>
lemma ereal_liminf_add_mono:
fixes u v::"nat \<Rightarrow> ereal"
assumes "\<not>((liminf u = \<infinity> \<and> liminf v = -\<infinity>) \<or> (liminf u = -\<infinity> \<and> liminf v = \<infinity>))"
shows "liminf (\<lambda>n. u n + v n) \<ge> liminf u + liminf v"
proof (cases)
assume "(liminf u = -\<infinity>) \<or> (liminf v = -\<infinity>)"
then have *: "liminf u + liminf v = -\<infinity>" using assms by auto
show ?thesis by (simp add: *)
next
assume "\<not>((liminf u = -\<infinity>) \<or> (liminf v = -\<infinity>))"
then have "liminf u > -\<infinity>" "liminf v > -\<infinity>" by auto
define w where "w = (\<lambda>n. u n + v n)"
obtain r where r: "strict_mono r" "(w o r) \<longlonglongrightarrow> liminf w" using liminf_subseq_lim by auto
obtain s where s: "strict_mono s" "(u o r o s) \<longlonglongrightarrow> liminf (u o r)" using liminf_subseq_lim by auto
obtain t where t: "strict_mono t" "(v o r o s o t) \<longlonglongrightarrow> liminf (v o r o s)" using liminf_subseq_lim by auto
define a where "a = r o s o t"
have "strict_mono a" using r s t by (simp add: a_def strict_mono_o)
have l:"(w o a) \<longlonglongrightarrow> liminf w"
"(u o a) \<longlonglongrightarrow> liminf (u o r)"
"(v o a) \<longlonglongrightarrow> liminf (v o r o s)"
apply (metis (no_types, lifting) r(2) s(1) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc)
apply (metis (no_types, lifting) s(2) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc)
apply (metis (no_types, lifting) t(2) a_def comp_assoc)
done
have "liminf (u o r) \<ge> liminf u" by (simp add: liminf_subseq_mono r(1))
then have a: "liminf (u o r) \<noteq> -\<infinity>" using \<open>liminf u > -\<infinity>\<close> by auto
have "liminf (v o r o s) \<ge> liminf v"
by (simp add: comp_assoc liminf_subseq_mono r(1) s(1) strict_mono_o)
then have b: "liminf (v o r o s) \<noteq> -\<infinity>" using \<open>liminf v > -\<infinity>\<close> by auto
have "(\<lambda>n. (u o a) n + (v o a) n) \<longlonglongrightarrow> liminf (u o r) + liminf (v o r o s)"
using l tendsto_add_ereal_general a b by fastforce
moreover have "(\<lambda>n. (u o a) n + (v o a) n) = (w o a)" unfolding w_def by auto
ultimately have "(w o a) \<longlonglongrightarrow> liminf (u o r) + liminf (v o r o s)" by simp
then have "liminf w = liminf (u o r) + liminf (v o r o s)" using l(1) LIMSEQ_unique by blast
then have "liminf w \<ge> liminf u + liminf v"
using \<open>liminf (u o r) \<ge> liminf u\<close> \<open>liminf (v o r o s) \<ge> liminf v\<close> add_mono by simp
then show ?thesis unfolding w_def by simp
qed
lemma ereal_limsup_lim_add:
fixes u v::"nat \<Rightarrow> ereal"
assumes "u \<longlonglongrightarrow> a" "abs(a) \<noteq> \<infinity>"
shows "limsup (\<lambda>n. u n + v n) = a + limsup v"
proof -
have "limsup u = a" using assms(1) using tendsto_iff_Liminf_eq_Limsup trivial_limit_at_top_linorder by blast
have "(\<lambda>n. -u n) \<longlonglongrightarrow> -a" using assms(1) by auto
then have "limsup (\<lambda>n. -u n) = -a" using tendsto_iff_Liminf_eq_Limsup trivial_limit_at_top_linorder by blast
have "limsup (\<lambda>n. u n + v n) \<le> limsup u + limsup v"
by (rule ereal_limsup_add_mono)
then have up: "limsup (\<lambda>n. u n + v n) \<le> a + limsup v" using \<open>limsup u = a\<close> by simp
have a: "limsup (\<lambda>n. (u n + v n) + (-u n)) \<le> limsup (\<lambda>n. u n + v n) + limsup (\<lambda>n. -u n)"
by (rule ereal_limsup_add_mono)
have "eventually (\<lambda>n. u n = ereal(real_of_ereal(u n))) sequentially" using assms
real_lim_then_eventually_real by auto
moreover have "\<And>x. x = ereal(real_of_ereal(x)) \<Longrightarrow> x + (-x) = 0"
by (metis plus_ereal.simps(1) right_minus uminus_ereal.simps(1) zero_ereal_def)
ultimately have "eventually (\<lambda>n. u n + (-u n) = 0) sequentially"
by (metis (mono_tags, lifting) eventually_mono)
moreover have "\<And>n. u n + (-u n) = 0 \<Longrightarrow> u n + v n + (-u n) = v n"
by (metis add.commute add.left_commute add.left_neutral)
ultimately have "eventually (\<lambda>n. u n + v n + (-u n) = v n) sequentially"
using eventually_mono by force
then have "limsup v = limsup (\<lambda>n. u n + v n + (-u n))" using Limsup_eq by force
then have "limsup v \<le> limsup (\<lambda>n. u n + v n) -a" using a \<open>limsup (\<lambda>n. -u n) = -a\<close> by (simp add: minus_ereal_def)
then have "limsup (\<lambda>n. u n + v n) \<ge> a + limsup v" using assms(2) by (metis add.commute ereal_le_minus)
then show ?thesis using up by simp
qed
lemma ereal_limsup_lim_mult:
fixes u v::"nat \<Rightarrow> ereal"
assumes "u \<longlonglongrightarrow> a" "a>0" "a \<noteq> \<infinity>"
shows "limsup (\<lambda>n. u n * v n) = a * limsup v"
proof -
define w where "w = (\<lambda>n. u n * v n)"
obtain r where r: "strict_mono r" "(v o r) \<longlonglongrightarrow> limsup v" using limsup_subseq_lim by auto
have "(u o r) \<longlonglongrightarrow> a" using assms(1) LIMSEQ_subseq_LIMSEQ r by auto
with tendsto_mult_ereal[OF this r(2)] have "(\<lambda>n. (u o r) n * (v o r) n) \<longlonglongrightarrow> a * limsup v" using assms(2) assms(3) by auto
moreover have "\<And>n. (w o r) n = (u o r) n * (v o r) n" unfolding w_def by auto
ultimately have "(w o r) \<longlonglongrightarrow> a * limsup v" unfolding w_def by presburger
then have "limsup (w o r) = a * limsup v" by (simp add: tendsto_iff_Liminf_eq_Limsup)
then have I: "limsup w \<ge> a * limsup v" by (metis limsup_subseq_mono r(1))
obtain s where s: "strict_mono s" "(w o s) \<longlonglongrightarrow> limsup w" using limsup_subseq_lim by auto
have *: "(u o s) \<longlonglongrightarrow> a" using assms(1) LIMSEQ_subseq_LIMSEQ s by auto
have "eventually (\<lambda>n. (u o s) n > 0) sequentially" using assms(2) * order_tendsto_iff by blast
moreover have "eventually (\<lambda>n. (u o s) n < \<infinity>) sequentially" using assms(3) * order_tendsto_iff by blast
moreover have "(w o s) n / (u o s) n = (v o s) n" if "(u o s) n > 0" "(u o s) n < \<infinity>" for n
unfolding w_def using that by (auto simp add: ereal_divide_eq)
ultimately have "eventually (\<lambda>n. (w o s) n / (u o s) n = (v o s) n) sequentially" using eventually_elim2 by force
moreover have "(\<lambda>n. (w o s) n / (u o s) n) \<longlonglongrightarrow> (limsup w) / a"
apply (rule tendsto_divide_ereal[OF s(2) *]) using assms(2) assms(3) by auto
ultimately have "(v o s) \<longlonglongrightarrow> (limsup w) / a" using Lim_transform_eventually by fastforce
then have "limsup (v o s) = (limsup w) / a" by (simp add: tendsto_iff_Liminf_eq_Limsup)
then have "limsup v \<ge> (limsup w) / a" by (metis limsup_subseq_mono s(1))
then have "a * limsup v \<ge> limsup w" using assms(2) assms(3) by (simp add: ereal_divide_le_pos)
then show ?thesis using I unfolding w_def by auto
qed
lemma ereal_liminf_lim_mult:
fixes u v::"nat \<Rightarrow> ereal"
assumes "u \<longlonglongrightarrow> a" "a>0" "a \<noteq> \<infinity>"
shows "liminf (\<lambda>n. u n * v n) = a * liminf v"
proof -
define w where "w = (\<lambda>n. u n * v n)"
obtain r where r: "strict_mono r" "(v o r) \<longlonglongrightarrow> liminf v" using liminf_subseq_lim by auto
have "(u o r) \<longlonglongrightarrow> a" using assms(1) LIMSEQ_subseq_LIMSEQ r by auto
with tendsto_mult_ereal[OF this r(2)] have "(\<lambda>n. (u o r) n * (v o r) n) \<longlonglongrightarrow> a * liminf v" using assms(2) assms(3) by auto
moreover have "\<And>n. (w o r) n = (u o r) n * (v o r) n" unfolding w_def by auto
ultimately have "(w o r) \<longlonglongrightarrow> a * liminf v" unfolding w_def by presburger
then have "liminf (w o r) = a * liminf v" by (simp add: tendsto_iff_Liminf_eq_Limsup)
then have I: "liminf w \<le> a * liminf v" by (metis liminf_subseq_mono r(1))
obtain s where s: "strict_mono s" "(w o s) \<longlonglongrightarrow> liminf w" using liminf_subseq_lim by auto
have *: "(u o s) \<longlonglongrightarrow> a" using assms(1) LIMSEQ_subseq_LIMSEQ s by auto
have "eventually (\<lambda>n. (u o s) n > 0) sequentially" using assms(2) * order_tendsto_iff by blast
moreover have "eventually (\<lambda>n. (u o s) n < \<infinity>) sequentially" using assms(3) * order_tendsto_iff by blast
moreover have "(w o s) n / (u o s) n = (v o s) n" if "(u o s) n > 0" "(u o s) n < \<infinity>" for n
unfolding w_def using that by (auto simp add: ereal_divide_eq)
ultimately have "eventually (\<lambda>n. (w o s) n / (u o s) n = (v o s) n) sequentially" using eventually_elim2 by force
moreover have "(\<lambda>n. (w o s) n / (u o s) n) \<longlonglongrightarrow> (liminf w) / a"
apply (rule tendsto_divide_ereal[OF s(2) *]) using assms(2) assms(3) by auto
ultimately have "(v o s) \<longlonglongrightarrow> (liminf w) / a" using Lim_transform_eventually by fastforce
then have "liminf (v o s) = (liminf w) / a" by (simp add: tendsto_iff_Liminf_eq_Limsup)
then have "liminf v \<le> (liminf w) / a" by (metis liminf_subseq_mono s(1))
then have "a * liminf v \<le> liminf w" using assms(2) assms(3) by (simp add: ereal_le_divide_pos)
then show ?thesis using I unfolding w_def by auto
qed
lemma ereal_liminf_lim_add:
fixes u v::"nat \<Rightarrow> ereal"
assumes "u \<longlonglongrightarrow> a" "abs(a) \<noteq> \<infinity>"
shows "liminf (\<lambda>n. u n + v n) = a + liminf v"
proof -
have "liminf u = a" using assms(1) tendsto_iff_Liminf_eq_Limsup trivial_limit_at_top_linorder by blast
then have *: "abs(liminf u) \<noteq> \<infinity>" using assms(2) by auto
have "(\<lambda>n. -u n) \<longlonglongrightarrow> -a" using assms(1) by auto
then have "liminf (\<lambda>n. -u n) = -a" using tendsto_iff_Liminf_eq_Limsup trivial_limit_at_top_linorder by blast
then have **: "abs(liminf (\<lambda>n. -u n)) \<noteq> \<infinity>" using assms(2) by auto
have "liminf (\<lambda>n. u n + v n) \<ge> liminf u + liminf v"
apply (rule ereal_liminf_add_mono) using * by auto
then have up: "liminf (\<lambda>n. u n + v n) \<ge> a + liminf v" using \<open>liminf u = a\<close> by simp
have a: "liminf (\<lambda>n. (u n + v n) + (-u n)) \<ge> liminf (\<lambda>n. u n + v n) + liminf (\<lambda>n. -u n)"
apply (rule ereal_liminf_add_mono) using ** by auto
have "eventually (\<lambda>n. u n = ereal(real_of_ereal(u n))) sequentially" using assms
real_lim_then_eventually_real by auto
moreover have "\<And>x. x = ereal(real_of_ereal(x)) \<Longrightarrow> x + (-x) = 0"
by (metis plus_ereal.simps(1) right_minus uminus_ereal.simps(1) zero_ereal_def)
ultimately have "eventually (\<lambda>n. u n + (-u n) = 0) sequentially"
by (metis (mono_tags, lifting) eventually_mono)
moreover have "\<And>n. u n + (-u n) = 0 \<Longrightarrow> u n + v n + (-u n) = v n"
by (metis add.commute add.left_commute add.left_neutral)
ultimately have "eventually (\<lambda>n. u n + v n + (-u n) = v n) sequentially"
using eventually_mono by force
then have "liminf v = liminf (\<lambda>n. u n + v n + (-u n))" using Liminf_eq by force
then have "liminf v \<ge> liminf (\<lambda>n. u n + v n) -a" using a \<open>liminf (\<lambda>n. -u n) = -a\<close> by (simp add: minus_ereal_def)
then have "liminf (\<lambda>n. u n + v n) \<le> a + liminf v" using assms(2) by (metis add.commute ereal_minus_le)
then show ?thesis using up by simp
qed
lemma ereal_liminf_limsup_add:
fixes u v::"nat \<Rightarrow> ereal"
shows "liminf (\<lambda>n. u n + v n) \<le> liminf u + limsup v"
proof (cases)
assume "limsup v = \<infinity> \<or> liminf u = \<infinity>"
then show ?thesis by auto
next
assume "\<not>(limsup v = \<infinity> \<or> liminf u = \<infinity>)"
then have "limsup v < \<infinity>" "liminf u < \<infinity>" by auto
define w where "w = (\<lambda>n. u n + v n)"
obtain r where r: "strict_mono r" "(u o r) \<longlonglongrightarrow> liminf u" using liminf_subseq_lim by auto
obtain s where s: "strict_mono s" "(w o r o s) \<longlonglongrightarrow> liminf (w o r)" using liminf_subseq_lim by auto
obtain t where t: "strict_mono t" "(v o r o s o t) \<longlonglongrightarrow> limsup (v o r o s)" using limsup_subseq_lim by auto
define a where "a = r o s o t"
have "strict_mono a" using r s t by (simp add: a_def strict_mono_o)
have l:"(u o a) \<longlonglongrightarrow> liminf u"
"(w o a) \<longlonglongrightarrow> liminf (w o r)"
"(v o a) \<longlonglongrightarrow> limsup (v o r o s)"
apply (metis (no_types, lifting) r(2) s(1) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc)
apply (metis (no_types, lifting) s(2) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc)
apply (metis (no_types, lifting) t(2) a_def comp_assoc)
done
have "liminf (w o r) \<ge> liminf w" by (simp add: liminf_subseq_mono r(1))
have "limsup (v o r o s) \<le> limsup v"
by (simp add: comp_assoc limsup_subseq_mono r(1) s(1) strict_mono_o)
then have b: "limsup (v o r o s) < \<infinity>" using \<open>limsup v < \<infinity>\<close> by auto
have "(\<lambda>n. (u o a) n + (v o a) n) \<longlonglongrightarrow> liminf u + limsup (v o r o s)"
apply (rule tendsto_add_ereal_general) using b \<open>liminf u < \<infinity>\<close> l(1) l(3) by force+
moreover have "(\<lambda>n. (u o a) n + (v o a) n) = (w o a)" unfolding w_def by auto
ultimately have "(w o a) \<longlonglongrightarrow> liminf u + limsup (v o r o s)" by simp
then have "liminf (w o r) = liminf u + limsup (v o r o s)" using l(2) using LIMSEQ_unique by blast
then have "liminf w \<le> liminf u + limsup v"
using \<open>liminf (w o r) \<ge> liminf w\<close> \<open>limsup (v o r o s) \<le> limsup v\<close>
by (metis add_mono_thms_linordered_semiring(2) le_less_trans not_less)
then show ?thesis unfolding w_def by simp
qed
lemma ereal_liminf_limsup_minus:
fixes u v::"nat \<Rightarrow> ereal"
shows "liminf (\<lambda>n. u n - v n) \<le> limsup u - limsup v"
unfolding minus_ereal_def
apply (subst add.commute)
apply (rule order_trans[OF ereal_liminf_limsup_add])
using ereal_Limsup_uminus[of sequentially "\<lambda>n. - v n"]
apply (simp add: add.commute)
done
lemma liminf_minus_ennreal:
fixes u v::"nat \<Rightarrow> ennreal"
shows "(\<And>n. v n \<le> u n) \<Longrightarrow> liminf (\<lambda>n. u n - v n) \<le> limsup u - limsup v"
unfolding liminf_SUP_INF limsup_INF_SUP
including ennreal.lifting
proof (transfer, clarsimp)
fix v u :: "nat \<Rightarrow> ereal" assume *: "\<forall>x. 0 \<le> v x" "\<forall>x. 0 \<le> u x" "\<And>n. v n \<le> u n"
moreover have "0 \<le> limsup u - limsup v"
using * by (intro ereal_diff_positive Limsup_mono always_eventually) simp
moreover have "0 \<le> Sup (u ` {x..})" for x
using * by (intro SUP_upper2[of x]) auto
moreover have "0 \<le> Sup (v ` {x..})" for x
using * by (intro SUP_upper2[of x]) auto
ultimately show "(SUP n. INF n\<in>{n..}. max 0 (u n - v n))
\<le> max 0 ((INF x. max 0 (Sup (u ` {x..}))) - (INF x. max 0 (Sup (v ` {x..}))))"
by (auto simp: * ereal_diff_positive max.absorb2 liminf_SUP_INF[symmetric] limsup_INF_SUP[symmetric] ereal_liminf_limsup_minus)
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
diff --git a/src/HOL/Analysis/FPS_Convergence.thy b/src/HOL/Analysis/FPS_Convergence.thy
--- a/src/HOL/Analysis/FPS_Convergence.thy
+++ b/src/HOL/Analysis/FPS_Convergence.thy
@@ -1,1056 +1,1056 @@
(*
Title: HOL/Analysis/FPS_Convergence.thy
Author: Manuel Eberl, TU München
Connection of formal power series and actual convergent power series on Banach spaces
(most notably the complex numbers).
*)
section \<open>Convergence of Formal Power Series\<close>
theory FPS_Convergence
imports
Generalised_Binomial_Theorem
"HOL-Computational_Algebra.Formal_Power_Series"
begin
text \<open>
In this theory, we will connect formal power series (which are algebraic objects) with analytic
functions. This will become more important in complex analysis, and indeed some of the less
trivial results will only be proven there.
\<close>
subsection\<^marker>\<open>tag unimportant\<close> \<open>Balls with extended real radius\<close>
(* TODO: This should probably go somewhere else *)
text \<open>
The following is a variant of \<^const>\<open>ball\<close> that also allows an infinite radius.
\<close>
definition eball :: "'a :: metric_space \<Rightarrow> ereal \<Rightarrow> 'a set" where
"eball z r = {z'. ereal (dist z z') < r}"
lemma in_eball_iff [simp]: "z \<in> eball z0 r \<longleftrightarrow> ereal (dist z0 z) < r"
by (simp add: eball_def)
lemma eball_ereal [simp]: "eball z (ereal r) = ball z r"
by auto
lemma eball_inf [simp]: "eball z \<infinity> = UNIV"
by auto
lemma eball_empty [simp]: "r \<le> 0 \<Longrightarrow> eball z r = {}"
proof safe
fix x assume "r \<le> 0" "x \<in> eball z r"
hence "dist z x < r" by simp
also have "\<dots> \<le> ereal 0" using \<open>r \<le> 0\<close> by (simp add: zero_ereal_def)
finally show "x \<in> {}" by simp
qed
lemma eball_conv_UNION_balls:
"eball z r = (\<Union>r'\<in>{r'. ereal r' < r}. ball z r')"
by (cases r) (use dense gt_ex in force)+
lemma eball_mono: "r \<le> r' \<Longrightarrow> eball z r \<le> eball z r'"
by auto
lemma ball_eball_mono: "ereal r \<le> r' \<Longrightarrow> ball z r \<le> eball z r'"
using eball_mono[of "ereal r" r'] by simp
lemma open_eball [simp, intro]: "open (eball z r)"
by (cases r) auto
subsection \<open>Basic properties of convergent power series\<close>
definition\<^marker>\<open>tag important\<close> fps_conv_radius :: "'a :: {banach, real_normed_div_algebra} fps \<Rightarrow> ereal" where
"fps_conv_radius f = conv_radius (fps_nth f)"
definition\<^marker>\<open>tag important\<close> eval_fps :: "'a :: {banach, real_normed_div_algebra} fps \<Rightarrow> 'a \<Rightarrow> 'a" where
"eval_fps f z = (\<Sum>n. fps_nth f n * z ^ n)"
lemma norm_summable_fps:
fixes f :: "'a :: {banach, real_normed_div_algebra} fps"
shows "norm z < fps_conv_radius f \<Longrightarrow> summable (\<lambda>n. norm (fps_nth f n * z ^ n))"
by (rule abs_summable_in_conv_radius) (simp_all add: fps_conv_radius_def)
lemma summable_fps:
fixes f :: "'a :: {banach, real_normed_div_algebra} fps"
shows "norm z < fps_conv_radius f \<Longrightarrow> summable (\<lambda>n. fps_nth f n * z ^ n)"
by (rule summable_in_conv_radius) (simp_all add: fps_conv_radius_def)
theorem sums_eval_fps:
fixes f :: "'a :: {banach, real_normed_div_algebra} fps"
assumes "norm z < fps_conv_radius f"
shows "(\<lambda>n. fps_nth f n * z ^ n) sums eval_fps f z"
using assms unfolding eval_fps_def fps_conv_radius_def
by (intro summable_sums summable_in_conv_radius) simp_all
lemma continuous_on_eval_fps:
fixes f :: "'a :: {banach, real_normed_div_algebra} fps"
shows "continuous_on (eball 0 (fps_conv_radius f)) (eval_fps f)"
proof (subst continuous_on_eq_continuous_at [OF open_eball], safe)
fix x :: 'a assume x: "x \<in> eball 0 (fps_conv_radius f)"
define r where "r = (if fps_conv_radius f = \<infinity> then norm x + 1 else
(norm x + real_of_ereal (fps_conv_radius f)) / 2)"
have r: "norm x < r \<and> ereal r < fps_conv_radius f"
using x by (cases "fps_conv_radius f")
(auto simp: r_def eball_def split: if_splits)
have "continuous_on (cball 0 r) (\<lambda>x. \<Sum>i. fps_nth f i * (x - 0) ^ i)"
by (rule powser_continuous_suminf) (insert r, auto simp: fps_conv_radius_def)
hence "continuous_on (cball 0 r) (eval_fps f)"
by (simp add: eval_fps_def)
thus "isCont (eval_fps f) x"
by (rule continuous_on_interior) (use r in auto)
qed
lemma continuous_on_eval_fps' [continuous_intros]:
assumes "continuous_on A g"
assumes "g ` A \<subseteq> eball 0 (fps_conv_radius f)"
shows "continuous_on A (\<lambda>x. eval_fps f (g x))"
using continuous_on_compose2[OF continuous_on_eval_fps assms] .
lemma has_field_derivative_powser:
fixes z :: "'a :: {banach, real_normed_field}"
assumes "ereal (norm z) < conv_radius f"
shows "((\<lambda>z. \<Sum>n. f n * z ^ n) has_field_derivative (\<Sum>n. diffs f n * z ^ n)) (at z within A)"
proof -
define K where "K = (if conv_radius f = \<infinity> then norm z + 1
else (norm z + real_of_ereal (conv_radius f)) / 2)"
have K: "norm z < K \<and> ereal K < conv_radius f"
using assms by (cases "conv_radius f") (auto simp: K_def)
have "0 \<le> norm z" by simp
also from K have "\<dots> < K" by simp
finally have K_pos: "K > 0" by simp
have "summable (\<lambda>n. f n * of_real K ^ n)"
using K and K_pos by (intro summable_in_conv_radius) auto
moreover from K and K_pos have "norm z < norm (of_real K :: 'a)" by auto
ultimately show ?thesis
by (rule has_field_derivative_at_within [OF termdiffs_strong])
qed
lemma has_field_derivative_eval_fps:
fixes z :: "'a :: {banach, real_normed_field}"
assumes "norm z < fps_conv_radius f"
shows "(eval_fps f has_field_derivative eval_fps (fps_deriv f) z) (at z within A)"
proof -
have "(eval_fps f has_field_derivative eval_fps (Abs_fps (diffs (fps_nth f))) z) (at z within A)"
using assms unfolding eval_fps_def fps_nth_Abs_fps fps_conv_radius_def
by (intro has_field_derivative_powser) auto
also have "Abs_fps (diffs (fps_nth f)) = fps_deriv f"
by (simp add: fps_eq_iff diffs_def)
finally show ?thesis .
qed
lemma holomorphic_on_eval_fps [holomorphic_intros]:
fixes z :: "'a :: {banach, real_normed_field}"
assumes "A \<subseteq> eball 0 (fps_conv_radius f)"
shows "eval_fps f holomorphic_on A"
proof (rule holomorphic_on_subset [OF _ assms])
show "eval_fps f holomorphic_on eball 0 (fps_conv_radius f)"
proof (subst holomorphic_on_open [OF open_eball], safe, goal_cases)
case (1 x)
thus ?case
by (intro exI[of _ "eval_fps (fps_deriv f) x"]) (auto intro: has_field_derivative_eval_fps)
qed
qed
lemma analytic_on_eval_fps:
fixes z :: "'a :: {banach, real_normed_field}"
assumes "A \<subseteq> eball 0 (fps_conv_radius f)"
shows "eval_fps f analytic_on A"
proof (rule analytic_on_subset [OF _ assms])
show "eval_fps f analytic_on eball 0 (fps_conv_radius f)"
using holomorphic_on_eval_fps[of "eball 0 (fps_conv_radius f)"]
by (subst analytic_on_open) auto
qed
lemma continuous_eval_fps [continuous_intros]:
fixes z :: "'a::{real_normed_field,banach}"
assumes "norm z < fps_conv_radius F"
shows "continuous (at z within A) (eval_fps F)"
proof -
from ereal_dense2[OF assms] obtain K :: real where K: "norm z < K" "K < fps_conv_radius F"
by auto
have "0 \<le> norm z" by simp
also have "norm z < K" by fact
finally have "K > 0" .
from K and \<open>K > 0\<close> have "summable (\<lambda>n. fps_nth F n * of_real K ^ n)"
by (intro summable_fps) auto
from this have "isCont (eval_fps F) z" unfolding eval_fps_def
by (rule isCont_powser) (use K in auto)
thus "continuous (at z within A) (eval_fps F)"
by (simp add: continuous_at_imp_continuous_within)
qed
subsection\<^marker>\<open>tag unimportant\<close> \<open>Lower bounds on radius of convergence\<close>
lemma fps_conv_radius_deriv:
fixes f :: "'a :: {banach, real_normed_field} fps"
shows "fps_conv_radius (fps_deriv f) \<ge> fps_conv_radius f"
unfolding fps_conv_radius_def
proof (rule conv_radius_geI_ex)
fix r :: real assume r: "r > 0" "ereal r < conv_radius (fps_nth f)"
define K where "K = (if conv_radius (fps_nth f) = \<infinity> then r + 1
else (real_of_ereal (conv_radius (fps_nth f)) + r) / 2)"
have K: "r < K \<and> ereal K < conv_radius (fps_nth f)"
using r by (cases "conv_radius (fps_nth f)") (auto simp: K_def)
have "summable (\<lambda>n. diffs (fps_nth f) n * of_real r ^ n)"
proof (rule termdiff_converges)
fix x :: 'a assume "norm x < K"
hence "ereal (norm x) < ereal K" by simp
also have "\<dots> < conv_radius (fps_nth f)" using K by simp
finally show "summable (\<lambda>n. fps_nth f n * x ^ n)"
by (intro summable_in_conv_radius) auto
qed (insert K r, auto)
also have "\<dots> = (\<lambda>n. fps_nth (fps_deriv f) n * of_real r ^ n)"
by (simp add: fps_deriv_def diffs_def)
finally show "\<exists>z::'a. norm z = r \<and> summable (\<lambda>n. fps_nth (fps_deriv f) n * z ^ n)"
using r by (intro exI[of _ "of_real r"]) auto
qed
lemma eval_fps_at_0: "eval_fps f 0 = fps_nth f 0"
by (simp add: eval_fps_def)
lemma fps_conv_radius_norm [simp]:
"fps_conv_radius (Abs_fps (\<lambda>n. norm (fps_nth f n))) = fps_conv_radius f"
by (simp add: fps_conv_radius_def)
lemma fps_conv_radius_const [simp]: "fps_conv_radius (fps_const c) = \<infinity>"
proof -
have "fps_conv_radius (fps_const c) = conv_radius (\<lambda>_. 0 :: 'a)"
unfolding fps_conv_radius_def
by (intro conv_radius_cong eventually_mono[OF eventually_gt_at_top[of 0]]) auto
thus ?thesis by simp
qed
lemma fps_conv_radius_0 [simp]: "fps_conv_radius 0 = \<infinity>"
by (simp only: fps_const_0_eq_0 [symmetric] fps_conv_radius_const)
lemma fps_conv_radius_1 [simp]: "fps_conv_radius 1 = \<infinity>"
by (simp only: fps_const_1_eq_1 [symmetric] fps_conv_radius_const)
lemma fps_conv_radius_numeral [simp]: "fps_conv_radius (numeral n) = \<infinity>"
by (simp add: numeral_fps_const)
lemma fps_conv_radius_fps_X_power [simp]: "fps_conv_radius (fps_X ^ n) = \<infinity>"
proof -
have "fps_conv_radius (fps_X ^ n :: 'a fps) = conv_radius (\<lambda>_. 0 :: 'a)"
unfolding fps_conv_radius_def
by (intro conv_radius_cong eventually_mono[OF eventually_gt_at_top[of n]])
(auto simp: fps_X_power_iff)
thus ?thesis by simp
qed
lemma fps_conv_radius_fps_X [simp]: "fps_conv_radius fps_X = \<infinity>"
using fps_conv_radius_fps_X_power[of 1] by (simp only: power_one_right)
lemma fps_conv_radius_shift [simp]:
"fps_conv_radius (fps_shift n f) = fps_conv_radius f"
by (simp add: fps_conv_radius_def fps_shift_def conv_radius_shift)
lemma fps_conv_radius_cmult_left:
"c \<noteq> 0 \<Longrightarrow> fps_conv_radius (fps_const c * f) = fps_conv_radius f"
unfolding fps_conv_radius_def by (simp add: conv_radius_cmult_left)
lemma fps_conv_radius_cmult_right:
"c \<noteq> 0 \<Longrightarrow> fps_conv_radius (f * fps_const c) = fps_conv_radius f"
unfolding fps_conv_radius_def by (simp add: conv_radius_cmult_right)
lemma fps_conv_radius_uminus [simp]:
"fps_conv_radius (-f) = fps_conv_radius f"
using fps_conv_radius_cmult_left[of "-1" f]
by (simp flip: fps_const_neg)
lemma fps_conv_radius_add: "fps_conv_radius (f + g) \<ge> min (fps_conv_radius f) (fps_conv_radius g)"
unfolding fps_conv_radius_def using conv_radius_add_ge[of "fps_nth f" "fps_nth g"]
by simp
lemma fps_conv_radius_diff: "fps_conv_radius (f - g) \<ge> min (fps_conv_radius f) (fps_conv_radius g)"
using fps_conv_radius_add[of f "-g"] by simp
lemma fps_conv_radius_mult: "fps_conv_radius (f * g) \<ge> min (fps_conv_radius f) (fps_conv_radius g)"
using conv_radius_mult_ge[of "fps_nth f" "fps_nth g"]
by (simp add: fps_mult_nth fps_conv_radius_def atLeast0AtMost)
lemma fps_conv_radius_power: "fps_conv_radius (f ^ n) \<ge> fps_conv_radius f"
proof (induction n)
case (Suc n)
hence "fps_conv_radius f \<le> min (fps_conv_radius f) (fps_conv_radius (f ^ n))"
by simp
also have "\<dots> \<le> fps_conv_radius (f * f ^ n)"
by (rule fps_conv_radius_mult)
finally show ?case by simp
qed simp_all
context
begin
lemma natfun_inverse_bound:
fixes f :: "'a :: {real_normed_field} fps"
assumes "fps_nth f 0 = 1" and "\<delta> > 0"
and summable: "summable (\<lambda>n. norm (fps_nth f (Suc n)) * \<delta> ^ Suc n)"
and le: "(\<Sum>n. norm (fps_nth f (Suc n)) * \<delta> ^ Suc n) \<le> 1"
shows "norm (natfun_inverse f n) \<le> inverse (\<delta> ^ n)"
proof (induction n rule: less_induct)
case (less m)
show ?case
proof (cases m)
case 0
thus ?thesis using assms by (simp add: field_split_simps norm_inverse norm_divide)
next
case [simp]: (Suc n)
have "norm (natfun_inverse f (Suc n)) =
norm (\<Sum>i = Suc 0..Suc n. fps_nth f i * natfun_inverse f (Suc n - i))"
(is "_ = norm ?S") using assms
by (simp add: field_simps norm_mult norm_divide del: sum.cl_ivl_Suc)
also have "norm ?S \<le> (\<Sum>i = Suc 0..Suc n. norm (fps_nth f i * natfun_inverse f (Suc n - i)))"
by (rule norm_sum)
also have "\<dots> \<le> (\<Sum>i = Suc 0..Suc n. norm (fps_nth f i) / \<delta> ^ (Suc n - i))"
proof (intro sum_mono, goal_cases)
case (1 i)
have "norm (fps_nth f i * natfun_inverse f (Suc n - i)) =
norm (fps_nth f i) * norm (natfun_inverse f (Suc n - i))"
by (simp add: norm_mult)
also have "\<dots> \<le> norm (fps_nth f i) * inverse (\<delta> ^ (Suc n - i))"
using 1 by (intro mult_left_mono less.IH) auto
also have "\<dots> = norm (fps_nth f i) / \<delta> ^ (Suc n - i)"
by (simp add: field_split_simps)
finally show ?case .
qed
also have "\<dots> = (\<Sum>i = Suc 0..Suc n. norm (fps_nth f i) * \<delta> ^ i) / \<delta> ^ Suc n"
by (subst sum_divide_distrib, rule sum.cong)
(insert \<open>\<delta> > 0\<close>, auto simp: field_simps power_diff)
also have "(\<Sum>i = Suc 0..Suc n. norm (fps_nth f i) * \<delta> ^ i) =
(\<Sum>i=0..n. norm (fps_nth f (Suc i)) * \<delta> ^ (Suc i))"
by (subst sum.atLeast_Suc_atMost_Suc_shift) simp_all
also have "{0..n} = {..<Suc n}" by auto
also have "(\<Sum>i< Suc n. norm (fps_nth f (Suc i)) * \<delta> ^ (Suc i)) \<le>
(\<Sum>n. norm (fps_nth f (Suc n)) * \<delta> ^ (Suc n))"
using \<open>\<delta> > 0\<close> by (intro sum_le_suminf ballI mult_nonneg_nonneg zero_le_power summable) auto
also have "\<dots> \<le> 1" by fact
finally show ?thesis using \<open>\<delta> > 0\<close>
by (simp add: divide_right_mono field_split_simps)
qed
qed
private lemma fps_conv_radius_inverse_pos_aux:
fixes f :: "'a :: {banach, real_normed_field} fps"
assumes "fps_nth f 0 = 1" "fps_conv_radius f > 0"
shows "fps_conv_radius (inverse f) > 0"
proof -
let ?R = "fps_conv_radius f"
define h where "h = Abs_fps (\<lambda>n. norm (fps_nth f n))"
have [simp]: "fps_conv_radius h = ?R" by (simp add: h_def)
have "continuous_on (eball 0 (fps_conv_radius h)) (eval_fps h)"
by (intro continuous_on_eval_fps)
hence *: "open (eval_fps h -` A \<inter> eball 0 ?R)" if "open A" for A
using that by (subst (asm) continuous_on_open_vimage) auto
have "open (eval_fps h -` {..<2} \<inter> eball 0 ?R)"
by (rule *) auto
moreover have "0 \<in> eval_fps h -` {..<2} \<inter> eball 0 ?R"
using assms by (auto simp: eball_def zero_ereal_def eval_fps_at_0 h_def)
ultimately obtain \<epsilon> where \<epsilon>: "\<epsilon> > 0" "ball 0 \<epsilon> \<subseteq> eval_fps h -` {..<2} \<inter> eball 0 ?R"
by (subst (asm) open_contains_ball_eq) blast+
define \<delta> where "\<delta> = real_of_ereal (min (ereal \<epsilon> / 2) (?R / 2))"
have \<delta>: "0 < \<delta> \<and> \<delta> < \<epsilon> \<and> ereal \<delta> < ?R"
using \<open>\<epsilon> > 0\<close> and assms by (cases ?R) (auto simp: \<delta>_def min_def)
have summable: "summable (\<lambda>n. norm (fps_nth f n) * \<delta> ^ n)"
using \<delta> by (intro summable_in_conv_radius) (simp_all add: fps_conv_radius_def)
hence "(\<lambda>n. norm (fps_nth f n) * \<delta> ^ n) sums eval_fps h \<delta>"
by (simp add: eval_fps_def summable_sums h_def)
hence "(\<lambda>n. norm (fps_nth f (Suc n)) * \<delta> ^ Suc n) sums (eval_fps h \<delta> - 1)"
by (subst sums_Suc_iff) (auto simp: assms)
moreover {
from \<delta> have "\<delta> \<in> ball 0 \<epsilon>" by auto
also have "\<dots> \<subseteq> eval_fps h -` {..<2} \<inter> eball 0 ?R" by fact
finally have "eval_fps h \<delta> < 2" by simp
}
ultimately have le: "(\<Sum>n. norm (fps_nth f (Suc n)) * \<delta> ^ Suc n) \<le> 1"
by (simp add: sums_iff)
from summable have summable: "summable (\<lambda>n. norm (fps_nth f (Suc n)) * \<delta> ^ Suc n)"
by (subst summable_Suc_iff)
have "0 < \<delta>" using \<delta> by blast
also have "\<delta> = inverse (limsup (\<lambda>n. ereal (inverse \<delta>)))"
using \<delta> by (subst Limsup_const) auto
also have "\<dots> \<le> conv_radius (natfun_inverse f)"
unfolding conv_radius_def
proof (intro ereal_inverse_antimono Limsup_mono
eventually_mono[OF eventually_gt_at_top[of 0]])
fix n :: nat assume n: "n > 0"
have "root n (norm (natfun_inverse f n)) \<le> root n (inverse (\<delta> ^ n))"
using n assms \<delta> le summable
by (intro real_root_le_mono natfun_inverse_bound) auto
also have "\<dots> = inverse \<delta>"
using n \<delta> by (simp add: power_inverse [symmetric] real_root_pos2)
finally show "ereal (inverse \<delta>) \<ge> ereal (root n (norm (natfun_inverse f n)))"
by (subst ereal_less_eq)
next
have "0 = limsup (\<lambda>n. 0::ereal)"
by (rule Limsup_const [symmetric]) auto
also have "\<dots> \<le> limsup (\<lambda>n. ereal (root n (norm (natfun_inverse f n))))"
by (intro Limsup_mono) (auto simp: real_root_ge_zero)
finally show "0 \<le> \<dots>" by simp
qed
also have "\<dots> = fps_conv_radius (inverse f)"
using assms by (simp add: fps_conv_radius_def fps_inverse_def)
finally show ?thesis by (simp add: zero_ereal_def)
qed
lemma fps_conv_radius_inverse_pos:
fixes f :: "'a :: {banach, real_normed_field} fps"
assumes "fps_nth f 0 \<noteq> 0" and "fps_conv_radius f > 0"
shows "fps_conv_radius (inverse f) > 0"
proof -
let ?c = "fps_nth f 0"
have "fps_conv_radius (inverse f) = fps_conv_radius (fps_const ?c * inverse f)"
using assms by (subst fps_conv_radius_cmult_left) auto
also have "fps_const ?c * inverse f = inverse (fps_const (inverse ?c) * f)"
using assms by (simp add: fps_inverse_mult fps_const_inverse)
also have "fps_conv_radius \<dots> > 0" using assms
by (intro fps_conv_radius_inverse_pos_aux)
(auto simp: fps_conv_radius_cmult_left)
finally show ?thesis .
qed
end
lemma fps_conv_radius_exp [simp]:
fixes c :: "'a :: {banach, real_normed_field}"
shows "fps_conv_radius (fps_exp c) = \<infinity>"
unfolding fps_conv_radius_def
proof (rule conv_radius_inftyI'')
fix z :: 'a
have "(\<lambda>n. norm (c * z) ^ n /\<^sub>R fact n) sums exp (norm (c * z))"
by (rule exp_converges)
also have "(\<lambda>n. norm (c * z) ^ n /\<^sub>R fact n) = (\<lambda>n. norm (fps_nth (fps_exp c) n * z ^ n))"
- by (rule ext) (simp add: norm_divide norm_mult norm_power field_split_simps power_mult_distrib)
+ by (rule ext) (simp add: norm_divide norm_mult norm_power field_split_simps)
finally have "summable \<dots>" by (simp add: sums_iff)
thus "summable (\<lambda>n. fps_nth (fps_exp c) n * z ^ n)"
by (rule summable_norm_cancel)
qed
subsection \<open>Evaluating power series\<close>
theorem eval_fps_deriv:
assumes "norm z < fps_conv_radius f"
shows "eval_fps (fps_deriv f) z = deriv (eval_fps f) z"
by (intro DERIV_imp_deriv [symmetric] has_field_derivative_eval_fps assms)
theorem fps_nth_conv_deriv:
fixes f :: "complex fps"
assumes "fps_conv_radius f > 0"
shows "fps_nth f n = (deriv ^^ n) (eval_fps f) 0 / fact n"
using assms
proof (induction n arbitrary: f)
case 0
thus ?case by (simp add: eval_fps_def)
next
case (Suc n f)
have "(deriv ^^ Suc n) (eval_fps f) 0 = (deriv ^^ n) (deriv (eval_fps f)) 0"
unfolding funpow_Suc_right o_def ..
also have "eventually (\<lambda>z::complex. z \<in> eball 0 (fps_conv_radius f)) (nhds 0)"
using Suc.prems by (intro eventually_nhds_in_open) (auto simp: zero_ereal_def)
hence "eventually (\<lambda>z. deriv (eval_fps f) z = eval_fps (fps_deriv f) z) (nhds 0)"
by eventually_elim (simp add: eval_fps_deriv)
hence "(deriv ^^ n) (deriv (eval_fps f)) 0 = (deriv ^^ n) (eval_fps (fps_deriv f)) 0"
by (intro higher_deriv_cong_ev refl)
also have "\<dots> / fact n = fps_nth (fps_deriv f) n"
using Suc.prems fps_conv_radius_deriv[of f]
by (intro Suc.IH [symmetric]) auto
also have "\<dots> / of_nat (Suc n) = fps_nth f (Suc n)"
by (simp add: fps_deriv_def del: of_nat_Suc)
finally show ?case by (simp add: field_split_simps)
qed
theorem eval_fps_eqD:
fixes f g :: "complex fps"
assumes "fps_conv_radius f > 0" "fps_conv_radius g > 0"
assumes "eventually (\<lambda>z. eval_fps f z = eval_fps g z) (nhds 0)"
shows "f = g"
proof (rule fps_ext)
fix n :: nat
have "fps_nth f n = (deriv ^^ n) (eval_fps f) 0 / fact n"
using assms by (intro fps_nth_conv_deriv)
also have "(deriv ^^ n) (eval_fps f) 0 = (deriv ^^ n) (eval_fps g) 0"
by (intro higher_deriv_cong_ev refl assms)
also have "\<dots> / fact n = fps_nth g n"
using assms by (intro fps_nth_conv_deriv [symmetric])
finally show "fps_nth f n = fps_nth g n" .
qed
lemma eval_fps_const [simp]:
fixes c :: "'a :: {banach, real_normed_div_algebra}"
shows "eval_fps (fps_const c) z = c"
proof -
have "(\<lambda>n::nat. if n \<in> {0} then c else 0) sums (\<Sum>n\<in>{0::nat}. c)"
by (rule sums_If_finite_set) auto
also have "?this \<longleftrightarrow> (\<lambda>n::nat. fps_nth (fps_const c) n * z ^ n) sums (\<Sum>n\<in>{0::nat}. c)"
by (intro sums_cong) auto
also have "(\<Sum>n\<in>{0::nat}. c) = c"
by simp
finally show ?thesis
by (simp add: eval_fps_def sums_iff)
qed
lemma eval_fps_0 [simp]:
"eval_fps (0 :: 'a :: {banach, real_normed_div_algebra} fps) z = 0"
by (simp only: fps_const_0_eq_0 [symmetric] eval_fps_const)
lemma eval_fps_1 [simp]:
"eval_fps (1 :: 'a :: {banach, real_normed_div_algebra} fps) z = 1"
by (simp only: fps_const_1_eq_1 [symmetric] eval_fps_const)
lemma eval_fps_numeral [simp]:
"eval_fps (numeral n :: 'a :: {banach, real_normed_div_algebra} fps) z = numeral n"
by (simp only: numeral_fps_const eval_fps_const)
lemma eval_fps_X_power [simp]:
"eval_fps (fps_X ^ m :: 'a :: {banach, real_normed_div_algebra} fps) z = z ^ m"
proof -
have "(\<lambda>n::nat. if n \<in> {m} then z ^ n else 0 :: 'a) sums (\<Sum>n\<in>{m::nat}. z ^ n)"
by (rule sums_If_finite_set) auto
also have "?this \<longleftrightarrow> (\<lambda>n::nat. fps_nth (fps_X ^ m) n * z ^ n) sums (\<Sum>n\<in>{m::nat}. z ^ n)"
by (intro sums_cong) (auto simp: fps_X_power_iff)
also have "(\<Sum>n\<in>{m::nat}. z ^ n) = z ^ m"
by simp
finally show ?thesis
by (simp add: eval_fps_def sums_iff)
qed
lemma eval_fps_X [simp]:
"eval_fps (fps_X :: 'a :: {banach, real_normed_div_algebra} fps) z = z"
using eval_fps_X_power[of 1 z] by (simp only: power_one_right)
lemma eval_fps_minus:
fixes f :: "'a :: {banach, real_normed_div_algebra} fps"
assumes "norm z < fps_conv_radius f"
shows "eval_fps (-f) z = -eval_fps f z"
using assms unfolding eval_fps_def
by (subst suminf_minus [symmetric]) (auto intro!: summable_fps)
lemma eval_fps_add:
fixes f g :: "'a :: {banach, real_normed_div_algebra} fps"
assumes "norm z < fps_conv_radius f" "norm z < fps_conv_radius g"
shows "eval_fps (f + g) z = eval_fps f z + eval_fps g z"
using assms unfolding eval_fps_def
by (subst suminf_add) (auto simp: ring_distribs intro!: summable_fps)
lemma eval_fps_diff:
fixes f g :: "'a :: {banach, real_normed_div_algebra} fps"
assumes "norm z < fps_conv_radius f" "norm z < fps_conv_radius g"
shows "eval_fps (f - g) z = eval_fps f z - eval_fps g z"
using assms unfolding eval_fps_def
by (subst suminf_diff) (auto simp: ring_distribs intro!: summable_fps)
lemma eval_fps_mult:
fixes f g :: "'a :: {banach, real_normed_div_algebra, comm_ring_1} fps"
assumes "norm z < fps_conv_radius f" "norm z < fps_conv_radius g"
shows "eval_fps (f * g) z = eval_fps f z * eval_fps g z"
proof -
have "eval_fps f z * eval_fps g z =
(\<Sum>k. \<Sum>i\<le>k. fps_nth f i * fps_nth g (k - i) * (z ^ i * z ^ (k - i)))"
unfolding eval_fps_def
proof (subst Cauchy_product)
show "summable (\<lambda>k. norm (fps_nth f k * z ^ k))" "summable (\<lambda>k. norm (fps_nth g k * z ^ k))"
by (rule norm_summable_fps assms)+
qed (simp_all add: algebra_simps)
also have "(\<lambda>k. \<Sum>i\<le>k. fps_nth f i * fps_nth g (k - i) * (z ^ i * z ^ (k - i))) =
(\<lambda>k. \<Sum>i\<le>k. fps_nth f i * fps_nth g (k - i) * z ^ k)"
by (intro ext sum.cong refl) (simp add: power_add [symmetric])
also have "suminf \<dots> = eval_fps (f * g) z"
by (simp add: eval_fps_def fps_mult_nth atLeast0AtMost sum_distrib_right)
finally show ?thesis ..
qed
lemma eval_fps_shift:
fixes f :: "'a :: {banach, real_normed_div_algebra, comm_ring_1} fps"
assumes "n \<le> subdegree f" "norm z < fps_conv_radius f"
shows "eval_fps (fps_shift n f) z = (if z = 0 then fps_nth f n else eval_fps f z / z ^ n)"
proof (cases "z = 0")
case False
have "eval_fps (fps_shift n f * fps_X ^ n) z = eval_fps (fps_shift n f) z * z ^ n"
using assms by (subst eval_fps_mult) simp_all
also from assms have "fps_shift n f * fps_X ^ n = f"
by (simp add: fps_shift_times_fps_X_power)
finally show ?thesis using False by (simp add: field_simps)
qed (simp_all add: eval_fps_at_0)
lemma eval_fps_exp [simp]:
fixes c :: "'a :: {banach, real_normed_field}"
shows "eval_fps (fps_exp c) z = exp (c * z)" unfolding eval_fps_def exp_def
- by (simp add: eval_fps_def exp_def scaleR_conv_of_real field_split_simps power_mult_distrib)
+ by (simp add: eval_fps_def exp_def scaleR_conv_of_real field_split_simps)
text \<open>
The case of division is more complicated and will therefore not be handled here.
Handling division becomes much more easy using complex analysis, and we will do so once
that is available.
\<close>
subsection \<open>Power series expansions of analytic functions\<close>
text \<open>
This predicate contains the notion that the given formal power series converges
in some disc of positive radius around the origin and is equal to the given complex
function there.
This relationship is unique in the sense that no complex function can have more than
one formal power series to which it expands, and if two holomorphic functions that are
holomorphic on a connected open set around the origin and have the same power series
expansion, they must be equal on that set.
More concrete statements about the radius of convergence can usually be made, but for
many purposes, the statment that the series converges to the function in some neighbourhood
of the origin is enough, and that can be shown almost fully automatically in most cases,
as there are straightforward introduction rules to show this.
In particular, when one wants to relate the coefficients of the power series to the
values of the derivatives of the function at the origin, or if one wants to approximate
the coefficients of the series with the singularities of the function, this predicate
is enough.
\<close>
definition\<^marker>\<open>tag important\<close>
has_fps_expansion :: "('a :: {banach,real_normed_div_algebra} \<Rightarrow> 'a) \<Rightarrow> 'a fps \<Rightarrow> bool"
(infixl "has'_fps'_expansion" 60)
where "(f has_fps_expansion F) \<longleftrightarrow>
fps_conv_radius F > 0 \<and> eventually (\<lambda>z. eval_fps F z = f z) (nhds 0)"
named_theorems fps_expansion_intros
lemma fps_nth_fps_expansion:
fixes f :: "complex \<Rightarrow> complex"
assumes "f has_fps_expansion F"
shows "fps_nth F n = (deriv ^^ n) f 0 / fact n"
proof -
have "fps_nth F n = (deriv ^^ n) (eval_fps F) 0 / fact n"
using assms by (intro fps_nth_conv_deriv) (auto simp: has_fps_expansion_def)
also have "(deriv ^^ n) (eval_fps F) 0 = (deriv ^^ n) f 0"
using assms by (intro higher_deriv_cong_ev) (auto simp: has_fps_expansion_def)
finally show ?thesis .
qed
lemma has_fps_expansion_imp_continuous:
fixes F :: "'a::{real_normed_field,banach} fps"
assumes "f has_fps_expansion F"
shows "continuous (at 0 within A) f"
proof -
from assms have "isCont (eval_fps F) 0"
by (intro continuous_eval_fps) (auto simp: has_fps_expansion_def zero_ereal_def)
also have "?this \<longleftrightarrow> isCont f 0" using assms
by (intro isCont_cong) (auto simp: has_fps_expansion_def)
finally have "isCont f 0" .
thus "continuous (at 0 within A) f"
by (simp add: continuous_at_imp_continuous_within)
qed
lemma has_fps_expansion_const [simp, intro, fps_expansion_intros]:
"(\<lambda>_. c) has_fps_expansion fps_const c"
by (auto simp: has_fps_expansion_def)
lemma has_fps_expansion_0 [simp, intro, fps_expansion_intros]:
"(\<lambda>_. 0) has_fps_expansion 0"
by (auto simp: has_fps_expansion_def)
lemma has_fps_expansion_1 [simp, intro, fps_expansion_intros]:
"(\<lambda>_. 1) has_fps_expansion 1"
by (auto simp: has_fps_expansion_def)
lemma has_fps_expansion_numeral [simp, intro, fps_expansion_intros]:
"(\<lambda>_. numeral n) has_fps_expansion numeral n"
by (auto simp: has_fps_expansion_def)
lemma has_fps_expansion_fps_X_power [fps_expansion_intros]:
"(\<lambda>x. x ^ n) has_fps_expansion (fps_X ^ n)"
by (auto simp: has_fps_expansion_def)
lemma has_fps_expansion_fps_X [fps_expansion_intros]:
"(\<lambda>x. x) has_fps_expansion fps_X"
by (auto simp: has_fps_expansion_def)
lemma has_fps_expansion_cmult_left [fps_expansion_intros]:
fixes c :: "'a :: {banach, real_normed_div_algebra, comm_ring_1}"
assumes "f has_fps_expansion F"
shows "(\<lambda>x. c * f x) has_fps_expansion fps_const c * F"
proof (cases "c = 0")
case False
from assms have "eventually (\<lambda>z. z \<in> eball 0 (fps_conv_radius F)) (nhds 0)"
by (intro eventually_nhds_in_open) (auto simp: has_fps_expansion_def zero_ereal_def)
moreover from assms have "eventually (\<lambda>z. eval_fps F z = f z) (nhds 0)"
by (auto simp: has_fps_expansion_def)
ultimately have "eventually (\<lambda>z. eval_fps (fps_const c * F) z = c * f z) (nhds 0)"
by eventually_elim (simp_all add: eval_fps_mult)
with assms and False show ?thesis
by (auto simp: has_fps_expansion_def fps_conv_radius_cmult_left)
qed auto
lemma has_fps_expansion_cmult_right [fps_expansion_intros]:
fixes c :: "'a :: {banach, real_normed_div_algebra, comm_ring_1}"
assumes "f has_fps_expansion F"
shows "(\<lambda>x. f x * c) has_fps_expansion F * fps_const c"
proof -
have "F * fps_const c = fps_const c * F"
by (intro fps_ext) (auto simp: mult.commute)
with has_fps_expansion_cmult_left [OF assms] show ?thesis
by (simp add: mult.commute)
qed
lemma has_fps_expansion_minus [fps_expansion_intros]:
assumes "f has_fps_expansion F"
shows "(\<lambda>x. - f x) has_fps_expansion -F"
proof -
from assms have "eventually (\<lambda>x. x \<in> eball 0 (fps_conv_radius F)) (nhds 0)"
by (intro eventually_nhds_in_open) (auto simp: has_fps_expansion_def zero_ereal_def)
moreover from assms have "eventually (\<lambda>x. eval_fps F x = f x) (nhds 0)"
by (auto simp: has_fps_expansion_def)
ultimately have "eventually (\<lambda>x. eval_fps (-F) x = -f x) (nhds 0)"
by eventually_elim (auto simp: eval_fps_minus)
thus ?thesis using assms by (auto simp: has_fps_expansion_def)
qed
lemma has_fps_expansion_add [fps_expansion_intros]:
assumes "f has_fps_expansion F" "g has_fps_expansion G"
shows "(\<lambda>x. f x + g x) has_fps_expansion F + G"
proof -
from assms have "0 < min (fps_conv_radius F) (fps_conv_radius G)"
by (auto simp: has_fps_expansion_def)
also have "\<dots> \<le> fps_conv_radius (F + G)"
by (rule fps_conv_radius_add)
finally have radius: "\<dots> > 0" .
from assms have "eventually (\<lambda>x. x \<in> eball 0 (fps_conv_radius F)) (nhds 0)"
"eventually (\<lambda>x. x \<in> eball 0 (fps_conv_radius G)) (nhds 0)"
by (intro eventually_nhds_in_open; force simp: has_fps_expansion_def zero_ereal_def)+
moreover have "eventually (\<lambda>x. eval_fps F x = f x) (nhds 0)"
and "eventually (\<lambda>x. eval_fps G x = g x) (nhds 0)"
using assms by (auto simp: has_fps_expansion_def)
ultimately have "eventually (\<lambda>x. eval_fps (F + G) x = f x + g x) (nhds 0)"
by eventually_elim (auto simp: eval_fps_add)
with radius show ?thesis by (auto simp: has_fps_expansion_def)
qed
lemma has_fps_expansion_diff [fps_expansion_intros]:
assumes "f has_fps_expansion F" "g has_fps_expansion G"
shows "(\<lambda>x. f x - g x) has_fps_expansion F - G"
using has_fps_expansion_add[of f F "\<lambda>x. - g x" "-G"] assms
by (simp add: has_fps_expansion_minus)
lemma has_fps_expansion_mult [fps_expansion_intros]:
fixes F G :: "'a :: {banach, real_normed_div_algebra, comm_ring_1} fps"
assumes "f has_fps_expansion F" "g has_fps_expansion G"
shows "(\<lambda>x. f x * g x) has_fps_expansion F * G"
proof -
from assms have "0 < min (fps_conv_radius F) (fps_conv_radius G)"
by (auto simp: has_fps_expansion_def)
also have "\<dots> \<le> fps_conv_radius (F * G)"
by (rule fps_conv_radius_mult)
finally have radius: "\<dots> > 0" .
from assms have "eventually (\<lambda>x. x \<in> eball 0 (fps_conv_radius F)) (nhds 0)"
"eventually (\<lambda>x. x \<in> eball 0 (fps_conv_radius G)) (nhds 0)"
by (intro eventually_nhds_in_open; force simp: has_fps_expansion_def zero_ereal_def)+
moreover have "eventually (\<lambda>x. eval_fps F x = f x) (nhds 0)"
and "eventually (\<lambda>x. eval_fps G x = g x) (nhds 0)"
using assms by (auto simp: has_fps_expansion_def)
ultimately have "eventually (\<lambda>x. eval_fps (F * G) x = f x * g x) (nhds 0)"
by eventually_elim (auto simp: eval_fps_mult)
with radius show ?thesis by (auto simp: has_fps_expansion_def)
qed
lemma has_fps_expansion_inverse [fps_expansion_intros]:
fixes F :: "'a :: {banach, real_normed_field} fps"
assumes "f has_fps_expansion F"
assumes "fps_nth F 0 \<noteq> 0"
shows "(\<lambda>x. inverse (f x)) has_fps_expansion inverse F"
proof -
have radius: "fps_conv_radius (inverse F) > 0"
using assms unfolding has_fps_expansion_def
by (intro fps_conv_radius_inverse_pos) auto
let ?R = "min (fps_conv_radius F) (fps_conv_radius (inverse F))"
from assms radius
have "eventually (\<lambda>x. x \<in> eball 0 (fps_conv_radius F)) (nhds 0)"
"eventually (\<lambda>x. x \<in> eball 0 (fps_conv_radius (inverse F))) (nhds 0)"
by (intro eventually_nhds_in_open; force simp: has_fps_expansion_def zero_ereal_def)+
moreover have "eventually (\<lambda>z. eval_fps F z = f z) (nhds 0)"
using assms by (auto simp: has_fps_expansion_def)
ultimately have "eventually (\<lambda>z. eval_fps (inverse F) z = inverse (f z)) (nhds 0)"
proof eventually_elim
case (elim z)
hence "eval_fps (inverse F * F) z = eval_fps (inverse F) z * f z"
by (subst eval_fps_mult) auto
also have "eval_fps (inverse F * F) z = 1"
using assms by (simp add: inverse_mult_eq_1)
finally show ?case by (auto simp: field_split_simps)
qed
with radius show ?thesis by (auto simp: has_fps_expansion_def)
qed
lemma has_fps_expansion_exp [fps_expansion_intros]:
fixes c :: "'a :: {banach, real_normed_field}"
shows "(\<lambda>x. exp (c * x)) has_fps_expansion fps_exp c"
by (auto simp: has_fps_expansion_def)
lemma has_fps_expansion_exp1 [fps_expansion_intros]:
"(\<lambda>x::'a :: {banach, real_normed_field}. exp x) has_fps_expansion fps_exp 1"
using has_fps_expansion_exp[of 1] by simp
lemma has_fps_expansion_exp_neg1 [fps_expansion_intros]:
"(\<lambda>x::'a :: {banach, real_normed_field}. exp (-x)) has_fps_expansion fps_exp (-1)"
using has_fps_expansion_exp[of "-1"] by simp
lemma has_fps_expansion_deriv [fps_expansion_intros]:
assumes "f has_fps_expansion F"
shows "deriv f has_fps_expansion fps_deriv F"
proof -
have "eventually (\<lambda>z. z \<in> eball 0 (fps_conv_radius F)) (nhds 0)"
using assms by (intro eventually_nhds_in_open)
(auto simp: has_fps_expansion_def zero_ereal_def)
moreover from assms have "eventually (\<lambda>z. eval_fps F z = f z) (nhds 0)"
by (auto simp: has_fps_expansion_def)
then obtain s where "open s" "0 \<in> s" and s: "\<And>w. w \<in> s \<Longrightarrow> eval_fps F w = f w"
by (auto simp: eventually_nhds)
hence "eventually (\<lambda>w. w \<in> s) (nhds 0)"
by (intro eventually_nhds_in_open) auto
ultimately have "eventually (\<lambda>z. eval_fps (fps_deriv F) z = deriv f z) (nhds 0)"
proof eventually_elim
case (elim z)
hence "eval_fps (fps_deriv F) z = deriv (eval_fps F) z"
by (simp add: eval_fps_deriv)
also have "eventually (\<lambda>w. w \<in> s) (nhds z)"
using elim and \<open>open s\<close> by (intro eventually_nhds_in_open) auto
hence "eventually (\<lambda>w. eval_fps F w = f w) (nhds z)"
by eventually_elim (simp add: s)
hence "deriv (eval_fps F) z = deriv f z"
by (intro deriv_cong_ev refl)
finally show ?case .
qed
with assms and fps_conv_radius_deriv[of F] show ?thesis
by (auto simp: has_fps_expansion_def)
qed
lemma fps_conv_radius_binomial:
fixes c :: "'a :: {real_normed_field,banach}"
shows "fps_conv_radius (fps_binomial c) = (if c \<in> \<nat> then \<infinity> else 1)"
unfolding fps_conv_radius_def by (simp add: conv_radius_gchoose)
lemma fps_conv_radius_ln:
fixes c :: "'a :: {banach, real_normed_field, field_char_0}"
shows "fps_conv_radius (fps_ln c) = (if c = 0 then \<infinity> else 1)"
proof (cases "c = 0")
case False
have "conv_radius (\<lambda>n. 1 / of_nat n :: 'a) = 1"
proof (rule conv_radius_ratio_limit_nonzero)
show "(\<lambda>n. norm (1 / of_nat n :: 'a) / norm (1 / of_nat (Suc n) :: 'a)) \<longlonglongrightarrow> 1"
using LIMSEQ_Suc_n_over_n by (simp add: norm_divide del: of_nat_Suc)
qed auto
also have "conv_radius (\<lambda>n. 1 / of_nat n :: 'a) =
conv_radius (\<lambda>n. if n = 0 then 0 else (- 1) ^ (n - 1) / of_nat n :: 'a)"
by (intro conv_radius_cong eventually_mono[OF eventually_gt_at_top[of 0]])
(simp add: norm_mult norm_divide norm_power)
finally show ?thesis using False unfolding fps_ln_def
by (subst fps_conv_radius_cmult_left) (simp_all add: fps_conv_radius_def)
qed (auto simp: fps_ln_def)
lemma fps_conv_radius_ln_nonzero [simp]:
assumes "c \<noteq> (0 :: 'a :: {banach,real_normed_field,field_char_0})"
shows "fps_conv_radius (fps_ln c) = 1"
using assms by (simp add: fps_conv_radius_ln)
lemma fps_conv_radius_sin [simp]:
fixes c :: "'a :: {banach, real_normed_field, field_char_0}"
shows "fps_conv_radius (fps_sin c) = \<infinity>"
proof (cases "c = 0")
case False
have "\<infinity> = conv_radius (\<lambda>n. of_real (sin_coeff n) :: 'a)"
proof (rule sym, rule conv_radius_inftyI'', rule summable_norm_cancel, goal_cases)
case (1 z)
show ?case using summable_norm_sin[of z] by (simp add: norm_mult)
qed
also have "\<dots> / norm c = conv_radius (\<lambda>n. c ^ n * of_real (sin_coeff n) :: 'a)"
using False by (subst conv_radius_mult_power) auto
also have "\<dots> = fps_conv_radius (fps_sin c)" unfolding fps_conv_radius_def
by (rule conv_radius_cong_weak) (auto simp add: fps_sin_def sin_coeff_def)
finally show ?thesis by simp
qed simp_all
lemma fps_conv_radius_cos [simp]:
fixes c :: "'a :: {banach, real_normed_field, field_char_0}"
shows "fps_conv_radius (fps_cos c) = \<infinity>"
proof (cases "c = 0")
case False
have "\<infinity> = conv_radius (\<lambda>n. of_real (cos_coeff n) :: 'a)"
proof (rule sym, rule conv_radius_inftyI'', rule summable_norm_cancel, goal_cases)
case (1 z)
show ?case using summable_norm_cos[of z] by (simp add: norm_mult)
qed
also have "\<dots> / norm c = conv_radius (\<lambda>n. c ^ n * of_real (cos_coeff n) :: 'a)"
using False by (subst conv_radius_mult_power) auto
also have "\<dots> = fps_conv_radius (fps_cos c)" unfolding fps_conv_radius_def
by (rule conv_radius_cong_weak) (auto simp add: fps_cos_def cos_coeff_def)
finally show ?thesis by simp
qed simp_all
lemma eval_fps_sin [simp]:
fixes z :: "'a :: {banach, real_normed_field, field_char_0}"
shows "eval_fps (fps_sin c) z = sin (c * z)"
proof -
have "(\<lambda>n. sin_coeff n *\<^sub>R (c * z) ^ n) sums sin (c * z)" by (rule sin_converges)
also have "(\<lambda>n. sin_coeff n *\<^sub>R (c * z) ^ n) = (\<lambda>n. fps_nth (fps_sin c) n * z ^ n)"
by (rule ext) (auto simp: sin_coeff_def fps_sin_def power_mult_distrib scaleR_conv_of_real)
finally show ?thesis by (simp add: sums_iff eval_fps_def)
qed
lemma eval_fps_cos [simp]:
fixes z :: "'a :: {banach, real_normed_field, field_char_0}"
shows "eval_fps (fps_cos c) z = cos (c * z)"
proof -
have "(\<lambda>n. cos_coeff n *\<^sub>R (c * z) ^ n) sums cos (c * z)" by (rule cos_converges)
also have "(\<lambda>n. cos_coeff n *\<^sub>R (c * z) ^ n) = (\<lambda>n. fps_nth (fps_cos c) n * z ^ n)"
by (rule ext) (auto simp: cos_coeff_def fps_cos_def power_mult_distrib scaleR_conv_of_real)
finally show ?thesis by (simp add: sums_iff eval_fps_def)
qed
lemma cos_eq_zero_imp_norm_ge:
assumes "cos (z :: complex) = 0"
shows "norm z \<ge> pi / 2"
proof -
from assms obtain n where "z = complex_of_real ((of_int n + 1 / 2) * pi)"
by (auto simp: cos_eq_0 algebra_simps)
also have "norm \<dots> = \<bar>real_of_int n + 1 / 2\<bar> * pi"
by (subst norm_of_real) (simp_all add: abs_mult)
also have "real_of_int n + 1 / 2 = of_int (2 * n + 1) / 2" by simp
also have "\<bar>\<dots>\<bar> = of_int \<bar>2 * n + 1\<bar> / 2" by (subst abs_divide) simp_all
also have "\<dots> * pi = of_int \<bar>2 * n + 1\<bar> * (pi / 2)" by simp
also have "\<dots> \<ge> of_int 1 * (pi / 2)"
by (intro mult_right_mono, subst of_int_le_iff) (auto simp: abs_if)
finally show ?thesis by simp
qed
lemma eval_fps_binomial:
fixes c :: complex
assumes "norm z < 1"
shows "eval_fps (fps_binomial c) z = (1 + z) powr c"
using gen_binomial_complex[OF assms] by (simp add: sums_iff eval_fps_def)
lemma has_fps_expansion_binomial_complex [fps_expansion_intros]:
fixes c :: complex
shows "(\<lambda>x. (1 + x) powr c) has_fps_expansion fps_binomial c"
proof -
have *: "eventually (\<lambda>z::complex. z \<in> eball 0 1) (nhds 0)"
by (intro eventually_nhds_in_open) auto
thus ?thesis
by (auto simp: has_fps_expansion_def eval_fps_binomial fps_conv_radius_binomial
intro!: eventually_mono [OF *])
qed
lemma has_fps_expansion_sin [fps_expansion_intros]:
fixes c :: "'a :: {banach, real_normed_field, field_char_0}"
shows "(\<lambda>x. sin (c * x)) has_fps_expansion fps_sin c"
by (auto simp: has_fps_expansion_def)
lemma has_fps_expansion_sin' [fps_expansion_intros]:
"(\<lambda>x::'a :: {banach, real_normed_field}. sin x) has_fps_expansion fps_sin 1"
using has_fps_expansion_sin[of 1] by simp
lemma has_fps_expansion_cos [fps_expansion_intros]:
fixes c :: "'a :: {banach, real_normed_field, field_char_0}"
shows "(\<lambda>x. cos (c * x)) has_fps_expansion fps_cos c"
by (auto simp: has_fps_expansion_def)
lemma has_fps_expansion_cos' [fps_expansion_intros]:
"(\<lambda>x::'a :: {banach, real_normed_field}. cos x) has_fps_expansion fps_cos 1"
using has_fps_expansion_cos[of 1] by simp
lemma has_fps_expansion_shift [fps_expansion_intros]:
fixes F :: "'a :: {banach, real_normed_field} fps"
assumes "f has_fps_expansion F" and "n \<le> subdegree F"
assumes "c = fps_nth F n"
shows "(\<lambda>x. if x = 0 then c else f x / x ^ n) has_fps_expansion (fps_shift n F)"
proof -
have "eventually (\<lambda>x. x \<in> eball 0 (fps_conv_radius F)) (nhds 0)"
using assms by (intro eventually_nhds_in_open) (auto simp: has_fps_expansion_def zero_ereal_def)
moreover have "eventually (\<lambda>x. eval_fps F x = f x) (nhds 0)"
using assms by (auto simp: has_fps_expansion_def)
ultimately have "eventually (\<lambda>x. eval_fps (fps_shift n F) x =
(if x = 0 then c else f x / x ^ n)) (nhds 0)"
by eventually_elim (auto simp: eval_fps_shift assms)
with assms show ?thesis by (auto simp: has_fps_expansion_def)
qed
lemma has_fps_expansion_divide [fps_expansion_intros]:
fixes F G :: "'a :: {banach, real_normed_field} fps"
assumes "f has_fps_expansion F" and "g has_fps_expansion G" and
"subdegree G \<le> subdegree F" "G \<noteq> 0"
"c = fps_nth F (subdegree G) / fps_nth G (subdegree G)"
shows "(\<lambda>x. if x = 0 then c else f x / g x) has_fps_expansion (F / G)"
proof -
define n where "n = subdegree G"
define F' and G' where "F' = fps_shift n F" and "G' = fps_shift n G"
have "F = F' * fps_X ^ n" "G = G' * fps_X ^ n" unfolding F'_def G'_def n_def
by (rule fps_shift_times_fps_X_power [symmetric] le_refl | fact)+
moreover from assms have "fps_nth G' 0 \<noteq> 0"
by (simp add: G'_def n_def)
ultimately have FG: "F / G = F' * inverse G'"
by (simp add: fps_divide_unit)
have "(\<lambda>x. (if x = 0 then fps_nth F n else f x / x ^ n) *
inverse (if x = 0 then fps_nth G n else g x / x ^ n)) has_fps_expansion F / G"
(is "?h has_fps_expansion _") unfolding FG F'_def G'_def n_def using \<open>G \<noteq> 0\<close>
by (intro has_fps_expansion_mult has_fps_expansion_inverse
has_fps_expansion_shift assms) auto
also have "?h = (\<lambda>x. if x = 0 then c else f x / g x)"
using assms(5) unfolding n_def
by (intro ext) (auto split: if_splits simp: field_simps)
finally show ?thesis .
qed
lemma has_fps_expansion_divide' [fps_expansion_intros]:
fixes F G :: "'a :: {banach, real_normed_field} fps"
assumes "f has_fps_expansion F" and "g has_fps_expansion G" and "fps_nth G 0 \<noteq> 0"
shows "(\<lambda>x. f x / g x) has_fps_expansion (F / G)"
proof -
have "(\<lambda>x. if x = 0 then fps_nth F 0 / fps_nth G 0 else f x / g x) has_fps_expansion (F / G)"
(is "?h has_fps_expansion _") using assms(3) by (intro has_fps_expansion_divide assms) auto
also from assms have "fps_nth F 0 = f 0" "fps_nth G 0 = g 0"
by (auto simp: has_fps_expansion_def eval_fps_at_0 dest: eventually_nhds_x_imp_x)
hence "?h = (\<lambda>x. f x / g x)" by auto
finally show ?thesis .
qed
lemma has_fps_expansion_tan [fps_expansion_intros]:
fixes c :: "'a :: {banach, real_normed_field, field_char_0}"
shows "(\<lambda>x. tan (c * x)) has_fps_expansion fps_tan c"
proof -
have "(\<lambda>x. sin (c * x) / cos (c * x)) has_fps_expansion fps_sin c / fps_cos c"
by (intro fps_expansion_intros) auto
thus ?thesis by (simp add: tan_def fps_tan_def)
qed
lemma has_fps_expansion_tan' [fps_expansion_intros]:
"tan has_fps_expansion fps_tan (1 :: 'a :: {banach, real_normed_field, field_char_0})"
using has_fps_expansion_tan[of 1] by simp
lemma has_fps_expansion_imp_holomorphic:
assumes "f has_fps_expansion F"
obtains s where "open s" "0 \<in> s" "f holomorphic_on s" "\<And>z. z \<in> s \<Longrightarrow> f z = eval_fps F z"
proof -
from assms obtain s where s: "open s" "0 \<in> s" "\<And>z. z \<in> s \<Longrightarrow> eval_fps F z = f z"
unfolding has_fps_expansion_def eventually_nhds by blast
let ?s' = "eball 0 (fps_conv_radius F) \<inter> s"
have "eval_fps F holomorphic_on ?s'"
by (intro holomorphic_intros) auto
also have "?this \<longleftrightarrow> f holomorphic_on ?s'"
using s by (intro holomorphic_cong) auto
finally show ?thesis using s assms
by (intro that[of ?s']) (auto simp: has_fps_expansion_def zero_ereal_def)
qed
end
\ No newline at end of file
diff --git a/src/HOL/Analysis/Fashoda_Theorem.thy b/src/HOL/Analysis/Fashoda_Theorem.thy
--- a/src/HOL/Analysis/Fashoda_Theorem.thy
+++ b/src/HOL/Analysis/Fashoda_Theorem.thy
@@ -1,781 +1,781 @@
(* 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\<^marker>\<open>tag important\<close> 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 add: interval_bij_def sum.distrib [symmetric] scaleR_add_left [symmetric]
fun_eq_iff intro!: sum.cong)
(simp add: algebra_simps diff_divide_distrib [symmetric])
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_sum 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_sum_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
proposition 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 infnorm_mul[unfolded scalar_mult_eq_scaleR]
by (simp add: real_abs_infnorm infnorm_eq_0)
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}"
proof
show "(\<lambda>x::real^2. x $ i) ` cbox (- 1) 1 \<subseteq> {-1..1}" for i
by (auto simp: mem_box_cart)
show "{-1..1} \<subseteq> (\<lambda>x::real^2. x $ i) ` cbox (- 1) 1" for i
by (clarsimp simp: image_iff mem_box_cart Bex_def) (metis (no_types, hide_lams) vec_component)
qed
{
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_box_cart atLeastAtMost_iff
by auto
} note x0 = this
have 1: "box (- 1) (1::real^2) \<noteq> {}"
unfolding interval_eq_empty_cart by auto
have "negatex (x + y) $ i = (negatex x + negatex y) $ i \<and> negatex (c *\<^sub>R x) $ i = (c *\<^sub>R negatex x) $ i"
for i x y c
using exhaust_2 [of i] by (auto simp: negatex_def)
then have "bounded_linear negatex"
by (simp add: bounded_linearI' vec_eq_iff)
then have 2: "continuous_on (cbox (- 1) 1) (negatex \<circ> sqprojection \<circ> ?F)"
apply (intro continuous_intros continuous_on_component)
unfolding * sqprojection_def
apply (intro assms continuous_intros)+
apply (simp_all add: infnorm_eq_0 x0 linear_continuous_on)
done
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"
by (rule x0) (use y in auto)
then have *: "infnorm (sqprojection (?F y)) = 1"
unfolding y o_def
by - (rule lem2[rule_format])
have inf1: "infnorm x = 1"
unfolding *[symmetric] y o_def
by (rule lem1[rule_format])
show "x \<in> cbox (-1) 1"
unfolding mem_box_cart interval_cbox_cart infnorm_2
proof
fix i
show "(- 1) $ i \<le> x $ i \<and> x $ i \<le> 1 $ i"
using exhaust_2 [of i] inf1 by (auto simp: infnorm_2)
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"
by (rule x0) (use x in auto)
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 iff: "0 < sqprojection x$i \<longleftrightarrow> 0 < x$i" "sqprojection x$i < 0 \<longleftrightarrow> x$i < 0" if "x \<noteq> 0" for x i
proof -
have "inverse (infnorm x) > 0"
by (simp add: infnorm_pos_lt that)
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
have x1: "x $ 1 \<in> {- 1..1::real}" "x $ 2 \<in> {- 1..1::real}"
using x(1) unfolding mem_box_cart by auto
then have nz: "f (x $ 1) - g (x $ 2) \<noteq> 0"
using as by auto
consider "x $ 1 = -1" | "x $ 1 = 1" | "x $ 2 = -1" | "x $ 2 = 1"
using nx unfolding infnorm_eq_1_2 by auto
then show False
proof cases
case 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]]
by (auto simp: negatex_def 1)
moreover
from x1 have "g (x $ 2) \<in> cbox (-1) 1"
using assms(2) by blast
ultimately show False
unfolding iff[OF nz] vector_component_simps * mem_box_cart
using not_le by auto
next
case 2
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]] 2
by (auto simp: negatex_def)
moreover have "g (x $ 2) \<in> cbox (-1) 1"
using assms(2) x1 by blast
ultimately show False
unfolding iff[OF nz] vector_component_simps * mem_box_cart
using not_le by auto
next
case 3
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]] 3 by (auto simp: negatex_def)
moreover
from x1 have "f (x $ 1) \<in> cbox (-1) 1"
using assms(1) by blast
ultimately show False
unfolding iff[OF nz] vector_component_simps * mem_box_cart
by (erule_tac x=2 in allE) auto
next
case 4
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]] 4 by (auto simp: negatex_def)
moreover
from x1 have "f (x $ 1) \<in> cbox (-1) 1"
using assms(1) by blast
ultimately show False
unfolding iff[OF nz] vector_component_simps * mem_box_cart
by (erule_tac x=2 in allE) auto
qed
qed
proposition 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
theorem 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_box_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_box_cart,rule_format,of "f 0" 1]
unfolding mem_box_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_box_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_box_cart,rule_format,of "g 0" 2]
unfolding mem_box_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: 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
proof (rule_tac z="interval_bij (- 1,1) (a,b) z" in that)
show "interval_bij (- 1, 1) (a, b) z \<in> path_image f"
using zf by (simp add: interval_bij_bij_cart[OF *] path_image_def)
show "interval_bij (- 1, 1) (a, b) z \<in> path_image g"
using zg by (simp add: interval_bij_bij_cart[OF *] path_image_def)
qed
qed
subsection\<^marker>\<open>tag unimportant\<close> \<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>(*FIXME change title? *)
corollary fashoda_interlace:
fixes a :: "real^2"
assumes "path f"
and "path g"
and paf: "path_image f \<subseteq> cbox a b"
and pag: "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_box_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)
+ by(auto simp add: path_image_join)
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)
+ by (auto simp add:less_eq_vec_def forall_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
show "path_image ?P1 \<subseteq> cbox ?a ?b" "path_image ?P2 \<subseteq> cbox ?a ?b"
unfolding P1P2 path_image_linepath using startfin paf pag
by (auto simp: mem_box_cart segment_horizontal segment_vertical forall_2)
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
proof (rule that[of z])
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_box_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_box_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) 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_box_cart forall_2]
then have "z$1 \<noteq> pathstart g$1"
using prems(1) 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_box_cart forall_2
by (simp only: segment_vertical segment_horizontal vector_2) (auto simp: assms)
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_box_cart forall_2
by (simp only: segment_vertical segment_horizontal vector_2) (auto simp: assms)
qed
qed
end
diff --git a/src/HOL/Analysis/Function_Metric.thy b/src/HOL/Analysis/Function_Metric.thy
--- a/src/HOL/Analysis/Function_Metric.thy
+++ b/src/HOL/Analysis/Function_Metric.thy
@@ -1,386 +1,386 @@
(* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr with additions from LCP
License: BSD
*)
section \<open>Metrics on product spaces\<close>
theory Function_Metric
imports
Function_Topology
Elementary_Metric_Spaces
begin
text \<open>In general, the product topology is not metrizable, unless the index set is countable.
When the index set is countable, essentially any (convergent) combination of the metrics on the
factors will do. We use below the simplest one, based on \<open>L\<^sup>1\<close>, but \<open>L\<^sup>2\<close> would also work,
for instance.
What is not completely trivial is that the distance thus defined induces the same topology
as the product topology. This is what we have to prove to show that we have an instance
of \<^class>\<open>metric_space\<close>.
The proofs below would work verbatim for general countable products of metric spaces. However,
since distances are only implemented in terms of type classes, we only develop the theory
for countable products of the same space.\<close>
instantiation "fun" :: (countable, metric_space) metric_space
begin
definition\<^marker>\<open>tag important\<close> dist_fun_def:
"dist x y = (\<Sum>n. (1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1)"
definition\<^marker>\<open>tag important\<close> uniformity_fun_def:
"(uniformity::(('a \<Rightarrow> 'b) \<times> ('a \<Rightarrow> 'b)) filter) = (INF e\<in>{0<..}. principal {(x, y). dist (x::('a\<Rightarrow>'b)) y < e})"
text \<open>Except for the first one, the auxiliary lemmas below are only useful when proving the
instance: once it is proved, they become trivial consequences of the general theory of metric
spaces. It would thus be desirable to hide them once the instance is proved, but I do not know how
to do this.\<close>
lemma dist_fun_le_dist_first_terms:
"dist x y \<le> 2 * Max {dist (x (from_nat n)) (y (from_nat n))|n. n \<le> N} + (1/2)^N"
proof -
have "(\<Sum>n. (1 / 2) ^ (n+Suc N) * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1)
= (\<Sum>n. (1 / 2) ^ (Suc N) * ((1/2) ^ n * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1))"
by (rule suminf_cong, simp add: power_add)
also have "... = (1/2)^(Suc N) * (\<Sum>n. (1 / 2) ^ n * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1)"
apply (rule suminf_mult)
by (rule summable_comparison_test'[of "\<lambda>n. (1/2)^n"], auto simp add: summable_geometric_iff)
also have "... \<le> (1/2)^(Suc N) * (\<Sum>n. (1 / 2) ^ n)"
apply (simp, rule suminf_le, simp)
by (rule summable_comparison_test'[of "\<lambda>n. (1/2)^n"], auto simp add: summable_geometric_iff)
also have "... = (1/2)^(Suc N) * 2"
using suminf_geometric[of "1/2"] by auto
also have "... = (1/2)^N"
by auto
finally have *: "(\<Sum>n. (1 / 2) ^ (n+Suc N) * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1) \<le> (1/2)^N"
by simp
define M where "M = Max {dist (x (from_nat n)) (y (from_nat n))|n. n \<le> N}"
have "dist (x (from_nat 0)) (y (from_nat 0)) \<le> M"
unfolding M_def by (rule Max_ge, auto)
then have [simp]: "M \<ge> 0" by (meson dual_order.trans zero_le_dist)
have "dist (x (from_nat n)) (y (from_nat n)) \<le> M" if "n\<le>N" for n
unfolding M_def apply (rule Max_ge) using that by auto
then have i: "min (dist (x (from_nat n)) (y (from_nat n))) 1 \<le> M" if "n\<le>N" for n
using that by force
have "(\<Sum>n< Suc N. (1 / 2) ^ n * min (dist (x (from_nat n)) (y (from_nat n))) 1) \<le>
(\<Sum>n< Suc N. M * (1 / 2) ^ n)"
by (rule sum_mono, simp add: i)
also have "... = M * (\<Sum>n<Suc N. (1 / 2) ^ n)"
by (rule sum_distrib_left[symmetric])
also have "... \<le> M * (\<Sum>n. (1 / 2) ^ n)"
by (rule mult_left_mono, rule sum_le_suminf, auto simp add: summable_geometric_iff)
also have "... = M * 2"
using suminf_geometric[of "1/2"] by auto
finally have **: "(\<Sum>n< Suc N. (1 / 2) ^ n * min (dist (x (from_nat n)) (y (from_nat n))) 1) \<le> 2 * M"
by simp
have "dist x y = (\<Sum>n. (1 / 2) ^ n * min (dist (x (from_nat n)) (y (from_nat n))) 1)"
unfolding dist_fun_def by simp
also have "... = (\<Sum>n. (1 / 2) ^ (n+Suc N) * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1)
+ (\<Sum>n<Suc N. (1 / 2) ^ n * min (dist (x (from_nat n)) (y (from_nat n))) 1)"
apply (rule suminf_split_initial_segment)
by (rule summable_comparison_test'[of "\<lambda>n. (1/2)^n"], auto simp add: summable_geometric_iff)
also have "... \<le> 2 * M + (1/2)^N"
using * ** by auto
finally show ?thesis unfolding M_def by simp
qed
lemma open_fun_contains_ball_aux:
assumes "open (U::(('a \<Rightarrow> 'b) set))"
"x \<in> U"
shows "\<exists>e>0. {y. dist x y < e} \<subseteq> U"
proof -
have *: "openin (product_topology (\<lambda>i. euclidean) UNIV) U"
using open_fun_def assms by auto
obtain X where H: "Pi\<^sub>E UNIV X \<subseteq> U"
"\<And>i. openin euclidean (X i)"
"finite {i. X i \<noteq> topspace euclidean}"
"x \<in> Pi\<^sub>E UNIV X"
using product_topology_open_contains_basis[OF * \<open>x \<in> U\<close>] by auto
define I where "I = {i. X i \<noteq> topspace euclidean}"
have "finite I" unfolding I_def using H(3) by auto
{
fix i
have "x i \<in> X i" using \<open>x \<in> U\<close> H by auto
then have "\<exists>e. e>0 \<and> ball (x i) e \<subseteq> X i"
using \<open>openin euclidean (X i)\<close> open_openin open_contains_ball by blast
then obtain e where "e>0" "ball (x i) e \<subseteq> X i" by blast
define f where "f = min e (1/2)"
have "f>0" "f<1" unfolding f_def using \<open>e>0\<close> by auto
moreover have "ball (x i) f \<subseteq> X i" unfolding f_def using \<open>ball (x i) e \<subseteq> X i\<close> by auto
ultimately have "\<exists>f. f > 0 \<and> f < 1 \<and> ball (x i) f \<subseteq> X i" by auto
} note * = this
have "\<exists>e. \<forall>i. e i > 0 \<and> e i < 1 \<and> ball (x i) (e i) \<subseteq> X i"
by (rule choice, auto simp add: *)
then obtain e where "\<And>i. e i > 0" "\<And>i. e i < 1" "\<And>i. ball (x i) (e i) \<subseteq> X i"
by blast
define m where "m = Min {(1/2)^(to_nat i) * e i|i. i \<in> I}"
have "m > 0" if "I\<noteq>{}"
unfolding m_def Min_gr_iff using \<open>finite I\<close> \<open>I \<noteq> {}\<close> \<open>\<And>i. e i > 0\<close> by auto
moreover have "{y. dist x y < m} \<subseteq> U"
proof (auto)
fix y assume "dist x y < m"
have "y i \<in> X i" if "i \<in> I" for i
proof -
have *: "summable (\<lambda>n. (1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1)"
by (rule summable_comparison_test'[of "\<lambda>n. (1/2)^n"], auto simp add: summable_geometric_iff)
define n where "n = to_nat i"
have "(1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1 < m"
using \<open>dist x y < m\<close> unfolding dist_fun_def using sum_le_suminf[OF *, of "{n}"] by auto
then have "(1/2)^(to_nat i) * min (dist (x i) (y i)) 1 < m"
using \<open>n = to_nat i\<close> by auto
also have "... \<le> (1/2)^(to_nat i) * e i"
unfolding m_def apply (rule Min_le) using \<open>finite I\<close> \<open>i \<in> I\<close> by auto
ultimately have "min (dist (x i) (y i)) 1 < e i"
by (auto simp add: field_split_simps)
then have "dist (x i) (y i) < e i"
using \<open>e i < 1\<close> by auto
then show "y i \<in> X i" using \<open>ball (x i) (e i) \<subseteq> X i\<close> by auto
qed
then have "y \<in> Pi\<^sub>E UNIV X" using H(1) unfolding I_def topspace_euclidean by (auto simp add: PiE_iff)
then show "y \<in> U" using \<open>Pi\<^sub>E UNIV X \<subseteq> U\<close> by auto
qed
ultimately have *: "\<exists>m>0. {y. dist x y < m} \<subseteq> U" if "I \<noteq> {}" using that by auto
have "U = UNIV" if "I = {}"
using that H(1) unfolding I_def topspace_euclidean by (auto simp add: PiE_iff)
then have "\<exists>m>0. {y. dist x y < m} \<subseteq> U" if "I = {}" using that zero_less_one by blast
then show "\<exists>m>0. {y. dist x y < m} \<subseteq> U" using * by blast
qed
lemma ball_fun_contains_open_aux:
fixes x::"('a \<Rightarrow> 'b)" and e::real
assumes "e>0"
shows "\<exists>U. open U \<and> x \<in> U \<and> U \<subseteq> {y. dist x y < e}"
proof -
have "\<exists>N::nat. 2^N > 8/e"
by (simp add: real_arch_pow)
then obtain N::nat where "2^N > 8/e" by auto
define f where "f = e/4"
have [simp]: "e>0" "f > 0" unfolding f_def using assms by auto
define X::"('a \<Rightarrow> 'b set)" where "X = (\<lambda>i. if (\<exists>n\<le>N. i = from_nat n) then {z. dist (x i) z < f} else UNIV)"
have "{i. X i \<noteq> UNIV} \<subseteq> from_nat`{0..N}"
unfolding X_def by auto
then have "finite {i. X i \<noteq> topspace euclidean}"
unfolding topspace_euclidean using finite_surj by blast
have "\<And>i. open (X i)"
unfolding X_def using metric_space_class.open_ball open_UNIV by auto
then have "\<And>i. openin euclidean (X i)"
using open_openin by auto
define U where "U = Pi\<^sub>E UNIV X"
have "open U"
unfolding open_fun_def product_topology_def apply (rule topology_generated_by_Basis)
unfolding U_def using \<open>\<And>i. openin euclidean (X i)\<close> \<open>finite {i. X i \<noteq> topspace euclidean}\<close>
by auto
moreover have "x \<in> U"
unfolding U_def X_def by (auto simp add: PiE_iff)
moreover have "dist x y < e" if "y \<in> U" for y
proof -
have *: "dist (x (from_nat n)) (y (from_nat n)) \<le> f" if "n \<le> N" for n
using \<open>y \<in> U\<close> unfolding U_def apply (auto simp add: PiE_iff)
unfolding X_def using that by (metis less_imp_le mem_Collect_eq)
have **: "Max {dist (x (from_nat n)) (y (from_nat n))|n. n \<le> N} \<le> f"
apply (rule Max.boundedI) using * by auto
have "dist x y \<le> 2 * Max {dist (x (from_nat n)) (y (from_nat n))|n. n \<le> N} + (1/2)^N"
by (rule dist_fun_le_dist_first_terms)
also have "... \<le> 2 * f + e / 8"
- apply (rule add_mono) using ** \<open>2^N > 8/e\<close> by(auto simp add: algebra_simps field_split_simps)
+ apply (rule add_mono) using ** \<open>2^N > 8/e\<close> by(auto simp add: field_split_simps)
also have "... \<le> e/2 + e/8"
unfolding f_def by auto
also have "... < e"
by auto
finally show "dist x y < e" by simp
qed
ultimately show ?thesis by auto
qed
lemma fun_open_ball_aux:
fixes U::"('a \<Rightarrow> 'b) set"
shows "open U \<longleftrightarrow> (\<forall>x\<in>U. \<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> y \<in> U)"
proof (auto)
assume "open U"
fix x assume "x \<in> U"
then show "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> y \<in> U"
using open_fun_contains_ball_aux[OF \<open>open U\<close> \<open>x \<in> U\<close>] by auto
next
assume H: "\<forall>x\<in>U. \<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> y \<in> U"
define K where "K = {V. open V \<and> V \<subseteq> U}"
{
fix x assume "x \<in> U"
then obtain e where e: "e>0" "{y. dist x y < e} \<subseteq> U" using H by blast
then obtain V where V: "open V" "x \<in> V" "V \<subseteq> {y. dist x y < e}"
using ball_fun_contains_open_aux[OF \<open>e>0\<close>, of x] by auto
have "V \<in> K"
unfolding K_def using e(2) V(1) V(3) by auto
then have "x \<in> (\<Union>K)" using \<open>x \<in> V\<close> by auto
}
then have "(\<Union>K) = U"
unfolding K_def by auto
moreover have "open (\<Union>K)"
unfolding K_def by auto
ultimately show "open U" by simp
qed
instance proof
fix x y::"'a \<Rightarrow> 'b" show "(dist x y = 0) = (x = y)"
proof
assume "x = y"
then show "dist x y = 0" unfolding dist_fun_def using \<open>x = y\<close> by auto
next
assume "dist x y = 0"
have *: "summable (\<lambda>n. (1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1)"
by (rule summable_comparison_test'[of "\<lambda>n. (1/2)^n"], auto simp add: summable_geometric_iff)
have "(1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1 = 0" for n
using \<open>dist x y = 0\<close> unfolding dist_fun_def by (simp add: "*" suminf_eq_zero_iff)
then have "dist (x (from_nat n)) (y (from_nat n)) = 0" for n
by (metis div_0 min_def nonzero_mult_div_cancel_left power_eq_0_iff
zero_eq_1_divide_iff zero_neq_numeral)
then have "x (from_nat n) = y (from_nat n)" for n
by auto
then have "x i = y i" for i
by (metis from_nat_to_nat)
then show "x = y"
by auto
qed
next
text\<open>The proof of the triangular inequality is trivial, modulo the fact that we are dealing
with infinite series, hence we should justify the convergence at each step. In this
respect, the following simplification rule is extremely handy.\<close>
have [simp]: "summable (\<lambda>n. (1/2)^n * min (dist (u (from_nat n)) (v (from_nat n))) 1)" for u v::"'a \<Rightarrow> 'b"
by (rule summable_comparison_test'[of "\<lambda>n. (1/2)^n"], auto simp add: summable_geometric_iff)
fix x y z::"'a \<Rightarrow> 'b"
{
fix n
have *: "dist (x (from_nat n)) (y (from_nat n)) \<le>
dist (x (from_nat n)) (z (from_nat n)) + dist (y (from_nat n)) (z (from_nat n))"
by (simp add: dist_triangle2)
have "0 \<le> dist (y (from_nat n)) (z (from_nat n))"
using zero_le_dist by blast
moreover
{
assume "min (dist (y (from_nat n)) (z (from_nat n))) 1 \<noteq> dist (y (from_nat n)) (z (from_nat n))"
then have "1 \<le> min (dist (x (from_nat n)) (z (from_nat n))) 1 + min (dist (y (from_nat n)) (z (from_nat n))) 1"
by (metis (no_types) diff_le_eq diff_self min_def zero_le_dist zero_le_one)
}
ultimately have "min (dist (x (from_nat n)) (y (from_nat n))) 1 \<le>
min (dist (x (from_nat n)) (z (from_nat n))) 1 + min (dist (y (from_nat n)) (z (from_nat n))) 1"
using * by linarith
} note ineq = this
have "dist x y = (\<Sum>n. (1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1)"
unfolding dist_fun_def by simp
also have "... \<le> (\<Sum>n. (1/2)^n * min (dist (x (from_nat n)) (z (from_nat n))) 1
+ (1/2)^n * min (dist (y (from_nat n)) (z (from_nat n))) 1)"
apply (rule suminf_le)
using ineq apply (metis (no_types, hide_lams) add.right_neutral distrib_left
le_divide_eq_numeral1(1) mult_2_right mult_left_mono zero_le_one zero_le_power)
by (auto simp add: summable_add)
also have "... = (\<Sum>n. (1/2)^n * min (dist (x (from_nat n)) (z (from_nat n))) 1)
+ (\<Sum>n. (1/2)^n * min (dist (y (from_nat n)) (z (from_nat n))) 1)"
by (rule suminf_add[symmetric], auto)
also have "... = dist x z + dist y z"
unfolding dist_fun_def by simp
finally show "dist x y \<le> dist x z + dist y z"
by simp
next
text\<open>Finally, we show that the topology generated by the distance and the product
topology coincide. This is essentially contained in Lemma \<open>fun_open_ball_aux\<close>,
except that the condition to prove is expressed with filters. To deal with this,
we copy some mumbo jumbo from Lemma \<open>eventually_uniformity_metric\<close> in
\<^file>\<open>~~/src/HOL/Real_Vector_Spaces.thy\<close>\<close>
fix U::"('a \<Rightarrow> 'b) set"
have "eventually P uniformity \<longleftrightarrow> (\<exists>e>0. \<forall>x (y::('a \<Rightarrow> 'b)). dist x y < e \<longrightarrow> P (x, y))" for P
unfolding uniformity_fun_def apply (subst eventually_INF_base)
by (auto simp: eventually_principal subset_eq intro: bexI[of _ "min _ _"])
then show "open U = (\<forall>x\<in>U. \<forall>\<^sub>F (x', y) in uniformity. x' = x \<longrightarrow> y \<in> U)"
unfolding fun_open_ball_aux by simp
qed (simp add: uniformity_fun_def)
end
text \<open>Nice properties of spaces are preserved under countable products. In addition
to first countability, second countability and metrizability, as we have seen above,
completeness is also preserved, and therefore being Polish.\<close>
instance "fun" :: (countable, complete_space) complete_space
proof
fix u::"nat \<Rightarrow> ('a \<Rightarrow> 'b)" assume "Cauchy u"
have "Cauchy (\<lambda>n. u n i)" for i
unfolding cauchy_def proof (intro impI allI)
fix e assume "e>(0::real)"
obtain k where "i = from_nat k"
using from_nat_to_nat[of i] by metis
have "(1/2)^k * min e 1 > 0" using \<open>e>0\<close> by auto
then have "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (u m) (u n) < (1/2)^k * min e 1"
using \<open>Cauchy u\<close> unfolding cauchy_def by blast
then obtain N::nat where C: "\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (u m) (u n) < (1/2)^k * min e 1"
by blast
have "\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (u m i) (u n i) < e"
proof (auto)
fix m n::nat assume "m \<ge> N" "n \<ge> N"
have "(1/2)^k * min (dist (u m i) (u n i)) 1
= sum (\<lambda>p. (1/2)^p * min (dist (u m (from_nat p)) (u n (from_nat p))) 1) {k}"
using \<open>i = from_nat k\<close> by auto
also have "... \<le> (\<Sum>p. (1/2)^p * min (dist (u m (from_nat p)) (u n (from_nat p))) 1)"
apply (rule sum_le_suminf)
by (rule summable_comparison_test'[of "\<lambda>n. (1/2)^n"], auto simp add: summable_geometric_iff)
also have "... = dist (u m) (u n)"
unfolding dist_fun_def by auto
also have "... < (1/2)^k * min e 1"
using C \<open>m \<ge> N\<close> \<open>n \<ge> N\<close> by auto
finally have "min (dist (u m i) (u n i)) 1 < min e 1"
- by (auto simp add: algebra_simps field_split_simps)
+ by (auto simp add: field_split_simps)
then show "dist (u m i) (u n i) < e" by auto
qed
then show "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (u m i) (u n i) < e"
by blast
qed
then have "\<exists>x. (\<lambda>n. u n i) \<longlonglongrightarrow> x" for i
using Cauchy_convergent convergent_def by auto
then have "\<exists>x. \<forall>i. (\<lambda>n. u n i) \<longlonglongrightarrow> x i"
using choice by force
then obtain x where *: "\<And>i. (\<lambda>n. u n i) \<longlonglongrightarrow> x i" by blast
have "u \<longlonglongrightarrow> x"
proof (rule metric_LIMSEQ_I)
fix e assume [simp]: "e>(0::real)"
have i: "\<exists>K. \<forall>n\<ge>K. dist (u n i) (x i) < e/4" for i
by (rule metric_LIMSEQ_D, auto simp add: *)
have "\<exists>K. \<forall>i. \<forall>n\<ge>K i. dist (u n i) (x i) < e/4"
apply (rule choice) using i by auto
then obtain K where K: "\<And>i n. n \<ge> K i \<Longrightarrow> dist (u n i) (x i) < e/4"
by blast
have "\<exists>N::nat. 2^N > 4/e"
by (simp add: real_arch_pow)
then obtain N::nat where "2^N > 4/e" by auto
define L where "L = Max {K (from_nat n)|n. n \<le> N}"
have "dist (u k) x < e" if "k \<ge> L" for k
proof -
have *: "dist (u k (from_nat n)) (x (from_nat n)) \<le> e / 4" if "n \<le> N" for n
proof -
have "K (from_nat n) \<le> L"
unfolding L_def apply (rule Max_ge) using \<open>n \<le> N\<close> by auto
then have "k \<ge> K (from_nat n)" using \<open>k \<ge> L\<close> by auto
then show ?thesis using K less_imp_le by auto
qed
have **: "Max {dist (u k (from_nat n)) (x (from_nat n)) |n. n \<le> N} \<le> e/4"
apply (rule Max.boundedI) using * by auto
have "dist (u k) x \<le> 2 * Max {dist (u k (from_nat n)) (x (from_nat n)) |n. n \<le> N} + (1/2)^N"
using dist_fun_le_dist_first_terms by auto
also have "... \<le> 2 * e/4 + e/4"
apply (rule add_mono)
- using ** \<open>2^N > 4/e\<close> less_imp_le by (auto simp add: algebra_simps field_split_simps)
+ using ** \<open>2^N > 4/e\<close> less_imp_le by (auto simp add: field_split_simps)
also have "... < e" by auto
finally show ?thesis by simp
qed
then show "\<exists>L. \<forall>k\<ge>L. dist (u k) x < e" by blast
qed
then show "convergent u" using convergent_def by blast
qed
instance "fun" :: (countable, polish_space) polish_space
by standard
end
\ No newline at end of file
diff --git a/src/HOL/Analysis/Function_Topology.thy b/src/HOL/Analysis/Function_Topology.thy
--- a/src/HOL/Analysis/Function_Topology.thy
+++ b/src/HOL/Analysis/Function_Topology.thy
@@ -1,1688 +1,1688 @@
(* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr with additions from LCP
License: BSD
*)
theory Function_Topology
imports
Elementary_Topology
Abstract_Limits
Connected
begin
section \<open>Function Topology\<close>
text \<open>We want to define the general product topology.
The product topology on a product of topological spaces is generated by
the sets which are products of open sets along finitely many coordinates, and the whole
space along the other coordinates. This is the coarsest topology for which the projection
to each factor is continuous.
To form a product of objects in Isabelle/HOL, all these objects should be subsets of a common type
'a. The product is then \<^term>\<open>Pi\<^sub>E I X\<close>, the set of elements from \<open>'i\<close> to \<open>'a\<close> such that the \<open>i\<close>-th
coordinate belongs to \<open>X i\<close> for all \<open>i \<in> I\<close>.
Hence, to form a product of topological spaces, all these spaces should be subsets of a common type.
This means that type classes can not be used to define such a product if one wants to take the
product of different topological spaces (as the type 'a can only be given one structure of
topological space using type classes). On the other hand, one can define different topologies (as
introduced in \<open>thy\<close>) on one type, and these topologies do not need to
share the same maximal open set. Hence, one can form a product of topologies in this sense, and
this works well. The big caveat is that it does not interact well with the main body of
topology in Isabelle/HOL defined in terms of type classes... For instance, continuity of maps
is not defined in this setting.
As the product of different topological spaces is very important in several areas of
mathematics (for instance adeles), I introduce below the product topology in terms of topologies,
and reformulate afterwards the consequences in terms of type classes (which are of course very
handy for applications).
Given this limitation, it looks to me that it would be very beneficial to revamp the theory
of topological spaces in Isabelle/HOL in terms of topologies, and keep the statements involving
type classes as consequences of more general statements in terms of topologies (but I am
probably too naive here).
Here is an example of a reformulation using topologies. Let
@{text [display]
\<open>continuous_map T1 T2 f =
((\<forall> U. openin T2 U \<longrightarrow> openin T1 (f-`U \<inter> topspace(T1)))
\<and> (f`(topspace T1) \<subseteq> (topspace T2)))\<close>}
be the natural continuity definition of a map from the topology \<open>T1\<close> to the topology \<open>T2\<close>. Then
the current \<open>continuous_on\<close> (with type classes) can be redefined as
@{text [display] \<open>continuous_on s f =
continuous_map (top_of_set s) (topology euclidean) f\<close>}
In fact, I need \<open>continuous_map\<close> to express the continuity of the projection on subfactors
for the product topology, in Lemma~\<open>continuous_on_restrict_product_topology\<close>, and I show
the above equivalence in Lemma~\<open>continuous_map_iff_continuous\<close>.
I only develop the basics of the product topology in this theory. The most important missing piece
is Tychonov theorem, stating that a product of compact spaces is always compact for the product
topology, even when the product is not finite (or even countable).
I realized afterwards that this theory has a lot in common with \<^file>\<open>~~/src/HOL/Library/Finite_Map.thy\<close>.
\<close>
subsection \<open>The product topology\<close>
text \<open>We can now define the product topology, as generated by
the sets which are products of open sets along finitely many coordinates, and the whole
space along the other coordinates. Equivalently, it is generated by sets which are one open
set along one single coordinate, and the whole space along other coordinates. In fact, this is only
equivalent for nonempty products, but for the empty product the first formulation is better
(the second one gives an empty product space, while an empty product should have exactly one
point, equal to \<open>undefined\<close> along all coordinates.
So, we use the first formulation, which moreover seems to give rise to more straightforward proofs.
\<close>
definition\<^marker>\<open>tag important\<close> product_topology::"('i \<Rightarrow> ('a topology)) \<Rightarrow> ('i set) \<Rightarrow> (('i \<Rightarrow> 'a) topology)"
where "product_topology T I =
topology_generated_by {(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)}}"
abbreviation powertop_real :: "'a set \<Rightarrow> ('a \<Rightarrow> real) topology"
where "powertop_real \<equiv> product_topology (\<lambda>i. euclideanreal)"
text \<open>The total set of the product topology is the product of the total sets
along each coordinate.\<close>
proposition product_topology:
"product_topology X I =
topology
(arbitrary union_of
((finite intersection_of
(\<lambda>F. \<exists>i U. F = {f. f i \<in> U} \<and> i \<in> I \<and> openin (X i) U))
relative_to (\<Pi>\<^sub>E i\<in>I. topspace (X i))))"
(is "_ = topology (_ union_of ((_ intersection_of ?\<Psi>) relative_to ?TOP))")
proof -
let ?\<Omega> = "(\<lambda>F. \<exists>Y. F = Pi\<^sub>E I Y \<and> (\<forall>i. openin (X i) (Y i)) \<and> finite {i. Y i \<noteq> topspace (X i)})"
have *: "(finite' intersection_of ?\<Omega>) A = (finite intersection_of ?\<Psi> relative_to ?TOP) A" for A
proof -
have 1: "\<exists>U. (\<exists>\<U>. finite \<U> \<and> \<U> \<subseteq> Collect ?\<Psi> \<and> \<Inter>\<U> = U) \<and> ?TOP \<inter> U = \<Inter>\<U>"
if \<U>: "\<U> \<subseteq> Collect ?\<Omega>" and "finite' \<U>" "A = \<Inter>\<U>" "\<U> \<noteq> {}" for \<U>
proof -
have "\<forall>U \<in> \<U>. \<exists>Y. U = Pi\<^sub>E I Y \<and> (\<forall>i. openin (X i) (Y i)) \<and> finite {i. Y i \<noteq> topspace (X i)}"
using \<U> by auto
then obtain Y where Y: "\<And>U. U \<in> \<U> \<Longrightarrow> U = Pi\<^sub>E I (Y U) \<and> (\<forall>i. openin (X i) (Y U i)) \<and> finite {i. (Y U) i \<noteq> topspace (X i)}"
by metis
obtain U where "U \<in> \<U>"
using \<open>\<U> \<noteq> {}\<close> by blast
let ?F = "\<lambda>U. (\<lambda>i. {f. f i \<in> Y U i}) ` {i \<in> I. Y U i \<noteq> topspace (X i)}"
show ?thesis
proof (intro conjI exI)
show "finite (\<Union>U\<in>\<U>. ?F U)"
using Y \<open>finite' \<U>\<close> by auto
show "?TOP \<inter> \<Inter>(\<Union>U\<in>\<U>. ?F U) = \<Inter>\<U>"
proof
have *: "f \<in> U"
if "U \<in> \<U>" and "\<forall>V\<in>\<U>. \<forall>i. i \<in> I \<and> Y V i \<noteq> topspace (X i) \<longrightarrow> f i \<in> Y V i"
and "\<forall>i\<in>I. f i \<in> topspace (X i)" and "f \<in> extensional I" for f U
proof -
have "Pi\<^sub>E I (Y U) = U"
using Y \<open>U \<in> \<U>\<close> by blast
then show "f \<in> U"
using that unfolding PiE_def Pi_def by blast
qed
show "?TOP \<inter> \<Inter>(\<Union>U\<in>\<U>. ?F U) \<subseteq> \<Inter>\<U>"
by (auto simp: PiE_iff *)
show "\<Inter>\<U> \<subseteq> ?TOP \<inter> \<Inter>(\<Union>U\<in>\<U>. ?F U)"
using Y openin_subset \<open>finite' \<U>\<close> by fastforce
qed
qed (use Y openin_subset in \<open>blast+\<close>)
qed
have 2: "\<exists>\<U>'. finite' \<U>' \<and> \<U>' \<subseteq> Collect ?\<Omega> \<and> \<Inter>\<U>' = ?TOP \<inter> \<Inter>\<U>"
if \<U>: "\<U> \<subseteq> Collect ?\<Psi>" and "finite \<U>" for \<U>
proof (cases "\<U>={}")
case True
then show ?thesis
apply (rule_tac x="{?TOP}" in exI, simp)
apply (rule_tac x="\<lambda>i. topspace (X i)" in exI)
apply force
done
next
case False
then obtain U where "U \<in> \<U>"
by blast
have "\<forall>U \<in> \<U>. \<exists>i Y. U = {f. f i \<in> Y} \<and> i \<in> I \<and> openin (X i) Y"
using \<U> by auto
then obtain J Y where
Y: "\<And>U. U \<in> \<U> \<Longrightarrow> U = {f. f (J U) \<in> Y U} \<and> J U \<in> I \<and> openin (X (J U)) (Y U)"
by metis
let ?G = "\<lambda>U. \<Pi>\<^sub>E i\<in>I. if i = J U then Y U else topspace (X i)"
show ?thesis
proof (intro conjI exI)
show "finite (?G ` \<U>)" "?G ` \<U> \<noteq> {}"
using \<open>finite \<U>\<close> \<open>U \<in> \<U>\<close> by blast+
have *: "\<And>U. U \<in> \<U> \<Longrightarrow> openin (X (J U)) (Y U)"
using Y by force
show "?G ` \<U> \<subseteq> {Pi\<^sub>E I Y |Y. (\<forall>i. openin (X i) (Y i)) \<and> finite {i. Y i \<noteq> topspace (X i)}}"
apply clarsimp
apply (rule_tac x=" (\<lambda>i. if i = J U then Y U else topspace (X i))" in exI)
apply (auto simp: *)
done
next
show "(\<Inter>U\<in>\<U>. ?G U) = ?TOP \<inter> \<Inter>\<U>"
proof
have "(\<Pi>\<^sub>E i\<in>I. if i = J U then Y U else topspace (X i)) \<subseteq> (\<Pi>\<^sub>E i\<in>I. topspace (X i))"
apply (clarsimp simp: PiE_iff Ball_def all_conj_distrib split: if_split_asm)
using Y \<open>U \<in> \<U>\<close> openin_subset by blast+
then have "(\<Inter>U\<in>\<U>. ?G U) \<subseteq> ?TOP"
using \<open>U \<in> \<U>\<close>
by fastforce
moreover have "(\<Inter>U\<in>\<U>. ?G U) \<subseteq> \<Inter>\<U>"
using PiE_mem Y by fastforce
ultimately show "(\<Inter>U\<in>\<U>. ?G U) \<subseteq> ?TOP \<inter> \<Inter>\<U>"
by auto
qed (use Y in fastforce)
qed
qed
show ?thesis
unfolding relative_to_def intersection_of_def
by (safe; blast dest!: 1 2)
qed
show ?thesis
unfolding product_topology_def generate_topology_on_eq
apply (rule arg_cong [where f = topology])
apply (rule arg_cong [where f = "(union_of)arbitrary"])
apply (force simp: *)
done
qed
lemma topspace_product_topology [simp]:
"topspace (product_topology T I) = (\<Pi>\<^sub>E i\<in>I. topspace(T i))"
proof
show "topspace (product_topology T I) \<subseteq> (\<Pi>\<^sub>E i\<in>I. topspace (T i))"
unfolding product_topology_def topology_generated_by_topspace
unfolding topspace_def by auto
have "(\<Pi>\<^sub>E i\<in>I. topspace (T i)) \<in> {(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)}}"
using openin_topspace not_finite_existsD by auto
then show "(\<Pi>\<^sub>E i\<in>I. topspace (T i)) \<subseteq> topspace (product_topology T I)"
- unfolding product_topology_def using PiE_def by (auto simp add: topology_generated_by_topspace)
+ unfolding product_topology_def using PiE_def by (auto)
qed
lemma product_topology_basis:
assumes "\<And>i. openin (T i) (X i)" "finite {i. X i \<noteq> topspace (T i)}"
shows "openin (product_topology T I) (\<Pi>\<^sub>E i\<in>I. X i)"
unfolding product_topology_def
by (rule topology_generated_by_Basis) (use assms in auto)
proposition product_topology_open_contains_basis:
assumes "openin (product_topology T I) U" "x \<in> U"
shows "\<exists>X. x \<in> (\<Pi>\<^sub>E i\<in>I. X i) \<and> (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)} \<and> (\<Pi>\<^sub>E i\<in>I. X i) \<subseteq> U"
proof -
have "generate_topology_on {(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)}} U"
using assms unfolding product_topology_def by (intro openin_topology_generated_by) auto
then have "\<And>x. x\<in>U \<Longrightarrow> \<exists>X. x \<in> (\<Pi>\<^sub>E i\<in>I. X i) \<and> (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)} \<and> (\<Pi>\<^sub>E i\<in>I. X i) \<subseteq> U"
proof induction
case (Int U V x)
then obtain XU XV where H:
"x \<in> Pi\<^sub>E I XU" "(\<forall>i. openin (T i) (XU i))" "finite {i. XU i \<noteq> topspace (T i)}" "Pi\<^sub>E I XU \<subseteq> U"
"x \<in> Pi\<^sub>E I XV" "(\<forall>i. openin (T i) (XV i))" "finite {i. XV i \<noteq> topspace (T i)}" "Pi\<^sub>E I XV \<subseteq> V"
by auto meson
define X where "X = (\<lambda>i. XU i \<inter> XV i)"
have "Pi\<^sub>E I X \<subseteq> Pi\<^sub>E I XU \<inter> Pi\<^sub>E I XV"
unfolding X_def by (auto simp add: PiE_iff)
then have "Pi\<^sub>E I X \<subseteq> U \<inter> V" using H by auto
moreover have "\<forall>i. openin (T i) (X i)"
unfolding X_def using H by auto
moreover have "finite {i. X i \<noteq> topspace (T i)}"
apply (rule rev_finite_subset[of "{i. XU i \<noteq> topspace (T i)} \<union> {i. XV i \<noteq> topspace (T i)}"])
unfolding X_def using H by auto
moreover have "x \<in> Pi\<^sub>E I X"
unfolding X_def using H by auto
ultimately show ?case
by auto
next
case (UN K x)
then obtain k where "k \<in> K" "x \<in> k" by auto
with UN have "\<exists>X. x \<in> Pi\<^sub>E I X \<and> (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)} \<and> Pi\<^sub>E I X \<subseteq> k"
by simp
then obtain X where "x \<in> Pi\<^sub>E I X \<and> (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)} \<and> Pi\<^sub>E I X \<subseteq> k"
by blast
then have "x \<in> Pi\<^sub>E I X \<and> (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)} \<and> Pi\<^sub>E I X \<subseteq> (\<Union>K)"
using \<open>k \<in> K\<close> by auto
then show ?case
by auto
qed auto
then show ?thesis using \<open>x \<in> U\<close> by auto
qed
lemma product_topology_empty_discrete:
"product_topology T {} = discrete_topology {(\<lambda>x. undefined)}"
by (simp add: subtopology_eq_discrete_topology_sing)
lemma openin_product_topology:
"openin (product_topology X I) =
arbitrary union_of
((finite intersection_of (\<lambda>F. (\<exists>i U. F = {f. f i \<in> U} \<and> i \<in> I \<and> openin (X i) U)))
relative_to topspace (product_topology X I))"
apply (subst product_topology)
apply (simp add: topology_inverse' [OF istopology_subbase])
done
lemma subtopology_PiE:
"subtopology (product_topology X I) (\<Pi>\<^sub>E i\<in>I. (S i)) = product_topology (\<lambda>i. subtopology (X i) (S i)) I"
proof -
let ?P = "\<lambda>F. \<exists>i U. F = {f. f i \<in> U} \<and> i \<in> I \<and> openin (X i) U"
let ?X = "\<Pi>\<^sub>E i\<in>I. topspace (X i)"
have "finite intersection_of ?P relative_to ?X \<inter> Pi\<^sub>E I S =
finite intersection_of (?P relative_to ?X \<inter> Pi\<^sub>E I S) relative_to ?X \<inter> Pi\<^sub>E I S"
by (rule finite_intersection_of_relative_to)
also have "\<dots> = finite intersection_of
((\<lambda>F. \<exists>i U. F = {f. f i \<in> U} \<and> i \<in> I \<and> (openin (X i) relative_to S i) U)
relative_to ?X \<inter> Pi\<^sub>E I S)
relative_to ?X \<inter> Pi\<^sub>E I S"
apply (rule arg_cong2 [where f = "(relative_to)"])
apply (rule arg_cong [where f = "(intersection_of)finite"])
apply (rule ext)
apply (auto simp: relative_to_def intersection_of_def)
done
finally
have "finite intersection_of ?P relative_to ?X \<inter> Pi\<^sub>E I S =
finite intersection_of
(\<lambda>F. \<exists>i U. F = {f. f i \<in> U} \<and> i \<in> I \<and> (openin (X i) relative_to S i) U)
relative_to ?X \<inter> Pi\<^sub>E I S"
by (metis finite_intersection_of_relative_to)
then show ?thesis
unfolding topology_eq
apply clarify
apply (simp add: openin_product_topology flip: openin_relative_to)
apply (simp add: arbitrary_union_of_relative_to flip: PiE_Int)
done
qed
lemma product_topology_base_alt:
"finite intersection_of (\<lambda>F. (\<exists>i U. F = {f. f i \<in> U} \<and> i \<in> I \<and> openin (X i) U))
relative_to (\<Pi>\<^sub>E i\<in>I. topspace (X i)) =
(\<lambda>F. (\<exists>U. F = Pi\<^sub>E I U \<and> finite {i \<in> I. U i \<noteq> topspace(X i)} \<and> (\<forall>i \<in> I. openin (X i) (U i))))"
(is "?lhs = ?rhs")
proof -
have "(\<forall>F. ?lhs F \<longrightarrow> ?rhs F)"
unfolding all_relative_to all_intersection_of topspace_product_topology
proof clarify
fix \<F>
assume "finite \<F>" and "\<F> \<subseteq> {{f. f i \<in> U} |i U. i \<in> I \<and> openin (X i) U}"
then show "\<exists>U. (\<Pi>\<^sub>E i\<in>I. topspace (X i)) \<inter> \<Inter>\<F> = Pi\<^sub>E I U \<and>
finite {i \<in> I. U i \<noteq> topspace (X i)} \<and> (\<forall>i\<in>I. openin (X i) (U i))"
proof (induction)
case (insert F \<F>)
then obtain U where eq: "(\<Pi>\<^sub>E i\<in>I. topspace (X i)) \<inter> \<Inter>\<F> = Pi\<^sub>E I U"
and fin: "finite {i \<in> I. U i \<noteq> topspace (X i)}"
and ope: "\<And>i. i \<in> I \<Longrightarrow> openin (X i) (U i)"
by auto
obtain i V where "F = {f. f i \<in> V}" "i \<in> I" "openin (X i) V"
using insert by auto
let ?U = "\<lambda>j. U j \<inter> (if j = i then V else topspace(X j))"
show ?case
proof (intro exI conjI)
show "(\<Pi>\<^sub>E i\<in>I. topspace (X i)) \<inter> \<Inter>(insert F \<F>) = Pi\<^sub>E I ?U"
using eq PiE_mem \<open>i \<in> I\<close> by (auto simp: \<open>F = {f. f i \<in> V}\<close>) fastforce
next
show "finite {i \<in> I. ?U i \<noteq> topspace (X i)}"
by (rule rev_finite_subset [OF finite.insertI [OF fin]]) auto
next
show "\<forall>i\<in>I. openin (X i) (?U i)"
by (simp add: \<open>openin (X i) V\<close> ope openin_Int)
qed
qed (auto intro: dest: not_finite_existsD)
qed
moreover have "(\<forall>F. ?rhs F \<longrightarrow> ?lhs F)"
proof clarify
fix U :: "'a \<Rightarrow> 'b set"
assume fin: "finite {i \<in> I. U i \<noteq> topspace (X i)}" and ope: "\<forall>i\<in>I. openin (X i) (U i)"
let ?U = "\<Inter>i\<in>{i \<in> I. U i \<noteq> topspace (X i)}. {x. x i \<in> U i}"
show "?lhs (Pi\<^sub>E I U)"
unfolding relative_to_def topspace_product_topology
proof (intro exI conjI)
show "(finite intersection_of (\<lambda>F. \<exists>i U. F = {f. f i \<in> U} \<and> i \<in> I \<and> openin (X i) U)) ?U"
using fin ope by (intro finite_intersection_of_Inter finite_intersection_of_inc) auto
show "(\<Pi>\<^sub>E i\<in>I. topspace (X i)) \<inter> ?U = Pi\<^sub>E I U"
using ope openin_subset by fastforce
qed
qed
ultimately show ?thesis
by meson
qed
corollary openin_product_topology_alt:
"openin (product_topology X I) S \<longleftrightarrow>
(\<forall>x \<in> S. \<exists>U. finite {i \<in> I. U i \<noteq> topspace(X i)} \<and>
(\<forall>i \<in> I. openin (X i) (U i)) \<and> x \<in> Pi\<^sub>E I U \<and> Pi\<^sub>E I U \<subseteq> S)"
unfolding openin_product_topology arbitrary_union_of_alt product_topology_base_alt topspace_product_topology
apply safe
apply (drule bspec; blast)+
done
lemma closure_of_product_topology:
"(product_topology X I) closure_of (PiE I S) = PiE I (\<lambda>i. (X i) closure_of (S i))"
proof -
have *: "(\<forall>T. f \<in> T \<and> openin (product_topology X I) T \<longrightarrow> (\<exists>y\<in>Pi\<^sub>E I S. y \<in> T))
\<longleftrightarrow> (\<forall>i \<in> I. \<forall>T. f i \<in> T \<and> openin (X i) T \<longrightarrow> S i \<inter> T \<noteq> {})"
(is "?lhs = ?rhs")
if top: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> topspace (X i)" and ext: "f \<in> extensional I" for f
proof
assume L[rule_format]: ?lhs
show ?rhs
proof clarify
fix i T
assume "i \<in> I" "f i \<in> T" "openin (X i) T" "S i \<inter> T = {}"
then have "openin (product_topology X I) ((\<Pi>\<^sub>E i\<in>I. topspace (X i)) \<inter> {x. x i \<in> T})"
by (force simp: openin_product_topology intro: arbitrary_union_of_inc relative_to_inc finite_intersection_of_inc)
then show "False"
using L [of "topspace (product_topology X I) \<inter> {f. f i \<in> T}"] \<open>S i \<inter> T = {}\<close> \<open>f i \<in> T\<close> \<open>i \<in> I\<close>
by (auto simp: top ext PiE_iff)
qed
next
assume R [rule_format]: ?rhs
show ?lhs
proof (clarsimp simp: openin_product_topology union_of_def arbitrary_def)
fix \<U> U
assume
\<U>: "\<U> \<subseteq> Collect
(finite intersection_of (\<lambda>F. \<exists>i U. F = {x. x i \<in> U} \<and> i \<in> I \<and> openin (X i) U) relative_to
(\<Pi>\<^sub>E i\<in>I. topspace (X i)))" and
"f \<in> U" "U \<in> \<U>"
then have "(finite intersection_of (\<lambda>F. \<exists>i U. F = {x. x i \<in> U} \<and> i \<in> I \<and> openin (X i) U)
relative_to (\<Pi>\<^sub>E i\<in>I. topspace (X i))) U"
by blast
with \<open>f \<in> U\<close> \<open>U \<in> \<U>\<close>
obtain \<T> where "finite \<T>"
and \<T>: "\<And>C. C \<in> \<T> \<Longrightarrow> \<exists>i \<in> I. \<exists>V. openin (X i) V \<and> C = {x. x i \<in> V}"
and "topspace (product_topology X I) \<inter> \<Inter> \<T> \<subseteq> U" "f \<in> topspace (product_topology X I) \<inter> \<Inter> \<T>"
apply (clarsimp simp add: relative_to_def intersection_of_def)
apply (rule that, auto dest!: subsetD)
done
then have "f \<in> PiE I (topspace \<circ> X)" "f \<in> \<Inter>\<T>" and subU: "PiE I (topspace \<circ> X) \<inter> \<Inter>\<T> \<subseteq> U"
by (auto simp: PiE_iff)
have *: "f i \<in> topspace (X i) \<inter> \<Inter>{U. openin (X i) U \<and> {x. x i \<in> U} \<in> \<T>}
\<and> openin (X i) (topspace (X i) \<inter> \<Inter>{U. openin (X i) U \<and> {x. x i \<in> U} \<in> \<T>})"
if "i \<in> I" for i
proof -
have "finite ((\<lambda>U. {x. x i \<in> U}) -` \<T>)"
proof (rule finite_vimageI [OF \<open>finite \<T>\<close>])
show "inj (\<lambda>U. {x. x i \<in> U})"
by (auto simp: inj_on_def)
qed
then have fin: "finite {U. openin (X i) U \<and> {x. x i \<in> U} \<in> \<T>}"
by (rule rev_finite_subset) auto
have "openin (X i) (\<Inter> (insert (topspace (X i)) {U. openin (X i) U \<and> {x. x i \<in> U} \<in> \<T>}))"
by (rule openin_Inter) (auto simp: fin)
then show ?thesis
using \<open>f \<in> \<Inter> \<T>\<close> by (fastforce simp: that top)
qed
define \<Phi> where "\<Phi> \<equiv> \<lambda>i. topspace (X i) \<inter> \<Inter>{U. openin (X i) U \<and> {f. f i \<in> U} \<in> \<T>}"
have "\<forall>i \<in> I. \<exists>x. x \<in> S i \<inter> \<Phi> i"
using R [OF _ *] unfolding \<Phi>_def by blast
then obtain \<theta> where \<theta> [rule_format]: "\<forall>i \<in> I. \<theta> i \<in> S i \<inter> \<Phi> i"
by metis
show "\<exists>y\<in>Pi\<^sub>E I S. \<exists>x\<in>\<U>. y \<in> x"
proof
show "\<exists>U \<in> \<U>. (\<lambda>i \<in> I. \<theta> i) \<in> U"
proof
have "restrict \<theta> I \<in> PiE I (topspace \<circ> X) \<inter> \<Inter>\<T>"
using \<T> by (fastforce simp: \<Phi>_def PiE_def dest: \<theta>)
then show "restrict \<theta> I \<in> U"
using subU by blast
qed (rule \<open>U \<in> \<U>\<close>)
next
show "(\<lambda>i \<in> I. \<theta> i) \<in> Pi\<^sub>E I S"
using \<theta> by simp
qed
qed
qed
show ?thesis
apply (simp add: * closure_of_def PiE_iff set_eq_iff cong: conj_cong)
by metis
qed
corollary closedin_product_topology:
"closedin (product_topology X I) (PiE I S) \<longleftrightarrow> PiE I S = {} \<or> (\<forall>i \<in> I. closedin (X i) (S i))"
apply (simp add: PiE_eq PiE_eq_empty_iff closure_of_product_topology flip: closure_of_eq)
apply (metis closure_of_empty)
done
corollary closedin_product_topology_singleton:
"f \<in> extensional I \<Longrightarrow> closedin (product_topology X I) {f} \<longleftrightarrow> (\<forall>i \<in> I. closedin (X i) {f i})"
using PiE_singleton closedin_product_topology [of X I]
by (metis (no_types, lifting) all_not_in_conv insertI1)
lemma product_topology_empty:
"product_topology X {} = topology (\<lambda>S. S \<in> {{},{\<lambda>k. undefined}})"
unfolding product_topology union_of_def intersection_of_def arbitrary_def relative_to_def
by (auto intro: arg_cong [where f=topology])
lemma openin_product_topology_empty: "openin (product_topology X {}) S \<longleftrightarrow> S \<in> {{},{\<lambda>k. undefined}}"
unfolding union_of_def intersection_of_def arbitrary_def relative_to_def openin_product_topology
by auto
subsubsection \<open>The basic property of the product topology is the continuity of projections:\<close>
lemma continuous_map_product_coordinates [simp]:
assumes "i \<in> I"
shows "continuous_map (product_topology T I) (T i) (\<lambda>x. x i)"
proof -
{
fix U assume "openin (T i) U"
define X where "X = (\<lambda>j. if j = i then U else topspace (T j))"
then have *: "(\<lambda>x. x i) -` U \<inter> (\<Pi>\<^sub>E i\<in>I. topspace (T i)) = (\<Pi>\<^sub>E j\<in>I. X j)"
unfolding X_def using assms openin_subset[OF \<open>openin (T i) U\<close>]
by (auto simp add: PiE_iff, auto, metis subsetCE)
have **: "(\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)}"
unfolding X_def using \<open>openin (T i) U\<close> by auto
have "openin (product_topology T I) ((\<lambda>x. x i) -` U \<inter> (\<Pi>\<^sub>E i\<in>I. topspace (T i)))"
unfolding product_topology_def
apply (rule topology_generated_by_Basis)
apply (subst *)
using ** by auto
}
then show ?thesis unfolding continuous_map_alt
by (auto simp add: assms PiE_iff)
qed
lemma continuous_map_coordinatewise_then_product [intro]:
assumes "\<And>i. i \<in> I \<Longrightarrow> continuous_map T1 (T i) (\<lambda>x. f x i)"
"\<And>i x. i \<notin> I \<Longrightarrow> x \<in> topspace T1 \<Longrightarrow> f x i = undefined"
shows "continuous_map T1 (product_topology T I) f"
unfolding product_topology_def
proof (rule continuous_on_generated_topo)
fix U assume "U \<in> {Pi\<^sub>E I X |X. (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)}}"
then obtain X where H: "U = Pi\<^sub>E I X" "\<And>i. openin (T i) (X i)" "finite {i. X i \<noteq> topspace (T i)}"
by blast
define J where "J = {i \<in> I. X i \<noteq> topspace (T i)}"
have "finite J" "J \<subseteq> I" unfolding J_def using H(3) by auto
have "(\<lambda>x. f x i)-`(topspace(T i)) \<inter> topspace T1 = topspace T1" if "i \<in> I" for i
using that assms(1) by (simp add: continuous_map_preimage_topspace)
then have *: "(\<lambda>x. f x i)-`(X i) \<inter> topspace T1 = topspace T1" if "i \<in> I-J" for i
using that unfolding J_def by auto
have "f-`U \<inter> topspace T1 = (\<Inter>i\<in>I. (\<lambda>x. f x i)-`(X i) \<inter> topspace T1) \<inter> (topspace T1)"
by (subst H(1), auto simp add: PiE_iff assms)
also have "... = (\<Inter>i\<in>J. (\<lambda>x. f x i)-`(X i) \<inter> topspace T1) \<inter> (topspace T1)"
using * \<open>J \<subseteq> I\<close> by auto
also have "openin T1 (...)"
apply (rule openin_INT)
apply (simp add: \<open>finite J\<close>)
using H(2) assms(1) \<open>J \<subseteq> I\<close> by auto
ultimately show "openin T1 (f-`U \<inter> topspace T1)" by simp
next
show "f `topspace T1 \<subseteq> \<Union>{Pi\<^sub>E I X |X. (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)}}"
apply (subst topology_generated_by_topspace[symmetric])
apply (subst product_topology_def[symmetric])
apply (simp only: topspace_product_topology)
apply (auto simp add: PiE_iff)
using assms unfolding continuous_map_def by auto
qed
lemma continuous_map_product_then_coordinatewise [intro]:
assumes "continuous_map T1 (product_topology T I) f"
shows "\<And>i. i \<in> I \<Longrightarrow> continuous_map T1 (T i) (\<lambda>x. f x i)"
"\<And>i x. i \<notin> I \<Longrightarrow> x \<in> topspace T1 \<Longrightarrow> f x i = undefined"
proof -
fix i assume "i \<in> I"
have "(\<lambda>x. f x i) = (\<lambda>y. y i) o f" by auto
also have "continuous_map T1 (T i) (...)"
apply (rule continuous_map_compose[of _ "product_topology T I"])
using assms \<open>i \<in> I\<close> by auto
ultimately show "continuous_map T1 (T i) (\<lambda>x. f x i)"
by simp
next
fix i x assume "i \<notin> I" "x \<in> topspace T1"
have "f x \<in> topspace (product_topology T I)"
using assms \<open>x \<in> topspace T1\<close> unfolding continuous_map_def by auto
then have "f x \<in> (\<Pi>\<^sub>E i\<in>I. topspace (T i))"
using topspace_product_topology by metis
then show "f x i = undefined"
using \<open>i \<notin> I\<close> by (auto simp add: PiE_iff extensional_def)
qed
lemma continuous_on_restrict:
assumes "J \<subseteq> I"
shows "continuous_map (product_topology T I) (product_topology T J) (\<lambda>x. restrict x J)"
proof (rule continuous_map_coordinatewise_then_product)
fix i assume "i \<in> J"
then have "(\<lambda>x. restrict x J i) = (\<lambda>x. x i)" unfolding restrict_def by auto
then show "continuous_map (product_topology T I) (T i) (\<lambda>x. restrict x J i)"
using \<open>i \<in> J\<close> \<open>J \<subseteq> I\<close> by auto
next
fix i assume "i \<notin> J"
then show "restrict x J i = undefined" for x::"'a \<Rightarrow> 'b"
unfolding restrict_def by auto
qed
subsubsection\<^marker>\<open>tag important\<close> \<open>Powers of a single topological space as a topological space, using type classes\<close>
instantiation "fun" :: (type, topological_space) topological_space
begin
definition\<^marker>\<open>tag important\<close> open_fun_def:
"open U = openin (product_topology (\<lambda>i. euclidean) UNIV) U"
instance proof
have "topspace (product_topology (\<lambda>(i::'a). euclidean::('b topology)) UNIV) = UNIV"
unfolding topspace_product_topology topspace_euclidean by auto
then show "open (UNIV::('a \<Rightarrow> 'b) set)"
unfolding open_fun_def by (metis openin_topspace)
qed (auto simp add: open_fun_def)
end
lemma open_PiE [intro?]:
fixes X::"'i \<Rightarrow> ('b::topological_space) set"
assumes "\<And>i. open (X i)" "finite {i. X i \<noteq> UNIV}"
shows "open (Pi\<^sub>E UNIV X)"
by (simp add: assms open_fun_def product_topology_basis)
lemma euclidean_product_topology:
"product_topology (\<lambda>i. euclidean::('b::topological_space) topology) UNIV = euclidean"
by (metis open_openin topology_eq open_fun_def)
proposition product_topology_basis':
fixes x::"'i \<Rightarrow> 'a" and U::"'i \<Rightarrow> ('b::topological_space) set"
assumes "finite I" "\<And>i. i \<in> I \<Longrightarrow> open (U i)"
shows "open {f. \<forall>i\<in>I. f (x i) \<in> U i}"
proof -
define J where "J = x`I"
define V where "V = (\<lambda>y. if y \<in> J then \<Inter>{U i|i. i\<in>I \<and> x i = y} else UNIV)"
define X where "X = (\<lambda>y. if y \<in> J then V y else UNIV)"
have *: "open (X i)" for i
unfolding X_def V_def using assms by auto
have **: "finite {i. X i \<noteq> UNIV}"
unfolding X_def V_def J_def using assms(1) by auto
have "open (Pi\<^sub>E UNIV X)"
by (simp add: "*" "**" open_PiE)
moreover have "Pi\<^sub>E UNIV X = {f. \<forall>i\<in>I. f (x i) \<in> U i}"
apply (auto simp add: PiE_iff) unfolding X_def V_def J_def
proof (auto)
fix f :: "'a \<Rightarrow> 'b" and i :: 'i
assume a1: "i \<in> I"
assume a2: "\<forall>i. f i \<in> (if i \<in> x`I then if i \<in> x`I then \<Inter>{U ia |ia. ia \<in> I \<and> x ia = i} else UNIV else UNIV)"
have f3: "x i \<in> x`I"
using a1 by blast
have "U i \<in> {U ia |ia. ia \<in> I \<and> x ia = x i}"
using a1 by blast
then show "f (x i) \<in> U i"
using f3 a2 by (meson Inter_iff)
qed
ultimately show ?thesis by simp
qed
text \<open>The results proved in the general situation of products of possibly different
spaces have their counterparts in this simpler setting.\<close>
lemma continuous_on_product_coordinates [simp]:
"continuous_on UNIV (\<lambda>x. x i::('b::topological_space))"
using continuous_map_product_coordinates [of _ UNIV "\<lambda>i. euclidean"]
by (metis (no_types) continuous_map_iff_continuous euclidean_product_topology iso_tuple_UNIV_I subtopology_UNIV)
lemma continuous_on_coordinatewise_then_product [continuous_intros]:
fixes f :: "'a::topological_space \<Rightarrow> 'b \<Rightarrow> 'c::topological_space"
assumes "\<And>i. continuous_on S (\<lambda>x. f x i)"
shows "continuous_on S f"
using continuous_map_coordinatewise_then_product [of UNIV, where T = "\<lambda>i. euclidean"]
by (metis UNIV_I assms continuous_map_iff_continuous euclidean_product_topology)
lemma continuous_on_product_then_coordinatewise_UNIV:
assumes "continuous_on UNIV f"
shows "continuous_on UNIV (\<lambda>x. f x i)"
unfolding continuous_map_iff_continuous2 [symmetric]
by (rule continuous_map_product_then_coordinatewise [where I=UNIV]) (use assms in \<open>auto simp: euclidean_product_topology\<close>)
lemma continuous_on_product_then_coordinatewise:
assumes "continuous_on S f"
shows "continuous_on S (\<lambda>x. f x i)"
proof -
have "continuous_on S ((\<lambda>q. q i) \<circ> f)"
by (metis assms continuous_on_compose continuous_on_id
continuous_on_product_then_coordinatewise_UNIV continuous_on_subset subset_UNIV)
then show ?thesis
by auto
qed
lemma continuous_on_coordinatewise_iff:
fixes f :: "('a \<Rightarrow> real) \<Rightarrow> 'b \<Rightarrow> real"
shows "continuous_on (A \<inter> S) f \<longleftrightarrow> (\<forall>i. continuous_on (A \<inter> S) (\<lambda>x. f x i))"
by (auto simp: continuous_on_product_then_coordinatewise continuous_on_coordinatewise_then_product)
lemma continuous_map_span_sum:
fixes B :: "'a::real_normed_vector set"
assumes biB: "\<And>i. i \<in> I \<Longrightarrow> b i \<in> B"
shows "continuous_map euclidean (top_of_set (span B)) (\<lambda>x. \<Sum>i\<in>I. x i *\<^sub>R b i)"
proof (rule continuous_map_euclidean_top_of_set)
show "(\<lambda>x. \<Sum>i\<in>I. x i *\<^sub>R b i) -` span B = UNIV"
by auto (meson biB lessThan_iff span_base span_scale span_sum)
show "continuous_on UNIV (\<lambda>x. \<Sum>i\<in> I. x i *\<^sub>R b i)"
by (intro continuous_intros) auto
qed
subsubsection\<^marker>\<open>tag important\<close> \<open>Topological countability for product spaces\<close>
text \<open>The next two lemmas are useful to prove first or second countability
of product spaces, but they have more to do with countability and could
be put in the corresponding theory.\<close>
lemma countable_nat_product_event_const:
fixes F::"'a set" and a::'a
assumes "a \<in> F" "countable F"
shows "countable {x::(nat \<Rightarrow> 'a). (\<forall>i. x i \<in> F) \<and> finite {i. x i \<noteq> a}}"
proof -
have *: "{x::(nat \<Rightarrow> 'a). (\<forall>i. x i \<in> F) \<and> finite {i. x i \<noteq> a}}
\<subseteq> (\<Union>N. {x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>N. x i = a)})"
using infinite_nat_iff_unbounded_le by fastforce
have "countable {x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>N. x i = a)}" for N::nat
proof (induction N)
case 0
have "{x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>(0::nat). x i = a)} = {(\<lambda>i. a)}"
using \<open>a \<in> F\<close> by auto
then show ?case by auto
next
case (Suc N)
define f::"((nat \<Rightarrow> 'a) \<times> 'a) \<Rightarrow> (nat \<Rightarrow> 'a)"
where "f = (\<lambda>(x, b). (\<lambda>i. if i = N then b else x i))"
have "{x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>Suc N. x i = a)} \<subseteq> f`({x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>N. x i = a)} \<times> F)"
proof (auto)
fix x assume H: "\<forall>i::nat. x i \<in> F" "\<forall>i\<ge>Suc N. x i = a"
define y where "y = (\<lambda>i. if i = N then a else x i)"
have "f (y, x N) = x"
unfolding f_def y_def by auto
moreover have "(y, x N) \<in> {x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>N. x i = a)} \<times> F"
unfolding y_def using H \<open>a \<in> F\<close> by auto
ultimately show "x \<in> f`({x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>N. x i = a)} \<times> F)"
by (metis (no_types, lifting) image_eqI)
qed
moreover have "countable ({x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>N. x i = a)} \<times> F)"
using Suc.IH assms(2) by auto
ultimately show ?case
by (meson countable_image countable_subset)
qed
then show ?thesis using countable_subset[OF *] by auto
qed
lemma countable_product_event_const:
fixes F::"('a::countable) \<Rightarrow> 'b set" and b::'b
assumes "\<And>i. countable (F i)"
shows "countable {f::('a \<Rightarrow> 'b). (\<forall>i. f i \<in> F i) \<and> (finite {i. f i \<noteq> b})}"
proof -
define G where "G = (\<Union>i. F i) \<union> {b}"
have "countable G" unfolding G_def using assms by auto
have "b \<in> G" unfolding G_def by auto
define pi where "pi = (\<lambda>(x::(nat \<Rightarrow> 'b)). (\<lambda> i::'a. x ((to_nat::('a \<Rightarrow> nat)) i)))"
have "{f::('a \<Rightarrow> 'b). (\<forall>i. f i \<in> F i) \<and> (finite {i. f i \<noteq> b})}
\<subseteq> pi`{g::(nat \<Rightarrow> 'b). (\<forall>j. g j \<in> G) \<and> (finite {j. g j \<noteq> b})}"
proof (auto)
fix f assume H: "\<forall>i. f i \<in> F i" "finite {i. f i \<noteq> b}"
define I where "I = {i. f i \<noteq> b}"
define g where "g = (\<lambda>j. if j \<in> to_nat`I then f (from_nat j) else b)"
have "{j. g j \<noteq> b} \<subseteq> to_nat`I" unfolding g_def by auto
then have "finite {j. g j \<noteq> b}"
unfolding I_def using H(2) using finite_surj by blast
moreover have "g j \<in> G" for j
unfolding g_def G_def using H by auto
ultimately have "g \<in> {g::(nat \<Rightarrow> 'b). (\<forall>j. g j \<in> G) \<and> (finite {j. g j \<noteq> b})}"
by auto
moreover have "f = pi g"
unfolding pi_def g_def I_def using H by fastforce
ultimately show "f \<in> pi`{g. (\<forall>j. g j \<in> G) \<and> finite {j. g j \<noteq> b}}"
by auto
qed
then show ?thesis
using countable_nat_product_event_const[OF \<open>b \<in> G\<close> \<open>countable G\<close>]
by (meson countable_image countable_subset)
qed
instance "fun" :: (countable, first_countable_topology) first_countable_topology
proof
fix x::"'a \<Rightarrow> 'b"
have "\<exists>A::('b \<Rightarrow> nat \<Rightarrow> 'b set). \<forall>x. (\<forall>i. x \<in> A x i \<and> open (A x i)) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>i. A x i \<subseteq> S))"
apply (rule choice) using first_countable_basis by auto
then obtain A::"('b \<Rightarrow> nat \<Rightarrow> 'b set)" where A: "\<And>x i. x \<in> A x i"
"\<And>x i. open (A x i)"
"\<And>x S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> (\<exists>i. A x i \<subseteq> S)"
by metis
text \<open>\<open>B i\<close> is a countable basis of neighborhoods of \<open>x\<^sub>i\<close>.\<close>
define B where "B = (\<lambda>i. (A (x i))`UNIV \<union> {UNIV})"
have "countable (B i)" for i unfolding B_def by auto
have open_B: "\<And>X i. X \<in> B i \<Longrightarrow> open X"
by (auto simp: B_def A)
define K where "K = {Pi\<^sub>E UNIV X | X. (\<forall>i. X i \<in> B i) \<and> finite {i. X i \<noteq> UNIV}}"
have "Pi\<^sub>E UNIV (\<lambda>i. UNIV) \<in> K"
unfolding K_def B_def by auto
then have "K \<noteq> {}" by auto
have "countable {X. (\<forall>i. X i \<in> B i) \<and> finite {i. X i \<noteq> UNIV}}"
apply (rule countable_product_event_const) using \<open>\<And>i. countable (B i)\<close> by auto
moreover have "K = (\<lambda>X. Pi\<^sub>E UNIV X)`{X. (\<forall>i. X i \<in> B i) \<and> finite {i. X i \<noteq> UNIV}}"
unfolding K_def by auto
ultimately have "countable K" by auto
have "x \<in> k" if "k \<in> K" for k
using that unfolding K_def B_def apply auto using A(1) by auto
have "open k" if "k \<in> K" for k
using that unfolding K_def by (blast intro: open_B open_PiE elim: )
have Inc: "\<exists>k\<in>K. k \<subseteq> U" if "open U \<and> x \<in> U" for U
proof -
have "openin (product_topology (\<lambda>i. euclidean) UNIV) U" "x \<in> U"
using \<open>open U \<and> x \<in> U\<close> unfolding open_fun_def by auto
with product_topology_open_contains_basis[OF this]
have "\<exists>X. x \<in> (\<Pi>\<^sub>E i\<in>UNIV. X i) \<and> (\<forall>i. open (X i)) \<and> finite {i. X i \<noteq> UNIV} \<and> (\<Pi>\<^sub>E i\<in>UNIV. X i) \<subseteq> U"
by simp
then obtain X where H: "x \<in> (\<Pi>\<^sub>E i\<in>UNIV. X i)"
"\<And>i. open (X i)"
"finite {i. X i \<noteq> UNIV}"
"(\<Pi>\<^sub>E i\<in>UNIV. X i) \<subseteq> U"
by auto
define I where "I = {i. X i \<noteq> UNIV}"
define Y where "Y = (\<lambda>i. if i \<in> I then (SOME y. y \<in> B i \<and> y \<subseteq> X i) else UNIV)"
have *: "\<exists>y. y \<in> B i \<and> y \<subseteq> X i" for i
unfolding B_def using A(3)[OF H(2)] H(1) by (metis PiE_E UNIV_I UnCI image_iff)
have **: "Y i \<in> B i \<and> Y i \<subseteq> X i" for i
apply (cases "i \<in> I")
unfolding Y_def using * that apply (auto)
apply (metis (no_types, lifting) someI, metis (no_types, lifting) someI_ex subset_iff)
unfolding B_def apply simp
unfolding I_def apply auto
done
have "{i. Y i \<noteq> UNIV} \<subseteq> I"
unfolding Y_def by auto
then have ***: "finite {i. Y i \<noteq> UNIV}"
unfolding I_def using H(3) rev_finite_subset by blast
have "(\<forall>i. Y i \<in> B i) \<and> finite {i. Y i \<noteq> UNIV}"
using ** *** by auto
then have "Pi\<^sub>E UNIV Y \<in> K"
unfolding K_def by auto
have "Y i \<subseteq> X i" for i
apply (cases "i \<in> I") using ** apply simp unfolding Y_def I_def by auto
then have "Pi\<^sub>E UNIV Y \<subseteq> Pi\<^sub>E UNIV X" by auto
then have "Pi\<^sub>E UNIV Y \<subseteq> U" using H(4) by auto
then show ?thesis using \<open>Pi\<^sub>E UNIV Y \<in> K\<close> by auto
qed
show "\<exists>L. (\<forall>(i::nat). x \<in> L i \<and> open (L i)) \<and> (\<forall>U. open U \<and> x \<in> U \<longrightarrow> (\<exists>i. L i \<subseteq> U))"
apply (rule first_countableI[of K])
using \<open>countable K\<close> \<open>\<And>k. k \<in> K \<Longrightarrow> x \<in> k\<close> \<open>\<And>k. k \<in> K \<Longrightarrow> open k\<close> Inc by auto
qed
proposition product_topology_countable_basis:
shows "\<exists>K::(('a::countable \<Rightarrow> 'b::second_countable_topology) set set).
topological_basis K \<and> countable K \<and>
(\<forall>k\<in>K. \<exists>X. (k = Pi\<^sub>E UNIV X) \<and> (\<forall>i. open (X i)) \<and> finite {i. X i \<noteq> UNIV})"
proof -
obtain B::"'b set set" where B: "countable B \<and> topological_basis B"
using ex_countable_basis by auto
then have "B \<noteq> {}" by (meson UNIV_I empty_iff open_UNIV topological_basisE)
define B2 where "B2 = B \<union> {UNIV}"
have "countable B2"
unfolding B2_def using B by auto
have "open U" if "U \<in> B2" for U
using that unfolding B2_def using B topological_basis_open by auto
define K where "K = {Pi\<^sub>E UNIV X | X. (\<forall>i::'a. X i \<in> B2) \<and> finite {i. X i \<noteq> UNIV}}"
have i: "\<forall>k\<in>K. \<exists>X. (k = Pi\<^sub>E UNIV X) \<and> (\<forall>i. open (X i)) \<and> finite {i. X i \<noteq> UNIV}"
unfolding K_def using \<open>\<And>U. U \<in> B2 \<Longrightarrow> open U\<close> by auto
have "countable {X. (\<forall>(i::'a). X i \<in> B2) \<and> finite {i. X i \<noteq> UNIV}}"
apply (rule countable_product_event_const) using \<open>countable B2\<close> by auto
moreover have "K = (\<lambda>X. Pi\<^sub>E UNIV X)`{X. (\<forall>i. X i \<in> B2) \<and> finite {i. X i \<noteq> UNIV}}"
unfolding K_def by auto
ultimately have ii: "countable K" by auto
have iii: "topological_basis K"
proof (rule topological_basisI)
fix U and x::"'a\<Rightarrow>'b" assume "open U" "x \<in> U"
then have "openin (product_topology (\<lambda>i. euclidean) UNIV) U"
unfolding open_fun_def by auto
with product_topology_open_contains_basis[OF this \<open>x \<in> U\<close>]
have "\<exists>X. x \<in> (\<Pi>\<^sub>E i\<in>UNIV. X i) \<and> (\<forall>i. open (X i)) \<and> finite {i. X i \<noteq> UNIV} \<and> (\<Pi>\<^sub>E i\<in>UNIV. X i) \<subseteq> U"
by simp
then obtain X where H: "x \<in> (\<Pi>\<^sub>E i\<in>UNIV. X i)"
"\<And>i. open (X i)"
"finite {i. X i \<noteq> UNIV}"
"(\<Pi>\<^sub>E i\<in>UNIV. X i) \<subseteq> U"
by auto
then have "x i \<in> X i" for i by auto
define I where "I = {i. X i \<noteq> UNIV}"
define Y where "Y = (\<lambda>i. if i \<in> I then (SOME y. y \<in> B2 \<and> y \<subseteq> X i \<and> x i \<in> y) else UNIV)"
have *: "\<exists>y. y \<in> B2 \<and> y \<subseteq> X i \<and> x i \<in> y" for i
unfolding B2_def using B \<open>open (X i)\<close> \<open>x i \<in> X i\<close> by (meson UnCI topological_basisE)
have **: "Y i \<in> B2 \<and> Y i \<subseteq> X i \<and> x i \<in> Y i" for i
using someI_ex[OF *]
apply (cases "i \<in> I")
unfolding Y_def using * apply (auto)
unfolding B2_def I_def by auto
have "{i. Y i \<noteq> UNIV} \<subseteq> I"
unfolding Y_def by auto
then have ***: "finite {i. Y i \<noteq> UNIV}"
unfolding I_def using H(3) rev_finite_subset by blast
have "(\<forall>i. Y i \<in> B2) \<and> finite {i. Y i \<noteq> UNIV}"
using ** *** by auto
then have "Pi\<^sub>E UNIV Y \<in> K"
unfolding K_def by auto
have "Y i \<subseteq> X i" for i
apply (cases "i \<in> I") using ** apply simp unfolding Y_def I_def by auto
then have "Pi\<^sub>E UNIV Y \<subseteq> Pi\<^sub>E UNIV X" by auto
then have "Pi\<^sub>E UNIV Y \<subseteq> U" using H(4) by auto
have "x \<in> Pi\<^sub>E UNIV Y"
using ** by auto
show "\<exists>V\<in>K. x \<in> V \<and> V \<subseteq> U"
using \<open>Pi\<^sub>E UNIV Y \<in> K\<close> \<open>Pi\<^sub>E UNIV Y \<subseteq> U\<close> \<open>x \<in> Pi\<^sub>E UNIV Y\<close> by auto
next
fix U assume "U \<in> K"
show "open U"
using \<open>U \<in> K\<close> unfolding open_fun_def K_def by clarify (metis \<open>U \<in> K\<close> i open_PiE open_fun_def)
qed
show ?thesis using i ii iii by auto
qed
instance "fun" :: (countable, second_countable_topology) second_countable_topology
apply standard
using product_topology_countable_basis topological_basis_imp_subbasis by auto
subsection\<open>The Alexander subbase theorem\<close>
theorem Alexander_subbase:
assumes X: "topology (arbitrary union_of (finite intersection_of (\<lambda>x. x \<in> \<B>) relative_to \<Union>\<B>)) = X"
and fin: "\<And>C. \<lbrakk>C \<subseteq> \<B>; \<Union>C = topspace X\<rbrakk> \<Longrightarrow> \<exists>C'. finite C' \<and> C' \<subseteq> C \<and> \<Union>C' = topspace X"
shows "compact_space X"
proof -
have U\<B>: "\<Union>\<B> = topspace X"
by (simp flip: X)
have False if \<U>: "\<forall>U\<in>\<U>. openin X U" and sub: "topspace X \<subseteq> \<Union>\<U>"
and neg: "\<And>\<F>. \<lbrakk>\<F> \<subseteq> \<U>; finite \<F>\<rbrakk> \<Longrightarrow> \<not> topspace X \<subseteq> \<Union>\<F>" for \<U>
proof -
define \<A> where "\<A> \<equiv> {\<C>. (\<forall>U \<in> \<C>. openin X U) \<and> topspace X \<subseteq> \<Union>\<C> \<and> (\<forall>\<F>. finite \<F> \<longrightarrow> \<F> \<subseteq> \<C> \<longrightarrow> ~(topspace X \<subseteq> \<Union>\<F>))}"
have 1: "\<A> \<noteq> {}"
unfolding \<A>_def using sub \<U> neg by force
have 2: "\<Union>\<C> \<in> \<A>" if "\<C>\<noteq>{}" and \<C>: "subset.chain \<A> \<C>" for \<C>
unfolding \<A>_def
proof (intro CollectI conjI ballI allI impI notI)
show "openin X U" if U: "U \<in> \<Union>\<C>" for U
using U \<C> unfolding \<A>_def subset_chain_def by force
have "\<C> \<subseteq> \<A>"
using subset_chain_def \<C> by blast
with that \<A>_def show UUC: "topspace X \<subseteq> \<Union>(\<Union>\<C>)"
by blast
show "False" if "finite \<F>" and "\<F> \<subseteq> \<Union>\<C>" and "topspace X \<subseteq> \<Union>\<F>" for \<F>
proof -
obtain \<B> where "\<B> \<in> \<C>" "\<F> \<subseteq> \<B>"
by (metis Sup_empty \<C> \<open>\<F> \<subseteq> \<Union>\<C>\<close> \<open>finite \<F>\<close> UUC empty_subsetI finite.emptyI finite_subset_Union_chain neg)
then show False
using \<A>_def \<open>\<C> \<subseteq> \<A>\<close> \<open>finite \<F>\<close> \<open>topspace X \<subseteq> \<Union>\<F>\<close> by blast
qed
qed
obtain \<K> where "\<K> \<in> \<A>" and "\<And>X. \<lbrakk>X\<in>\<A>; \<K> \<subseteq> X\<rbrakk> \<Longrightarrow> X = \<K>"
using subset_Zorn_nonempty [OF 1 2] by metis
then have *: "\<And>\<W>. \<lbrakk>\<And>W. W\<in>\<W> \<Longrightarrow> openin X W; topspace X \<subseteq> \<Union>\<W>; \<K> \<subseteq> \<W>;
\<And>\<F>. \<lbrakk>finite \<F>; \<F> \<subseteq> \<W>; topspace X \<subseteq> \<Union>\<F>\<rbrakk> \<Longrightarrow> False\<rbrakk>
\<Longrightarrow> \<W> = \<K>"
and ope: "\<forall>U\<in>\<K>. openin X U" and top: "topspace X \<subseteq> \<Union>\<K>"
and non: "\<And>\<F>. \<lbrakk>finite \<F>; \<F> \<subseteq> \<K>; topspace X \<subseteq> \<Union>\<F>\<rbrakk> \<Longrightarrow> False"
unfolding \<A>_def by simp_all metis+
then obtain x where "x \<in> topspace X" "x \<notin> \<Union>(\<B> \<inter> \<K>)"
proof -
have "\<Union>(\<B> \<inter> \<K>) \<noteq> \<Union>\<B>"
by (metis \<open>\<Union>\<B> = topspace X\<close> fin inf.bounded_iff non order_refl)
then have "\<exists>a. a \<notin> \<Union>(\<B> \<inter> \<K>) \<and> a \<in> \<Union>\<B>"
by blast
then show ?thesis
using that by (metis U\<B>)
qed
obtain C where C: "openin X C" "C \<in> \<K>" "x \<in> C"
using \<open>x \<in> topspace X\<close> ope top by auto
then have "C \<subseteq> topspace X"
by (metis openin_subset)
then have "(arbitrary union_of (finite intersection_of (\<lambda>x. x \<in> \<B>) relative_to \<Union>\<B>)) C"
using openin_subbase C unfolding X [symmetric] by blast
moreover have "C \<noteq> topspace X"
using \<open>\<K> \<in> \<A>\<close> \<open>C \<in> \<K>\<close> unfolding \<A>_def by blast
ultimately obtain \<V> W where W: "(finite intersection_of (\<lambda>x. x \<in> \<B>) relative_to topspace X) W"
and "x \<in> W" "W \<in> \<V>" "\<Union>\<V> \<noteq> topspace X" "C = \<Union>\<V>"
using C by (auto simp: union_of_def U\<B>)
then have "\<Union>\<V> \<subseteq> topspace X"
by (metis \<open>C \<subseteq> topspace X\<close>)
then have "topspace X \<notin> \<V>"
using \<open>\<Union>\<V> \<noteq> topspace X\<close> by blast
then obtain \<B>' where \<B>': "finite \<B>'" "\<B>' \<subseteq> \<B>" "x \<in> \<Inter>\<B>'" "W = topspace X \<inter> \<Inter>\<B>'"
using W \<open>x \<in> W\<close> unfolding relative_to_def intersection_of_def by auto
then have "\<Inter>\<B>' \<subseteq> \<Union>\<B>"
using \<open>W \<in> \<V>\<close> \<open>\<Union>\<V> \<noteq> topspace X\<close> \<open>\<Union>\<V> \<subseteq> topspace X\<close> by blast
then have "\<Inter>\<B>' \<subseteq> C"
using U\<B> \<open>C = \<Union>\<V>\<close> \<open>W = topspace X \<inter> \<Inter>\<B>'\<close> \<open>W \<in> \<V>\<close> by auto
have "\<forall>b \<in> \<B>'. \<exists>C'. finite C' \<and> C' \<subseteq> \<K> \<and> topspace X \<subseteq> \<Union>(insert b C')"
proof
fix b
assume "b \<in> \<B>'"
have "insert b \<K> = \<K>" if neg: "\<not> (\<exists>C'. finite C' \<and> C' \<subseteq> \<K> \<and> topspace X \<subseteq> \<Union>(insert b C'))"
proof (rule *)
show "openin X W" if "W \<in> insert b \<K>" for W
using that
proof
have "b \<in> \<B>"
using \<open>b \<in> \<B>'\<close> \<open>\<B>' \<subseteq> \<B>\<close> by blast
then have "\<exists>\<U>. finite \<U> \<and> \<U> \<subseteq> \<B> \<and> \<Inter>\<U> = b"
by (rule_tac x="{b}" in exI) auto
moreover have "\<Union>\<B> \<inter> b = b"
using \<B>'(2) \<open>b \<in> \<B>'\<close> by auto
ultimately show "openin X W" if "W = b"
using that \<open>b \<in> \<B>'\<close>
apply (simp add: openin_subbase flip: X)
apply (auto simp: arbitrary_def intersection_of_def relative_to_def intro!: union_of_inc)
done
show "openin X W" if "W \<in> \<K>"
by (simp add: \<open>W \<in> \<K>\<close> ope)
qed
next
show "topspace X \<subseteq> \<Union> (insert b \<K>)"
using top by auto
next
show False if "finite \<F>" and "\<F> \<subseteq> insert b \<K>" "topspace X \<subseteq> \<Union>\<F>" for \<F>
proof -
have "insert b (\<F> \<inter> \<K>) = \<F>"
using non that by blast
then show False
by (metis Int_lower2 finite_insert neg that(1) that(3))
qed
qed auto
then show "\<exists>C'. finite C' \<and> C' \<subseteq> \<K> \<and> topspace X \<subseteq> \<Union>(insert b C')"
using \<open>b \<in> \<B>'\<close> \<open>x \<notin> \<Union>(\<B> \<inter> \<K>)\<close> \<B>'
by (metis IntI InterE Union_iff subsetD insertI1)
qed
then obtain F where F: "\<forall>b \<in> \<B>'. finite (F b) \<and> F b \<subseteq> \<K> \<and> topspace X \<subseteq> \<Union>(insert b (F b))"
by metis
let ?\<D> = "insert C (\<Union>(F ` \<B>'))"
show False
proof (rule non)
have "topspace X \<subseteq> (\<Inter>b \<in> \<B>'. \<Union>(insert b (F b)))"
using F by (simp add: INT_greatest)
also have "\<dots> \<subseteq> \<Union>?\<D>"
using \<open>\<Inter>\<B>' \<subseteq> C\<close> by force
finally show "topspace X \<subseteq> \<Union>?\<D>" .
show "?\<D> \<subseteq> \<K>"
using \<open>C \<in> \<K>\<close> F by auto
show "finite ?\<D>"
using \<open>finite \<B>'\<close> F by auto
qed
qed
then show ?thesis
by (force simp: compact_space_def compactin_def)
qed
corollary Alexander_subbase_alt:
assumes "U \<subseteq> \<Union>\<B>"
and fin: "\<And>C. \<lbrakk>C \<subseteq> \<B>; U \<subseteq> \<Union>C\<rbrakk> \<Longrightarrow> \<exists>C'. finite C' \<and> C' \<subseteq> C \<and> U \<subseteq> \<Union>C'"
and X: "topology
(arbitrary union_of
(finite intersection_of (\<lambda>x. x \<in> \<B>) relative_to U)) = X"
shows "compact_space X"
proof -
have "topspace X = U"
using X topspace_subbase by fastforce
have eq: "\<Union> (Collect ((\<lambda>x. x \<in> \<B>) relative_to U)) = U"
unfolding relative_to_def
using \<open>U \<subseteq> \<Union>\<B>\<close> by blast
have *: "\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<C> \<and> \<Union>\<F> = topspace X"
if "\<C> \<subseteq> Collect ((\<lambda>x. x \<in> \<B>) relative_to topspace X)" and UC: "\<Union>\<C> = topspace X" for \<C>
proof -
have "\<C> \<subseteq> (\<lambda>U. topspace X \<inter> U) ` \<B>"
using that by (auto simp: relative_to_def)
then obtain \<B>' where "\<B>' \<subseteq> \<B>" and \<B>': "\<C> = (\<inter>) (topspace X) ` \<B>'"
by (auto simp: subset_image_iff)
moreover have "U \<subseteq> \<Union>\<B>'"
using \<B>' \<open>topspace X = U\<close> UC by auto
ultimately obtain \<C>' where "finite \<C>'" "\<C>' \<subseteq> \<B>'" "U \<subseteq> \<Union>\<C>'"
using fin [of \<B>'] \<open>topspace X = U\<close> \<open>U \<subseteq> \<Union>\<B>\<close> by blast
then show ?thesis
unfolding \<B>' ex_finite_subset_image \<open>topspace X = U\<close> by auto
qed
show ?thesis
apply (rule Alexander_subbase [where \<B> = "Collect ((\<lambda>x. x \<in> \<B>) relative_to (topspace X))"])
apply (simp flip: X)
apply (metis finite_intersection_of_relative_to eq)
apply (blast intro: *)
done
qed
proposition continuous_map_componentwise:
"continuous_map X (product_topology Y I) f \<longleftrightarrow>
f ` (topspace X) \<subseteq> extensional I \<and> (\<forall>k \<in> I. continuous_map X (Y k) (\<lambda>x. f x k))"
(is "?lhs \<longleftrightarrow> _ \<and> ?rhs")
proof (cases "\<forall>x \<in> topspace X. f x \<in> extensional I")
case True
then have "f ` (topspace X) \<subseteq> extensional I"
by force
moreover have ?rhs if L: ?lhs
proof -
have "openin X {x \<in> topspace X. f x k \<in> U}" if "k \<in> I" and "openin (Y k) U" for k U
proof -
have "openin (product_topology Y I) ({Y. Y k \<in> U} \<inter> (\<Pi>\<^sub>E i\<in>I. topspace (Y i)))"
apply (simp add: openin_product_topology flip: arbitrary_union_of_relative_to)
apply (simp add: relative_to_def)
using that apply (blast intro: arbitrary_union_of_inc finite_intersection_of_inc)
done
with that have "openin X {x \<in> topspace X. f x \<in> ({Y. Y k \<in> U} \<inter> (\<Pi>\<^sub>E i\<in>I. topspace (Y i)))}"
using L unfolding continuous_map_def by blast
moreover have "{x \<in> topspace X. f x \<in> ({Y. Y k \<in> U} \<inter> (\<Pi>\<^sub>E i\<in>I. topspace (Y i)))} = {x \<in> topspace X. f x k \<in> U}"
using L by (auto simp: continuous_map_def)
ultimately show ?thesis
by metis
qed
with that
show ?thesis
by (auto simp: continuous_map_def)
qed
moreover have ?lhs if ?rhs
proof -
have 1: "\<And>x. x \<in> topspace X \<Longrightarrow> f x \<in> (\<Pi>\<^sub>E i\<in>I. topspace (Y i))"
using that True by (auto simp: continuous_map_def PiE_iff)
have 2: "{x \<in> S. \<exists>T\<in>\<T>. f x \<in> T} = (\<Union>T\<in>\<T>. {x \<in> S. f x \<in> T})" for S \<T>
by blast
have 3: "{x \<in> S. \<forall>U\<in>\<U>. f x \<in> U} = (\<Inter> (insert S ((\<lambda>U. {x \<in> S. f x \<in> U}) ` \<U>)))" for S \<U>
by blast
show ?thesis
unfolding continuous_map_def openin_product_topology arbitrary_def
proof (clarsimp simp: all_union_of 1 2)
fix \<T>
assume \<T>: "\<T> \<subseteq> Collect (finite intersection_of (\<lambda>F. \<exists>i U. F = {f. f i \<in> U} \<and> i \<in> I \<and> openin (Y i) U)
relative_to (\<Pi>\<^sub>E i\<in>I. topspace (Y i)))"
show "openin X (\<Union>T\<in>\<T>. {x \<in> topspace X. f x \<in> T})"
proof (rule openin_Union; clarify)
fix S T
assume "T \<in> \<T>"
obtain \<U> where "T = (\<Pi>\<^sub>E i\<in>I. topspace (Y i)) \<inter> \<Inter>\<U>" and "finite \<U>"
"\<U> \<subseteq> {{f. f i \<in> U} |i U. i \<in> I \<and> openin (Y i) U}"
using subsetD [OF \<T> \<open>T \<in> \<T>\<close>] by (auto simp: intersection_of_def relative_to_def)
with that show "openin X {x \<in> topspace X. f x \<in> T}"
apply (simp add: continuous_map_def 1 cong: conj_cong)
unfolding 3
apply (rule openin_Inter; auto)
done
qed
qed
qed
ultimately show ?thesis
by metis
next
case False
then show ?thesis
by (auto simp: continuous_map_def PiE_def)
qed
lemma continuous_map_componentwise_UNIV:
"continuous_map X (product_topology Y UNIV) f \<longleftrightarrow> (\<forall>k. continuous_map X (Y k) (\<lambda>x. f x k))"
by (simp add: continuous_map_componentwise)
lemma continuous_map_product_projection [continuous_intros]:
"k \<in> I \<Longrightarrow> continuous_map (product_topology X I) (X k) (\<lambda>x. x k)"
using continuous_map_componentwise [of "product_topology X I" X I id] by simp
declare continuous_map_from_subtopology [OF continuous_map_product_projection, continuous_intros]
proposition open_map_product_projection:
assumes "i \<in> I"
shows "open_map (product_topology Y I) (Y i) (\<lambda>f. f i)"
unfolding openin_product_topology all_union_of arbitrary_def open_map_def image_Union
proof clarify
fix \<V>
assume \<V>: "\<V> \<subseteq> Collect
(finite intersection_of
(\<lambda>F. \<exists>i U. F = {f. f i \<in> U} \<and> i \<in> I \<and> openin (Y i) U) relative_to
topspace (product_topology Y I))"
show "openin (Y i) (\<Union>x\<in>\<V>. (\<lambda>f. f i) ` x)"
proof (rule openin_Union, clarify)
fix S V
assume "V \<in> \<V>"
obtain \<F> where "finite \<F>"
and V: "V = (\<Pi>\<^sub>E i\<in>I. topspace (Y i)) \<inter> \<Inter>\<F>"
and \<F>: "\<F> \<subseteq> {{f. f i \<in> U} |i U. i \<in> I \<and> openin (Y i) U}"
using subsetD [OF \<V> \<open>V \<in> \<V>\<close>]
by (auto simp: intersection_of_def relative_to_def)
show "openin (Y i) ((\<lambda>f. f i) ` V)"
proof (subst openin_subopen; clarify)
fix x f
assume "f \<in> V"
let ?T = "{a \<in> topspace(Y i).
(\<lambda>j. if j = i then a
else if j \<in> I then f j else undefined) \<in> (\<Pi>\<^sub>E i\<in>I. topspace (Y i)) \<inter> \<Inter>\<F>}"
show "\<exists>T. openin (Y i) T \<and> f i \<in> T \<and> T \<subseteq> (\<lambda>f. f i) ` V"
proof (intro exI conjI)
show "openin (Y i) ?T"
proof (rule openin_continuous_map_preimage)
have "continuous_map (Y i) (Y k) (\<lambda>x. if k = i then x else f k)" if "k \<in> I" for k
proof (cases "k=i")
case True
then show ?thesis
by (metis (mono_tags) continuous_map_id eq_id_iff)
next
case False
then show ?thesis
by simp (metis IntD1 PiE_iff V \<open>f \<in> V\<close> that)
qed
then show "continuous_map (Y i) (product_topology Y I)
(\<lambda>x j. if j = i then x else if j \<in> I then f j else undefined)"
by (auto simp: continuous_map_componentwise assms extensional_def)
next
have "openin (product_topology Y I) (\<Pi>\<^sub>E i\<in>I. topspace (Y i))"
by (metis openin_topspace topspace_product_topology)
moreover have "openin (product_topology Y I) (\<Inter>B\<in>\<F>. (\<Pi>\<^sub>E i\<in>I. topspace (Y i)) \<inter> B)"
if "\<F> \<noteq> {}"
proof -
show ?thesis
proof (rule openin_Inter)
show "\<And>X. X \<in> (\<inter>) (\<Pi>\<^sub>E i\<in>I. topspace (Y i)) ` \<F> \<Longrightarrow> openin (product_topology Y I) X"
unfolding openin_product_topology relative_to_def
apply (clarify intro!: arbitrary_union_of_inc)
apply (rename_tac F)
apply (rule_tac x=F in exI)
using subsetD [OF \<F>]
apply (force intro: finite_intersection_of_inc)
done
qed (use \<open>finite \<F>\<close> \<open>\<F> \<noteq> {}\<close> in auto)
qed
ultimately show "openin (product_topology Y I) ((\<Pi>\<^sub>E i\<in>I. topspace (Y i)) \<inter> \<Inter>\<F>)"
by (auto simp only: Int_Inter_eq split: if_split)
qed
next
have eqf: "(\<lambda>j. if j = i then f i else if j \<in> I then f j else undefined) = f"
using PiE_arb V \<open>f \<in> V\<close> by force
show "f i \<in> ?T"
using V assms \<open>f \<in> V\<close> by (auto simp: PiE_iff eqf)
next
show "?T \<subseteq> (\<lambda>f. f i) ` V"
unfolding V by (auto simp: intro!: rev_image_eqI)
qed
qed
qed
qed
lemma retraction_map_product_projection:
assumes "i \<in> I"
shows "(retraction_map (product_topology X I) (X i) (\<lambda>x. x i) \<longleftrightarrow>
(topspace (product_topology X I) = {}) \<longrightarrow> topspace (X i) = {})"
(is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
using retraction_imp_surjective_map by force
next
assume R: ?rhs
show ?lhs
proof (cases "topspace (product_topology X I) = {}")
case True
then show ?thesis
using R by (auto simp: retraction_map_def retraction_maps_def continuous_map_on_empty)
next
case False
have *: "\<exists>g. continuous_map (X i) (product_topology X I) g \<and> (\<forall>x\<in>topspace (X i). g x i = x)"
if z: "z \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))" for z
proof -
have cm: "continuous_map (X i) (X j) (\<lambda>x. if j = i then x else z j)" if "j \<in> I" for j
using \<open>j \<in> I\<close> z by (case_tac "j = i") auto
show ?thesis
using \<open>i \<in> I\<close> that
by (rule_tac x="\<lambda>x j. if j = i then x else z j" in exI) (auto simp: continuous_map_componentwise PiE_iff extensional_def cm)
qed
show ?thesis
using \<open>i \<in> I\<close> False
by (auto simp: retraction_map_def retraction_maps_def assms continuous_map_product_projection *)
qed
qed
subsection \<open>Open Pi-sets in the product topology\<close>
proposition openin_PiE_gen:
"openin (product_topology X I) (PiE I S) \<longleftrightarrow>
PiE I S = {} \<or>
finite {i \<in> I. ~(S i = topspace(X i))} \<and> (\<forall>i \<in> I. openin (X i) (S i))"
(is "?lhs \<longleftrightarrow> _ \<or> ?rhs")
proof (cases "PiE I S = {}")
case False
moreover have "?lhs = ?rhs"
proof
assume L: ?lhs
moreover
obtain z where z: "z \<in> PiE I S"
using False by blast
ultimately obtain U where fin: "finite {i \<in> I. U i \<noteq> topspace (X i)}"
and "Pi\<^sub>E I U \<noteq> {}"
and sub: "Pi\<^sub>E I U \<subseteq> Pi\<^sub>E I S"
by (fastforce simp add: openin_product_topology_alt)
then have *: "\<And>i. i \<in> I \<Longrightarrow> U i \<subseteq> S i"
by (simp add: subset_PiE)
show ?rhs
proof (intro conjI ballI)
show "finite {i \<in> I. S i \<noteq> topspace (X i)}"
apply (rule finite_subset [OF _ fin], clarify)
using *
by (metis False L openin_subset topspace_product_topology subset_PiE subset_antisym)
next
fix i :: "'a"
assume "i \<in> I"
then show "openin (X i) (S i)"
using open_map_product_projection [of i I X] L
apply (simp add: open_map_def)
apply (drule_tac x="PiE I S" in spec)
apply (simp add: False image_projection_PiE split: if_split_asm)
done
qed
next
assume ?rhs
then show ?lhs
apply (simp only: openin_product_topology)
apply (rule arbitrary_union_of_inc)
apply (auto simp: product_topology_base_alt)
done
qed
ultimately show ?thesis
by simp
qed simp
corollary openin_PiE:
"finite I \<Longrightarrow> openin (product_topology X I) (PiE I S) \<longleftrightarrow> PiE I S = {} \<or> (\<forall>i \<in> I. openin (X i) (S i))"
by (simp add: openin_PiE_gen)
proposition compact_space_product_topology:
"compact_space(product_topology X I) \<longleftrightarrow>
topspace(product_topology X I) = {} \<or> (\<forall>i \<in> I. compact_space(X i))"
(is "?lhs = ?rhs")
proof (cases "topspace(product_topology X I) = {}")
case False
then obtain z where z: "z \<in> (\<Pi>\<^sub>E i\<in>I. topspace(X i))"
by auto
show ?thesis
proof
assume L: ?lhs
show ?rhs
proof (clarsimp simp add: False compact_space_def)
fix i
assume "i \<in> I"
with L have "continuous_map (product_topology X I) (X i) (\<lambda>f. f i)"
by (simp add: continuous_map_product_projection)
moreover
have "\<And>x. x \<in> topspace (X i) \<Longrightarrow> x \<in> (\<lambda>f. f i) ` (\<Pi>\<^sub>E i\<in>I. topspace (X i))"
using \<open>i \<in> I\<close> z
apply (rule_tac x="\<lambda>j. if j = i then x else if j \<in> I then z j else undefined" in image_eqI, auto)
done
then have "(\<lambda>f. f i) ` (\<Pi>\<^sub>E i\<in>I. topspace (X i)) = topspace (X i)"
using \<open>i \<in> I\<close> z by auto
ultimately show "compactin (X i) (topspace (X i))"
by (metis L compact_space_def image_compactin topspace_product_topology)
qed
next
assume R: ?rhs
show ?lhs
proof (cases "I = {}")
case True
with R show ?thesis
by (simp add: compact_space_def)
next
case False
then obtain i where "i \<in> I"
by blast
show ?thesis
using R
proof
assume com [rule_format]: "\<forall>i\<in>I. compact_space (X i)"
let ?\<C> = "{{f. f i \<in> U} |i U. i \<in> I \<and> openin (X i) U}"
show "compact_space (product_topology X I)"
proof (rule Alexander_subbase_alt)
show "topspace (product_topology X I) \<subseteq> \<Union>?\<C>"
unfolding topspace_product_topology using \<open>i \<in> I\<close> by blast
next
fix C
assume Csub: "C \<subseteq> ?\<C>" and UC: "topspace (product_topology X I) \<subseteq> \<Union>C"
define \<D> where "\<D> \<equiv> \<lambda>i. {U. openin (X i) U \<and> {f. f i \<in> U} \<in> C}"
show "\<exists>C'. finite C' \<and> C' \<subseteq> C \<and> topspace (product_topology X I) \<subseteq> \<Union>C'"
proof (cases "\<exists>i. i \<in> I \<and> topspace (X i) \<subseteq> \<Union>(\<D> i)")
case True
then obtain i where "i \<in> I"
and i: "topspace (X i) \<subseteq> \<Union>(\<D> i)"
unfolding \<D>_def by blast
then have *: "\<And>\<U>. \<lbrakk>Ball \<U> (openin (X i)); topspace (X i) \<subseteq> \<Union>\<U>\<rbrakk> \<Longrightarrow>
\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> topspace (X i) \<subseteq> \<Union>\<F>"
using com [OF \<open>i \<in> I\<close>] by (auto simp: compact_space_def compactin_def)
have "topspace (X i) \<subseteq> \<Union>(\<D> i)"
using i by auto
with * obtain \<F> where "finite \<F> \<and> \<F> \<subseteq> (\<D> i) \<and> topspace (X i) \<subseteq> \<Union>\<F>"
unfolding \<D>_def by fastforce
with \<open>i \<in> I\<close> show ?thesis
unfolding \<D>_def
by (rule_tac x="(\<lambda>U. {x. x i \<in> U}) ` \<F>" in exI) auto
next
case False
then have "\<forall>i \<in> I. \<exists>y. y \<in> topspace (X i) \<and> y \<notin> \<Union>(\<D> i)"
by force
then obtain g where g: "\<And>i. i \<in> I \<Longrightarrow> g i \<in> topspace (X i) \<and> g i \<notin> \<Union>(\<D> i)"
by metis
then have "(\<lambda>i. if i \<in> I then g i else undefined) \<in> topspace (product_topology X I)"
by (simp add: PiE_I)
moreover have "(\<lambda>i. if i \<in> I then g i else undefined) \<notin> \<Union>C"
using Csub g unfolding \<D>_def by force
ultimately show ?thesis
using UC by blast
qed
qed (simp add: product_topology)
qed (simp add: compact_space_topspace_empty)
qed
qed
qed (simp add: compact_space_topspace_empty)
corollary compactin_PiE:
"compactin (product_topology X I) (PiE I S) \<longleftrightarrow>
PiE I S = {} \<or> (\<forall>i \<in> I. compactin (X i) (S i))"
by (auto simp: compactin_subspace subtopology_PiE subset_PiE compact_space_product_topology
PiE_eq_empty_iff)
lemma in_product_topology_closure_of:
"z \<in> (product_topology X I) closure_of S
\<Longrightarrow> i \<in> I \<Longrightarrow> z i \<in> ((X i) closure_of ((\<lambda>x. x i) ` S))"
using continuous_map_product_projection
by (force simp: continuous_map_eq_image_closure_subset image_subset_iff)
lemma homeomorphic_space_singleton_product:
"product_topology X {k} homeomorphic_space (X k)"
unfolding homeomorphic_space
apply (rule_tac x="\<lambda>x. x k" in exI)
apply (rule bijective_open_imp_homeomorphic_map)
apply (simp_all add: continuous_map_product_projection open_map_product_projection)
unfolding PiE_over_singleton_iff
apply (auto simp: image_iff inj_on_def)
done
subsection\<open>Relationship with connected spaces, paths, etc.\<close>
proposition connected_space_product_topology:
"connected_space(product_topology X I) \<longleftrightarrow>
(\<Pi>\<^sub>E i\<in>I. topspace (X i)) = {} \<or> (\<forall>i \<in> I. connected_space(X i))"
(is "?lhs \<longleftrightarrow> ?eq \<or> ?rhs")
proof (cases ?eq)
case False
moreover have "?lhs = ?rhs"
proof
assume ?lhs
moreover
have "connectedin(X i) (topspace(X i))"
if "i \<in> I" and ci: "connectedin(product_topology X I) (topspace(product_topology X I))" for i
proof -
have cm: "continuous_map (product_topology X I) (X i) (\<lambda>f. f i)"
by (simp add: \<open>i \<in> I\<close> continuous_map_product_projection)
show ?thesis
using connectedin_continuous_map_image [OF cm ci] \<open>i \<in> I\<close>
by (simp add: False image_projection_PiE)
qed
ultimately show ?rhs
by (meson connectedin_topspace)
next
assume cs [rule_format]: ?rhs
have False
if disj: "U \<inter> V = {}" and subUV: "(\<Pi>\<^sub>E i\<in>I. topspace (X i)) \<subseteq> U \<union> V"
and U: "openin (product_topology X I) U"
and V: "openin (product_topology X I) V"
and "U \<noteq> {}" "V \<noteq> {}"
for U V
proof -
obtain f where "f \<in> U"
using \<open>U \<noteq> {}\<close> by blast
then have f: "f \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))"
using U openin_subset by fastforce
have "U \<subseteq> topspace(product_topology X I)" "V \<subseteq> topspace(product_topology X I)"
using U V openin_subset by blast+
moreover have "(\<Pi>\<^sub>E i\<in>I. topspace (X i)) \<subseteq> U"
proof -
obtain C where "(finite intersection_of (\<lambda>F. \<exists>i U. F = {x. x i \<in> U} \<and> i \<in> I \<and> openin (X i) U) relative_to
(\<Pi>\<^sub>E i\<in>I. topspace (X i))) C" "C \<subseteq> U" "f \<in> C"
using U \<open>f \<in> U\<close> unfolding openin_product_topology union_of_def by auto
then obtain \<T> where "finite \<T>"
and t: "\<And>C. C \<in> \<T> \<Longrightarrow> \<exists>i u. (i \<in> I \<and> openin (X i) u) \<and> C = {x. x i \<in> u}"
and subU: "topspace (product_topology X I) \<inter> \<Inter>\<T> \<subseteq> U"
and ftop: "f \<in> topspace (product_topology X I)"
and fint: "f \<in> \<Inter> \<T>"
by (fastforce simp: relative_to_def intersection_of_def subset_iff)
let ?L = "\<Union>C\<in>\<T>. {i. (\<lambda>x. x i) ` C \<subset> topspace (X i)}"
obtain L where "finite L"
and L: "\<And>i U. \<lbrakk>i \<in> I; openin (X i) U; U \<subset> topspace(X i); {x. x i \<in> U} \<in> \<T>\<rbrakk> \<Longrightarrow> i \<in> L"
proof
show "finite ?L"
proof (rule finite_Union)
fix M
assume "M \<in> (\<lambda>C. {i. (\<lambda>x. x i) ` C \<subset> topspace (X i)}) ` \<T>"
then obtain C where "C \<in> \<T>" and C: "M = {i. (\<lambda>x. x i) ` C \<subset> topspace (X i)}"
by blast
then obtain j V where "j \<in> I" and ope: "openin (X j) V" and Ceq: "C = {x. x j \<in> V}"
using t by meson
then have "f j \<in> V"
using \<open>C \<in> \<T>\<close> fint by force
then have "(\<lambda>x. x k) ` {x. x j \<in> V} = UNIV" if "k \<noteq> j" for k
using that
apply (clarsimp simp add: set_eq_iff)
apply (rule_tac x="\<lambda>m. if m = k then x else f m" in image_eqI, auto)
done
then have "{i. (\<lambda>x. x i) ` C \<subset> topspace (X i)} \<subseteq> {j}"
using Ceq by auto
then show "finite M"
using C finite_subset by fastforce
qed (use \<open>finite \<T>\<close> in blast)
next
fix i U
assume "i \<in> I" and ope: "openin (X i) U" and psub: "U \<subset> topspace (X i)" and int: "{x. x i \<in> U} \<in> \<T>"
then show "i \<in> ?L"
by (rule_tac a="{x. x i \<in> U}" in UN_I) (force+)
qed
show ?thesis
proof
fix h
assume h: "h \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))"
define g where "g \<equiv> \<lambda>i. if i \<in> L then f i else h i"
have gin: "g \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))"
unfolding g_def using f h by auto
moreover have "g \<in> X" if "X \<in> \<T>" for X
using fint openin_subset t [OF that] L g_def h that by fastforce
ultimately have "g \<in> U"
using subU by auto
have "h \<in> U" if "finite M" "h \<in> PiE I (topspace \<circ> X)" "{i \<in> I. h i \<noteq> g i} \<subseteq> M" for M h
using that
proof (induction arbitrary: h)
case empty
then show ?case
using PiE_ext \<open>g \<in> U\<close> gin by force
next
case (insert i M)
define f where "f \<equiv> \<lambda>j. if j = i then g i else h j"
have fin: "f \<in> PiE I (topspace \<circ> X)"
unfolding f_def using gin insert.prems(1) by auto
have subM: "{j \<in> I. f j \<noteq> g j} \<subseteq> M"
unfolding f_def using insert.prems(2) by auto
have "f \<in> U"
using insert.IH [OF fin subM] .
show ?case
proof (cases "h \<in> V")
case True
show ?thesis
proof (cases "i \<in> I")
case True
let ?U = "{x \<in> topspace(X i). (\<lambda>j. if j = i then x else h j) \<in> U}"
let ?V = "{x \<in> topspace(X i). (\<lambda>j. if j = i then x else h j) \<in> V}"
have False
proof (rule connected_spaceD [OF cs [OF \<open>i \<in> I\<close>]])
have "\<And>k. k \<in> I \<Longrightarrow> continuous_map (X i) (X k) (\<lambda>x. if k = i then x else h k)"
using continuous_map_eq_topcontinuous_at insert.prems(1) topcontinuous_at_def by fastforce
then have cm: "continuous_map (X i) (product_topology X I) (\<lambda>x j. if j = i then x else h j)"
using \<open>i \<in> I\<close> insert.prems(1)
by (auto simp: continuous_map_componentwise extensional_def)
show "openin (X i) ?U"
by (rule openin_continuous_map_preimage [OF cm U])
show "openin (X i) ?V"
by (rule openin_continuous_map_preimage [OF cm V])
show "topspace (X i) \<subseteq> ?U \<union> ?V"
proof clarsimp
fix x
assume "x \<in> topspace (X i)" and "(\<lambda>j. if j = i then x else h j) \<notin> V"
with True subUV \<open>h \<in> Pi\<^sub>E I (topspace \<circ> X)\<close>
show "(\<lambda>j. if j = i then x else h j) \<in> U"
by (drule_tac c="(\<lambda>j. if j = i then x else h j)" in subsetD) auto
qed
show "?U \<inter> ?V = {}"
using disj by blast
show "?U \<noteq> {}"
using \<open>U \<noteq> {}\<close> f_def
proof -
have "(\<lambda>j. if j = i then g i else h j) \<in> U"
using \<open>f \<in> U\<close> f_def by blast
moreover have "f i \<in> topspace (X i)"
by (metis PiE_iff True comp_apply fin)
ultimately have "\<exists>b. b \<in> topspace (X i) \<and> (\<lambda>a. if a = i then b else h a) \<in> U"
using f_def by auto
then show ?thesis
by blast
qed
have "(\<lambda>j. if j = i then h i else h j) = h"
by force
moreover have "h i \<in> topspace (X i)"
using True insert.prems(1) by auto
ultimately show "?V \<noteq> {}"
using \<open>h \<in> V\<close> by force
qed
then show ?thesis ..
next
case False
show ?thesis
proof (cases "h = f")
case True
show ?thesis
by (rule insert.IH [OF insert.prems(1)]) (simp add: True subM)
next
case False
then show ?thesis
using gin insert.prems(1) \<open>i \<notin> I\<close> unfolding f_def by fastforce
qed
qed
next
case False
then show ?thesis
using subUV insert.prems(1) by auto
qed
qed
then show "h \<in> U"
unfolding g_def using PiE_iff \<open>finite L\<close> h by fastforce
qed
qed
ultimately show ?thesis
using disj inf_absorb2 \<open>V \<noteq> {}\<close> by fastforce
qed
then show ?lhs
unfolding connected_space_def
by auto
qed
ultimately show ?thesis
by simp
qed (simp add: connected_space_topspace_empty)
lemma connectedin_PiE:
"connectedin (product_topology X I) (PiE I S) \<longleftrightarrow>
PiE I S = {} \<or> (\<forall>i \<in> I. connectedin (X i) (S i))"
by (fastforce simp add: connectedin_def subtopology_PiE connected_space_product_topology subset_PiE PiE_eq_empty_iff)
lemma path_connected_space_product_topology:
"path_connected_space(product_topology X I) \<longleftrightarrow>
topspace(product_topology X I) = {} \<or> (\<forall>i \<in> I. path_connected_space(X i))"
(is "?lhs \<longleftrightarrow> ?eq \<or> ?rhs")
proof (cases ?eq)
case False
moreover have "?lhs = ?rhs"
proof
assume L: ?lhs
show ?rhs
proof (clarsimp simp flip: path_connectedin_topspace)
fix i :: "'a"
assume "i \<in> I"
have cm: "continuous_map (product_topology X I) (X i) (\<lambda>f. f i)"
by (simp add: \<open>i \<in> I\<close> continuous_map_product_projection)
show "path_connectedin (X i) (topspace (X i))"
using path_connectedin_continuous_map_image [OF cm L [unfolded path_connectedin_topspace [symmetric]]]
by (metis \<open>i \<in> I\<close> False retraction_imp_surjective_map retraction_map_product_projection)
qed
next
assume R [rule_format]: ?rhs
show ?lhs
unfolding path_connected_space_def topspace_product_topology
proof clarify
fix x y
assume x: "x \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))" and y: "y \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))"
have "\<forall>i. \<exists>g. i \<in> I \<longrightarrow> pathin (X i) g \<and> g 0 = x i \<and> g 1 = y i"
using PiE_mem R path_connected_space_def x y by force
then obtain g where g: "\<And>i. i \<in> I \<Longrightarrow> pathin (X i) (g i) \<and> g i 0 = x i \<and> g i 1 = y i"
by metis
with x y show "\<exists>g. pathin (product_topology X I) g \<and> g 0 = x \<and> g 1 = y"
apply (rule_tac x="\<lambda>a. \<lambda>i \<in> I. g i a" in exI)
apply (force simp: pathin_def continuous_map_componentwise)
done
qed
qed
ultimately show ?thesis
by simp
qed (simp add: path_connected_space_topspace_empty)
lemma path_connectedin_PiE:
"path_connectedin (product_topology X I) (PiE I S) \<longleftrightarrow>
PiE I S = {} \<or> (\<forall>i \<in> I. path_connectedin (X i) (S i))"
by (fastforce simp add: path_connectedin_def subtopology_PiE path_connected_space_product_topology subset_PiE PiE_eq_empty_iff topspace_subtopology_subset)
subsection \<open>Projections from a function topology to a component\<close>
lemma quotient_map_product_projection:
assumes "i \<in> I"
shows "quotient_map(product_topology X I) (X i) (\<lambda>x. x i) \<longleftrightarrow>
(topspace(product_topology X I) = {} \<longrightarrow> topspace(X i) = {})"
(is "?lhs = ?rhs")
proof
assume ?lhs with assms show ?rhs
by (auto simp: continuous_open_quotient_map open_map_product_projection)
next
assume ?rhs with assms show ?lhs
by (auto simp: Abstract_Topology.retraction_imp_quotient_map retraction_map_product_projection)
qed
lemma product_topology_homeomorphic_component:
assumes "i \<in> I" "\<And>j. \<lbrakk>j \<in> I; j \<noteq> i\<rbrakk> \<Longrightarrow> \<exists>a. topspace(X j) = {a}"
shows "product_topology X I homeomorphic_space (X i)"
proof -
have "quotient_map (product_topology X I) (X i) (\<lambda>x. x i)"
using assms by (force simp add: quotient_map_product_projection PiE_eq_empty_iff)
moreover
have "inj_on (\<lambda>x. x i) (\<Pi>\<^sub>E i\<in>I. topspace (X i))"
using assms by (auto simp: inj_on_def PiE_iff) (metis extensionalityI singletonD)
ultimately show ?thesis
unfolding homeomorphic_space_def
by (rule_tac x="\<lambda>x. x i" in exI) (simp add: homeomorphic_map_def flip: homeomorphic_map_maps)
qed
lemma topological_property_of_product_component:
assumes major: "P (product_topology X I)"
and minor: "\<And>z i. \<lbrakk>z \<in> (\<Pi>\<^sub>E i\<in>I. topspace(X i)); P(product_topology X I); i \<in> I\<rbrakk>
\<Longrightarrow> P(subtopology (product_topology X I) (PiE I (\<lambda>j. if j = i then topspace(X i) else {z j})))"
(is "\<And>z i. \<lbrakk>_; _; _\<rbrakk> \<Longrightarrow> P (?SX z i)")
and PQ: "\<And>X X'. X homeomorphic_space X' \<Longrightarrow> (P X \<longleftrightarrow> Q X')"
shows "(\<Pi>\<^sub>E i\<in>I. topspace(X i)) = {} \<or> (\<forall>i \<in> I. Q(X i))"
proof -
have "Q(X i)" if "(\<Pi>\<^sub>E i\<in>I. topspace(X i)) \<noteq> {}" "i \<in> I" for i
proof -
from that obtain f where f: "f \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))"
by force
have "?SX f i homeomorphic_space X i"
apply (simp add: subtopology_PiE )
using product_topology_homeomorphic_component [OF \<open>i \<in> I\<close>, of "\<lambda>j. subtopology (X j) (if j = i then topspace (X i) else {f j})"]
using f by fastforce
then show ?thesis
using minor [OF f major \<open>i \<in> I\<close>] PQ by auto
qed
then show ?thesis by metis
qed
end
diff --git a/src/HOL/Analysis/Further_Topology.thy b/src/HOL/Analysis/Further_Topology.thy
--- a/src/HOL/Analysis/Further_Topology.thy
+++ b/src/HOL/Analysis/Further_Topology.thy
@@ -1,5712 +1,5712 @@
section \<open>Extending Continous Maps, Invariance of Domain, etc\<close> (*FIX rename? *)
text\<open>Ported from HOL Light (moretop.ml) by L C Paulson\<close>
theory Further_Topology
imports Weierstrass_Theorems Polytope Complex_Transcendental Equivalence_Lebesgue_Henstock_Integration Retracts
begin
subsection\<open>A map from a sphere to a higher dimensional sphere is nullhomotopic\<close>
lemma spheremap_lemma1:
fixes f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space"
assumes "subspace S" "subspace T" and dimST: "dim S < dim T"
and "S \<subseteq> T"
and diff_f: "f differentiable_on sphere 0 1 \<inter> S"
shows "f ` (sphere 0 1 \<inter> S) \<noteq> sphere 0 1 \<inter> T"
proof
assume fim: "f ` (sphere 0 1 \<inter> S) = sphere 0 1 \<inter> T"
have inS: "\<And>x. \<lbrakk>x \<in> S; x \<noteq> 0\<rbrakk> \<Longrightarrow> (x /\<^sub>R norm x) \<in> S"
using subspace_mul \<open>subspace S\<close> by blast
have subS01: "(\<lambda>x. x /\<^sub>R norm x) ` (S - {0}) \<subseteq> sphere 0 1 \<inter> S"
using \<open>subspace S\<close> subspace_mul by fastforce
then have diff_f': "f differentiable_on (\<lambda>x. x /\<^sub>R norm x) ` (S - {0})"
by (rule differentiable_on_subset [OF diff_f])
define g where "g \<equiv> \<lambda>x. norm x *\<^sub>R f(inverse(norm x) *\<^sub>R x)"
have gdiff: "g differentiable_on S - {0}"
unfolding g_def
by (rule diff_f' derivative_intros differentiable_on_compose [where f=f] | force)+
have geq: "g ` (S - {0}) = T - {0}"
proof
have "g ` (S - {0}) \<subseteq> T"
apply (auto simp: g_def subspace_mul [OF \<open>subspace T\<close>])
apply (metis (mono_tags, lifting) DiffI subS01 subspace_mul [OF \<open>subspace T\<close>] fim image_subset_iff inf_le2 singletonD)
done
moreover have "g ` (S - {0}) \<subseteq> UNIV - {0}"
proof (clarsimp simp: g_def)
fix y
assume "y \<in> S" and f0: "f (y /\<^sub>R norm y) = 0"
then have "y \<noteq> 0 \<Longrightarrow> y /\<^sub>R norm y \<in> sphere 0 1 \<inter> S"
by (auto simp: subspace_mul [OF \<open>subspace S\<close>])
then show "y = 0"
by (metis fim f0 Int_iff image_iff mem_sphere_0 norm_eq_zero zero_neq_one)
qed
ultimately show "g ` (S - {0}) \<subseteq> T - {0}"
by auto
next
have *: "sphere 0 1 \<inter> T \<subseteq> f ` (sphere 0 1 \<inter> S)"
using fim by (simp add: image_subset_iff)
have "x \<in> (\<lambda>x. norm x *\<^sub>R f (x /\<^sub>R norm x)) ` (S - {0})"
if "x \<in> T" "x \<noteq> 0" for x
proof -
have "x /\<^sub>R norm x \<in> T"
using \<open>subspace T\<close> subspace_mul that by blast
then show ?thesis
using * [THEN subsetD, of "x /\<^sub>R norm x"] that apply clarsimp
apply (rule_tac x="norm x *\<^sub>R xa" in image_eqI, simp)
apply (metis norm_eq_zero right_inverse scaleR_one scaleR_scaleR)
using \<open>subspace S\<close> subspace_mul apply force
done
qed
then have "T - {0} \<subseteq> (\<lambda>x. norm x *\<^sub>R f (x /\<^sub>R norm x)) ` (S - {0})"
by force
then show "T - {0} \<subseteq> g ` (S - {0})"
by (simp add: g_def)
qed
define T' where "T' \<equiv> {y. \<forall>x \<in> T. orthogonal x y}"
have "subspace T'"
by (simp add: subspace_orthogonal_to_vectors T'_def)
have dim_eq: "dim T' + dim T = DIM('a)"
using dim_subspace_orthogonal_to_vectors [of T UNIV] \<open>subspace T\<close>
by (simp add: T'_def)
have "\<exists>v1 v2. v1 \<in> span T \<and> (\<forall>w \<in> span T. orthogonal v2 w) \<and> x = v1 + v2" for x
by (force intro: orthogonal_subspace_decomp_exists [of T x])
then obtain p1 p2 where p1span: "p1 x \<in> span T"
and "\<And>w. w \<in> span T \<Longrightarrow> orthogonal (p2 x) w"
and eq: "p1 x + p2 x = x" for x
by metis
then have p1: "\<And>z. p1 z \<in> T" and ortho: "\<And>w. w \<in> T \<Longrightarrow> orthogonal (p2 x) w" for x
using span_eq_iff \<open>subspace T\<close> by blast+
then have p2: "\<And>z. p2 z \<in> T'"
by (simp add: T'_def orthogonal_commute)
have p12_eq: "\<And>x y. \<lbrakk>x \<in> T; y \<in> T'\<rbrakk> \<Longrightarrow> p1(x + y) = x \<and> p2(x + y) = y"
proof (rule orthogonal_subspace_decomp_unique [OF eq p1span, where T=T'])
show "\<And>x y. \<lbrakk>x \<in> T; y \<in> T'\<rbrakk> \<Longrightarrow> p2 (x + y) \<in> span T'"
using span_eq_iff p2 \<open>subspace T'\<close> by blast
show "\<And>a b. \<lbrakk>a \<in> T; b \<in> T'\<rbrakk> \<Longrightarrow> orthogonal a b"
using T'_def by blast
qed (auto simp: span_base)
then have "\<And>c x. p1 (c *\<^sub>R x) = c *\<^sub>R p1 x \<and> p2 (c *\<^sub>R x) = c *\<^sub>R p2 x"
proof -
fix c :: real and x :: 'a
have f1: "c *\<^sub>R x = c *\<^sub>R p1 x + c *\<^sub>R p2 x"
by (metis eq pth_6)
have f2: "c *\<^sub>R p2 x \<in> T'"
by (simp add: \<open>subspace T'\<close> p2 subspace_scale)
have "c *\<^sub>R p1 x \<in> T"
by (metis (full_types) assms(2) p1span span_eq_iff subspace_scale)
then show "p1 (c *\<^sub>R x) = c *\<^sub>R p1 x \<and> p2 (c *\<^sub>R x) = c *\<^sub>R p2 x"
using f2 f1 p12_eq by presburger
qed
moreover have lin_add: "\<And>x y. p1 (x + y) = p1 x + p1 y \<and> p2 (x + y) = p2 x + p2 y"
proof (rule orthogonal_subspace_decomp_unique [OF _ p1span, where T=T'])
show "\<And>x y. p1 (x + y) + p2 (x + y) = p1 x + p1 y + (p2 x + p2 y)"
by (simp add: add.assoc add.left_commute eq)
show "\<And>a b. \<lbrakk>a \<in> T; b \<in> T'\<rbrakk> \<Longrightarrow> orthogonal a b"
using T'_def by blast
qed (auto simp: p1span p2 span_base span_add)
ultimately have "linear p1" "linear p2"
by unfold_locales auto
have "(\<lambda>z. g (p1 z)) differentiable_on {x + y |x y. x \<in> S - {0} \<and> y \<in> T'}"
apply (rule differentiable_on_compose [where f=g])
apply (rule linear_imp_differentiable_on [OF \<open>linear p1\<close>])
apply (rule differentiable_on_subset [OF gdiff])
using p12_eq \<open>S \<subseteq> T\<close> apply auto
done
then have diff: "(\<lambda>x. g (p1 x) + p2 x) differentiable_on {x + y |x y. x \<in> S - {0} \<and> y \<in> T'}"
by (intro derivative_intros linear_imp_differentiable_on [OF \<open>linear p2\<close>])
have "dim {x + y |x y. x \<in> S - {0} \<and> y \<in> T'} \<le> dim {x + y |x y. x \<in> S \<and> y \<in> T'}"
by (blast intro: dim_subset)
also have "... = dim S + dim T' - dim (S \<inter> T')"
using dim_sums_Int [OF \<open>subspace S\<close> \<open>subspace T'\<close>]
by (simp add: algebra_simps)
also have "... < DIM('a)"
using dimST dim_eq by auto
finally have neg: "negligible {x + y |x y. x \<in> S - {0} \<and> y \<in> T'}"
by (rule negligible_lowdim)
have "negligible ((\<lambda>x. g (p1 x) + p2 x) ` {x + y |x y. x \<in> S - {0} \<and> y \<in> T'})"
by (rule negligible_differentiable_image_negligible [OF order_refl neg diff])
then have "negligible {x + y |x y. x \<in> g ` (S - {0}) \<and> y \<in> T'}"
proof (rule negligible_subset)
have "\<lbrakk>t' \<in> T'; s \<in> S; s \<noteq> 0\<rbrakk>
\<Longrightarrow> g s + t' \<in> (\<lambda>x. g (p1 x) + p2 x) `
{x + t' |x t'. x \<in> S \<and> x \<noteq> 0 \<and> t' \<in> T'}" for t' s
apply (rule_tac x="s + t'" in image_eqI)
using \<open>S \<subseteq> T\<close> p12_eq by auto
then show "{x + y |x y. x \<in> g ` (S - {0}) \<and> y \<in> T'}
\<subseteq> (\<lambda>x. g (p1 x) + p2 x) ` {x + y |x y. x \<in> S - {0} \<and> y \<in> T'}"
by auto
qed
moreover have "- T' \<subseteq> {x + y |x y. x \<in> g ` (S - {0}) \<and> y \<in> T'}"
proof clarsimp
fix z assume "z \<notin> T'"
show "\<exists>x y. z = x + y \<and> x \<in> g ` (S - {0}) \<and> y \<in> T'"
apply (rule_tac x="p1 z" in exI)
apply (rule_tac x="p2 z" in exI)
apply (simp add: p1 eq p2 geq)
by (metis \<open>z \<notin> T'\<close> add.left_neutral eq p2)
qed
ultimately have "negligible (-T')"
using negligible_subset by blast
moreover have "negligible T'"
using negligible_lowdim
by (metis add.commute assms(3) diff_add_inverse2 diff_self_eq_0 dim_eq le_add1 le_antisym linordered_semidom_class.add_diff_inverse not_less0)
ultimately have "negligible (-T' \<union> T')"
by (metis negligible_Un_eq)
then show False
using negligible_Un_eq non_negligible_UNIV by simp
qed
lemma spheremap_lemma2:
fixes f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space"
assumes ST: "subspace S" "subspace T" "dim S < dim T"
and "S \<subseteq> T"
and contf: "continuous_on (sphere 0 1 \<inter> S) f"
and fim: "f ` (sphere 0 1 \<inter> S) \<subseteq> sphere 0 1 \<inter> T"
shows "\<exists>c. homotopic_with_canon (\<lambda>x. True) (sphere 0 1 \<inter> S) (sphere 0 1 \<inter> T) f (\<lambda>x. c)"
proof -
have [simp]: "\<And>x. \<lbrakk>norm x = 1; x \<in> S\<rbrakk> \<Longrightarrow> norm (f x) = 1"
using fim by (simp add: image_subset_iff)
have "compact (sphere 0 1 \<inter> S)"
by (simp add: \<open>subspace S\<close> closed_subspace compact_Int_closed)
then obtain g where pfg: "polynomial_function g" and gim: "g ` (sphere 0 1 \<inter> S) \<subseteq> T"
and g12: "\<And>x. x \<in> sphere 0 1 \<inter> S \<Longrightarrow> norm(f x - g x) < 1/2"
apply (rule Stone_Weierstrass_polynomial_function_subspace [OF _ contf _ \<open>subspace T\<close>, of "1/2"])
using fim apply auto
done
have gnz: "g x \<noteq> 0" if "x \<in> sphere 0 1 \<inter> S" for x
proof -
have "norm (f x) = 1"
using fim that by (simp add: image_subset_iff)
then show ?thesis
using g12 [OF that] by auto
qed
have diffg: "g differentiable_on sphere 0 1 \<inter> S"
by (metis pfg differentiable_on_polynomial_function)
define h where "h \<equiv> \<lambda>x. inverse(norm(g x)) *\<^sub>R g x"
have h: "x \<in> sphere 0 1 \<inter> S \<Longrightarrow> h x \<in> sphere 0 1 \<inter> T" for x
unfolding h_def
using gnz [of x]
by (auto simp: subspace_mul [OF \<open>subspace T\<close>] subsetD [OF gim])
have diffh: "h differentiable_on sphere 0 1 \<inter> S"
unfolding h_def
apply (intro derivative_intros diffg differentiable_on_compose [OF diffg])
using gnz apply auto
done
have homfg: "homotopic_with_canon (\<lambda>z. True) (sphere 0 1 \<inter> S) (T - {0}) f g"
proof (rule homotopic_with_linear [OF contf])
show "continuous_on (sphere 0 1 \<inter> S) g"
using pfg by (simp add: differentiable_imp_continuous_on diffg)
next
have non0fg: "0 \<notin> closed_segment (f x) (g x)" if "norm x = 1" "x \<in> S" for x
proof -
have "f x \<in> sphere 0 1"
using fim that by (simp add: image_subset_iff)
moreover have "norm(f x - g x) < 1/2"
apply (rule g12)
using that by force
ultimately show ?thesis
by (auto simp: norm_minus_commute dest: segment_bound)
qed
show "\<And>x. x \<in> sphere 0 1 \<inter> S \<Longrightarrow> closed_segment (f x) (g x) \<subseteq> T - {0}"
apply (simp add: subset_Diff_insert non0fg)
apply (simp add: segment_convex_hull)
apply (rule hull_minimal)
using fim image_eqI gim apply force
apply (rule subspace_imp_convex [OF \<open>subspace T\<close>])
done
qed
obtain d where d: "d \<in> (sphere 0 1 \<inter> T) - h ` (sphere 0 1 \<inter> S)"
using h spheremap_lemma1 [OF ST \<open>S \<subseteq> T\<close> diffh] by force
then have non0hd: "0 \<notin> closed_segment (h x) (- d)" if "norm x = 1" "x \<in> S" for x
using midpoint_between [of 0 "h x" "-d"] that h [of x]
by (auto simp: between_mem_segment midpoint_def)
have conth: "continuous_on (sphere 0 1 \<inter> S) h"
using differentiable_imp_continuous_on diffh by blast
have hom_hd: "homotopic_with_canon (\<lambda>z. True) (sphere 0 1 \<inter> S) (T - {0}) h (\<lambda>x. -d)"
apply (rule homotopic_with_linear [OF conth continuous_on_const])
apply (simp add: subset_Diff_insert non0hd)
apply (simp add: segment_convex_hull)
apply (rule hull_minimal)
using h d apply (force simp: subspace_neg [OF \<open>subspace T\<close>])
apply (rule subspace_imp_convex [OF \<open>subspace T\<close>])
done
have conT0: "continuous_on (T - {0}) (\<lambda>y. inverse(norm y) *\<^sub>R y)"
by (intro continuous_intros) auto
have sub0T: "(\<lambda>y. y /\<^sub>R norm y) ` (T - {0}) \<subseteq> sphere 0 1 \<inter> T"
by (fastforce simp: assms(2) subspace_mul)
obtain c where homhc: "homotopic_with_canon (\<lambda>z. True) (sphere 0 1 \<inter> S) (sphere 0 1 \<inter> T) h (\<lambda>x. c)"
apply (rule_tac c="-d" in that)
apply (rule homotopic_with_eq)
apply (rule homotopic_compose_continuous_left [OF hom_hd conT0 sub0T])
using d apply (auto simp: h_def)
done
show ?thesis
apply (rule_tac x=c in exI)
apply (rule homotopic_with_trans [OF _ homhc])
apply (rule homotopic_with_eq)
apply (rule homotopic_compose_continuous_left [OF homfg conT0 sub0T])
apply (auto simp: h_def)
done
qed
lemma spheremap_lemma3:
assumes "bounded S" "convex S" "subspace U" and affSU: "aff_dim S \<le> dim U"
obtains T where "subspace T" "T \<subseteq> U" "S \<noteq> {} \<Longrightarrow> aff_dim T = aff_dim S"
"(rel_frontier S) homeomorphic (sphere 0 1 \<inter> T)"
proof (cases "S = {}")
case True
with \<open>subspace U\<close> subspace_0 show ?thesis
by (rule_tac T = "{0}" in that) auto
next
case False
then obtain a where "a \<in> S"
by auto
then have affS: "aff_dim S = int (dim ((\<lambda>x. -a+x) ` S))"
by (metis hull_inc aff_dim_eq_dim)
with affSU have "dim ((\<lambda>x. -a+x) ` S) \<le> dim U"
by linarith
with choose_subspace_of_subspace
obtain T where "subspace T" "T \<subseteq> span U" and dimT: "dim T = dim ((\<lambda>x. -a+x) ` S)" .
show ?thesis
proof (rule that [OF \<open>subspace T\<close>])
show "T \<subseteq> U"
using span_eq_iff \<open>subspace U\<close> \<open>T \<subseteq> span U\<close> by blast
show "aff_dim T = aff_dim S"
using dimT \<open>subspace T\<close> affS aff_dim_subspace by fastforce
show "rel_frontier S homeomorphic sphere 0 1 \<inter> T"
proof -
have "aff_dim (ball 0 1 \<inter> T) = aff_dim (T)"
by (metis IntI interior_ball \<open>subspace T\<close> aff_dim_convex_Int_nonempty_interior centre_in_ball empty_iff inf_commute subspace_0 subspace_imp_convex zero_less_one)
then have affS_eq: "aff_dim S = aff_dim (ball 0 1 \<inter> T)"
using \<open>aff_dim T = aff_dim S\<close> by simp
have "rel_frontier S homeomorphic rel_frontier(ball 0 1 \<inter> T)"
apply (rule homeomorphic_rel_frontiers_convex_bounded_sets [OF \<open>convex S\<close> \<open>bounded S\<close>])
apply (simp add: \<open>subspace T\<close> convex_Int subspace_imp_convex)
apply (simp add: bounded_Int)
apply (rule affS_eq)
done
also have "... = frontier (ball 0 1) \<inter> T"
apply (rule convex_affine_rel_frontier_Int [OF convex_ball])
apply (simp add: \<open>subspace T\<close> subspace_imp_affine)
using \<open>subspace T\<close> subspace_0 by force
also have "... = sphere 0 1 \<inter> T"
by auto
finally show ?thesis .
qed
qed
qed
proposition inessential_spheremap_lowdim_gen:
fixes f :: "'M::euclidean_space \<Rightarrow> 'a::euclidean_space"
assumes "convex S" "bounded S" "convex T" "bounded T"
and affST: "aff_dim S < aff_dim T"
and contf: "continuous_on (rel_frontier S) f"
and fim: "f ` (rel_frontier S) \<subseteq> rel_frontier T"
obtains c where "homotopic_with_canon (\<lambda>z. True) (rel_frontier S) (rel_frontier T) f (\<lambda>x. c)"
proof (cases "S = {}")
case True
then show ?thesis
by (simp add: that)
next
case False
then show ?thesis
proof (cases "T = {}")
case True
then show ?thesis
using fim that by auto
next
case False
obtain T':: "'a set"
where "subspace T'" and affT': "aff_dim T' = aff_dim T"
and homT: "rel_frontier T homeomorphic sphere 0 1 \<inter> T'"
apply (rule spheremap_lemma3 [OF \<open>bounded T\<close> \<open>convex T\<close> subspace_UNIV, where 'b='a])
apply (simp add: aff_dim_le_DIM)
using \<open>T \<noteq> {}\<close> by blast
with homeomorphic_imp_homotopy_eqv
have relT: "sphere 0 1 \<inter> T' homotopy_eqv rel_frontier T"
using homotopy_equivalent_space_sym by blast
have "aff_dim S \<le> int (dim T')"
using affT' \<open>subspace T'\<close> affST aff_dim_subspace by force
with spheremap_lemma3 [OF \<open>bounded S\<close> \<open>convex S\<close> \<open>subspace T'\<close>] \<open>S \<noteq> {}\<close>
obtain S':: "'a set" where "subspace S'" "S' \<subseteq> T'"
and affS': "aff_dim S' = aff_dim S"
and homT: "rel_frontier S homeomorphic sphere 0 1 \<inter> S'"
by metis
with homeomorphic_imp_homotopy_eqv
have relS: "sphere 0 1 \<inter> S' homotopy_eqv rel_frontier S"
using homotopy_equivalent_space_sym by blast
have dimST': "dim S' < dim T'"
by (metis \<open>S' \<subseteq> T'\<close> \<open>subspace S'\<close> \<open>subspace T'\<close> affS' affST affT' less_irrefl not_le subspace_dim_equal)
have "\<exists>c. homotopic_with_canon (\<lambda>z. True) (rel_frontier S) (rel_frontier T) f (\<lambda>x. c)"
apply (rule homotopy_eqv_homotopic_triviality_null_imp [OF relT contf fim])
apply (rule homotopy_eqv_cohomotopic_triviality_null[OF relS, THEN iffD1, rule_format])
apply (metis dimST' \<open>subspace S'\<close> \<open>subspace T'\<close> \<open>S' \<subseteq> T'\<close> spheremap_lemma2, blast)
done
with that show ?thesis by blast
qed
qed
lemma inessential_spheremap_lowdim:
fixes f :: "'M::euclidean_space \<Rightarrow> 'a::euclidean_space"
assumes
"DIM('M) < DIM('a)" and f: "continuous_on (sphere a r) f" "f ` (sphere a r) \<subseteq> (sphere b s)"
obtains c where "homotopic_with_canon (\<lambda>z. True) (sphere a r) (sphere b s) f (\<lambda>x. c)"
proof (cases "s \<le> 0")
case True then show ?thesis
by (meson nullhomotopic_into_contractible f contractible_sphere that)
next
case False
show ?thesis
proof (cases "r \<le> 0")
case True then show ?thesis
by (meson f nullhomotopic_from_contractible contractible_sphere that)
next
case False
with \<open>\<not> s \<le> 0\<close> have "r > 0" "s > 0" by auto
show ?thesis
apply (rule inessential_spheremap_lowdim_gen [of "cball a r" "cball b s" f])
using \<open>0 < r\<close> \<open>0 < s\<close> assms(1)
apply (simp_all add: f aff_dim_cball)
using that by blast
qed
qed
subsection\<open> Some technical lemmas about extending maps from cell complexes\<close>
lemma extending_maps_Union_aux:
assumes fin: "finite \<F>"
and "\<And>S. S \<in> \<F> \<Longrightarrow> closed S"
and "\<And>S T. \<lbrakk>S \<in> \<F>; T \<in> \<F>; S \<noteq> T\<rbrakk> \<Longrightarrow> S \<inter> T \<subseteq> K"
and "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>g. continuous_on S g \<and> g ` S \<subseteq> T \<and> (\<forall>x \<in> S \<inter> K. g x = h x)"
shows "\<exists>g. continuous_on (\<Union>\<F>) g \<and> g ` (\<Union>\<F>) \<subseteq> T \<and> (\<forall>x \<in> \<Union>\<F> \<inter> K. g x = h x)"
using assms
proof (induction \<F>)
case empty show ?case by simp
next
case (insert S \<F>)
then obtain f where contf: "continuous_on (S) f" and fim: "f ` S \<subseteq> T" and feq: "\<forall>x \<in> S \<inter> K. f x = h x"
by (meson insertI1)
obtain g where contg: "continuous_on (\<Union>\<F>) g" and gim: "g ` \<Union>\<F> \<subseteq> T" and geq: "\<forall>x \<in> \<Union>\<F> \<inter> K. g x = h x"
using insert by auto
have fg: "f x = g x" if "x \<in> T" "T \<in> \<F>" "x \<in> S" for x T
proof -
have "T \<inter> S \<subseteq> K \<or> S = T"
using that by (metis (no_types) insert.prems(2) insertCI)
then show ?thesis
using UnionI feq geq \<open>S \<notin> \<F>\<close> subsetD that by fastforce
qed
show ?case
apply (rule_tac x="\<lambda>x. if x \<in> S then f x else g x" in exI, simp)
apply (intro conjI continuous_on_cases)
apply (simp_all add: insert closed_Union contf contg)
using fim gim feq geq
apply (force simp: insert closed_Union contf contg inf_commute intro: fg)+
done
qed
lemma extending_maps_Union:
assumes fin: "finite \<F>"
and "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>g. continuous_on S g \<and> g ` S \<subseteq> T \<and> (\<forall>x \<in> S \<inter> K. g x = h x)"
and "\<And>S. S \<in> \<F> \<Longrightarrow> closed S"
and K: "\<And>X Y. \<lbrakk>X \<in> \<F>; Y \<in> \<F>; \<not> X \<subseteq> Y; \<not> Y \<subseteq> X\<rbrakk> \<Longrightarrow> X \<inter> Y \<subseteq> K"
shows "\<exists>g. continuous_on (\<Union>\<F>) g \<and> g ` (\<Union>\<F>) \<subseteq> T \<and> (\<forall>x \<in> \<Union>\<F> \<inter> K. g x = h x)"
apply (simp add: Union_maximal_sets [OF fin, symmetric])
apply (rule extending_maps_Union_aux)
apply (simp_all add: Union_maximal_sets [OF fin] assms)
by (metis K psubsetI)
lemma extend_map_lemma:
assumes "finite \<F>" "\<G> \<subseteq> \<F>" "convex T" "bounded T"
and poly: "\<And>X. X \<in> \<F> \<Longrightarrow> polytope X"
and aff: "\<And>X. X \<in> \<F> - \<G> \<Longrightarrow> aff_dim X < aff_dim T"
and face: "\<And>S T. \<lbrakk>S \<in> \<F>; T \<in> \<F>\<rbrakk> \<Longrightarrow> (S \<inter> T) face_of S \<and> (S \<inter> T) face_of T"
and contf: "continuous_on (\<Union>\<G>) f" and fim: "f ` (\<Union>\<G>) \<subseteq> rel_frontier T"
obtains g where "continuous_on (\<Union>\<F>) g" "g ` (\<Union>\<F>) \<subseteq> rel_frontier T" "\<And>x. x \<in> \<Union>\<G> \<Longrightarrow> g x = f x"
proof (cases "\<F> - \<G> = {}")
case True
then have "\<Union>\<F> \<subseteq> \<Union>\<G>"
by (simp add: Union_mono)
then show ?thesis
apply (rule_tac g=f in that)
using contf continuous_on_subset apply blast
using fim apply blast
by simp
next
case False
then have "0 \<le> aff_dim T"
by (metis aff aff_dim_empty aff_dim_geq aff_dim_negative_iff all_not_in_conv not_less)
then obtain i::nat where i: "int i = aff_dim T"
by (metis nonneg_eq_int)
have Union_empty_eq: "\<Union>{D. D = {} \<and> P D} = {}" for P :: "'a set \<Rightarrow> bool"
by auto
have extendf: "\<exists>g. continuous_on (\<Union>(\<G> \<union> {D. \<exists>C \<in> \<F>. D face_of C \<and> aff_dim D < i})) g \<and>
g ` (\<Union> (\<G> \<union> {D. \<exists>C \<in> \<F>. D face_of C \<and> aff_dim D < i})) \<subseteq> rel_frontier T \<and>
(\<forall>x \<in> \<Union>\<G>. g x = f x)"
if "i \<le> aff_dim T" for i::nat
using that
proof (induction i)
case 0 then show ?case
apply (simp add: Union_empty_eq)
apply (rule_tac x=f in exI)
apply (intro conjI)
using contf continuous_on_subset apply blast
using fim apply blast
by simp
next
case (Suc p)
with \<open>bounded T\<close> have "rel_frontier T \<noteq> {}"
by (auto simp: rel_frontier_eq_empty affine_bounded_eq_lowdim [of T])
then obtain t where t: "t \<in> rel_frontier T" by auto
have ple: "int p \<le> aff_dim T" using Suc.prems by force
obtain h where conth: "continuous_on (\<Union>(\<G> \<union> {D. \<exists>C \<in> \<F>. D face_of C \<and> aff_dim D < p})) h"
and him: "h ` (\<Union> (\<G> \<union> {D. \<exists>C \<in> \<F>. D face_of C \<and> aff_dim D < p}))
\<subseteq> rel_frontier T"
and heq: "\<And>x. x \<in> \<Union>\<G> \<Longrightarrow> h x = f x"
using Suc.IH [OF ple] by auto
let ?Faces = "{D. \<exists>C \<in> \<F>. D face_of C \<and> aff_dim D \<le> p}"
have extendh: "\<exists>g. continuous_on D g \<and>
g ` D \<subseteq> rel_frontier T \<and>
(\<forall>x \<in> D \<inter> \<Union>(\<G> \<union> {D. \<exists>C \<in> \<F>. D face_of C \<and> aff_dim D < p}). g x = h x)"
if D: "D \<in> \<G> \<union> ?Faces" for D
proof (cases "D \<subseteq> \<Union>(\<G> \<union> {D. \<exists>C \<in> \<F>. D face_of C \<and> aff_dim D < p})")
case True
then show ?thesis
apply (rule_tac x=h in exI)
apply (intro conjI)
apply (blast intro: continuous_on_subset [OF conth])
using him apply blast
by simp
next
case False
note notDsub = False
show ?thesis
proof (cases "\<exists>a. D = {a}")
case True
then obtain a where "D = {a}" by auto
with notDsub t show ?thesis
by (rule_tac x="\<lambda>x. t" in exI) simp
next
case False
have "D \<noteq> {}" using notDsub by auto
have Dnotin: "D \<notin> \<G> \<union> {D. \<exists>C \<in> \<F>. D face_of C \<and> aff_dim D < p}"
using notDsub by auto
then have "D \<notin> \<G>" by simp
have "D \<in> ?Faces - {D. \<exists>C \<in> \<F>. D face_of C \<and> aff_dim D < p}"
using Dnotin that by auto
then obtain C where "C \<in> \<F>" "D face_of C" and affD: "aff_dim D = int p"
by auto
then have "bounded D"
using face_of_polytope_polytope poly polytope_imp_bounded by blast
then have [simp]: "\<not> affine D"
using affine_bounded_eq_trivial False \<open>D \<noteq> {}\<close> \<open>bounded D\<close> by blast
have "{F. F facet_of D} \<subseteq> {E. E face_of C \<and> aff_dim E < int p}"
apply clarify
apply (metis \<open>D face_of C\<close> affD eq_iff face_of_trans facet_of_def zle_diff1_eq)
done
moreover have "polyhedron D"
using \<open>C \<in> \<F>\<close> \<open>D face_of C\<close> face_of_polytope_polytope poly polytope_imp_polyhedron by auto
ultimately have relf_sub: "rel_frontier D \<subseteq> \<Union> {E. E face_of C \<and> aff_dim E < p}"
by (simp add: rel_frontier_of_polyhedron Union_mono)
then have him_relf: "h ` rel_frontier D \<subseteq> rel_frontier T"
using \<open>C \<in> \<F>\<close> him by blast
have "convex D"
by (simp add: \<open>polyhedron D\<close> polyhedron_imp_convex)
have affD_lessT: "aff_dim D < aff_dim T"
using Suc.prems affD by linarith
have contDh: "continuous_on (rel_frontier D) h"
using \<open>C \<in> \<F>\<close> relf_sub by (blast intro: continuous_on_subset [OF conth])
then have *: "(\<exists>c. homotopic_with_canon (\<lambda>x. True) (rel_frontier D) (rel_frontier T) h (\<lambda>x. c)) =
(\<exists>g. continuous_on UNIV g \<and> range g \<subseteq> rel_frontier T \<and>
(\<forall>x\<in>rel_frontier D. g x = h x))"
apply (rule nullhomotopic_into_rel_frontier_extension [OF closed_rel_frontier])
apply (simp_all add: assms rel_frontier_eq_empty him_relf)
done
have "(\<exists>c. homotopic_with_canon (\<lambda>x. True) (rel_frontier D)
(rel_frontier T) h (\<lambda>x. c))"
by (metis inessential_spheremap_lowdim_gen
[OF \<open>convex D\<close> \<open>bounded D\<close> \<open>convex T\<close> \<open>bounded T\<close> affD_lessT contDh him_relf])
then obtain g where contg: "continuous_on UNIV g"
and gim: "range g \<subseteq> rel_frontier T"
and gh: "\<And>x. x \<in> rel_frontier D \<Longrightarrow> g x = h x"
by (metis *)
have "D \<inter> E \<subseteq> rel_frontier D"
if "E \<in> \<G> \<union> {D. Bex \<F> ((face_of) D) \<and> aff_dim D < int p}" for E
proof (rule face_of_subset_rel_frontier)
show "D \<inter> E face_of D"
using that \<open>C \<in> \<F>\<close> \<open>D face_of C\<close> face
apply auto
apply (meson face_of_Int_subface \<open>\<G> \<subseteq> \<F>\<close> face_of_refl_eq poly polytope_imp_convex subsetD)
using face_of_Int_subface apply blast
done
show "D \<inter> E \<noteq> D"
using that notDsub by auto
qed
then show ?thesis
apply (rule_tac x=g in exI)
apply (intro conjI ballI)
using continuous_on_subset contg apply blast
using gim apply blast
using gh by fastforce
qed
qed
have intle: "i < 1 + int j \<longleftrightarrow> i \<le> int j" for i j
by auto
have "finite \<G>"
using \<open>finite \<F>\<close> \<open>\<G> \<subseteq> \<F>\<close> rev_finite_subset by blast
then have fin: "finite (\<G> \<union> ?Faces)"
apply simp
apply (rule_tac B = "\<Union>{{D. D face_of C}| C. C \<in> \<F>}" in finite_subset)
by (auto simp: \<open>finite \<F>\<close> finite_polytope_faces poly)
have clo: "closed S" if "S \<in> \<G> \<union> ?Faces" for S
using that \<open>\<G> \<subseteq> \<F>\<close> face_of_polytope_polytope poly polytope_imp_closed by blast
have K: "X \<inter> Y \<subseteq> \<Union>(\<G> \<union> {D. \<exists>C\<in>\<F>. D face_of C \<and> aff_dim D < int p})"
if "X \<in> \<G> \<union> ?Faces" "Y \<in> \<G> \<union> ?Faces" "\<not> Y \<subseteq> X" for X Y
proof -
have ff: "X \<inter> Y face_of X \<and> X \<inter> Y face_of Y"
if XY: "X face_of D" "Y face_of E" and DE: "D \<in> \<F>" "E \<in> \<F>" for D E
apply (rule face_of_Int_subface [OF _ _ XY])
apply (auto simp: face DE)
done
show ?thesis
using that
apply auto
apply (drule_tac x="X \<inter> Y" in spec, safe)
using ff face_of_imp_convex [of X] face_of_imp_convex [of Y]
apply (fastforce dest: face_of_aff_dim_lt)
by (meson face_of_trans ff)
qed
obtain g where "continuous_on (\<Union>(\<G> \<union> ?Faces)) g"
"g ` \<Union>(\<G> \<union> ?Faces) \<subseteq> rel_frontier T"
"(\<forall>x \<in> \<Union>(\<G> \<union> ?Faces) \<inter>
\<Union>(\<G> \<union> {D. \<exists>C\<in>\<F>. D face_of C \<and> aff_dim D < p}). g x = h x)"
apply (rule exE [OF extending_maps_Union [OF fin extendh clo K]], blast+)
done
then show ?case
apply (simp add: intle local.heq [symmetric], blast)
done
qed
have eq: "\<Union>(\<G> \<union> {D. \<exists>C \<in> \<F>. D face_of C \<and> aff_dim D < i}) = \<Union>\<F>"
proof
show "\<Union>(\<G> \<union> {D. \<exists>C\<in>\<F>. D face_of C \<and> aff_dim D < int i}) \<subseteq> \<Union>\<F>"
apply (rule Union_subsetI)
using \<open>\<G> \<subseteq> \<F>\<close> face_of_imp_subset apply force
done
show "\<Union>\<F> \<subseteq> \<Union>(\<G> \<union> {D. \<exists>C\<in>\<F>. D face_of C \<and> aff_dim D < i})"
apply (rule Union_mono)
using face apply (fastforce simp: aff i)
done
qed
have "int i \<le> aff_dim T" by (simp add: i)
then show ?thesis
using extendf [of i] unfolding eq by (metis that)
qed
lemma extend_map_lemma_cofinite0:
assumes "finite \<F>"
and "pairwise (\<lambda>S T. S \<inter> T \<subseteq> K) \<F>"
and "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>a g. a \<notin> U \<and> continuous_on (S - {a}) g \<and> g ` (S - {a}) \<subseteq> T \<and> (\<forall>x \<in> S \<inter> K. g x = h x)"
and "\<And>S. S \<in> \<F> \<Longrightarrow> closed S"
shows "\<exists>C g. finite C \<and> disjnt C U \<and> card C \<le> card \<F> \<and>
continuous_on (\<Union>\<F> - C) g \<and> g ` (\<Union>\<F> - C) \<subseteq> T
\<and> (\<forall>x \<in> (\<Union>\<F> - C) \<inter> K. g x = h x)"
using assms
proof induction
case empty then show ?case
by force
next
case (insert X \<F>)
then have "closed X" and clo: "\<And>X. X \<in> \<F> \<Longrightarrow> closed X"
and \<F>: "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>a g. a \<notin> U \<and> continuous_on (S - {a}) g \<and> g ` (S - {a}) \<subseteq> T \<and> (\<forall>x \<in> S \<inter> K. g x = h x)"
and pwX: "\<And>Y. Y \<in> \<F> \<and> Y \<noteq> X \<longrightarrow> X \<inter> Y \<subseteq> K \<and> Y \<inter> X \<subseteq> K"
and pwF: "pairwise (\<lambda> S T. S \<inter> T \<subseteq> K) \<F>"
by (simp_all add: pairwise_insert)
obtain C g where C: "finite C" "disjnt C U" "card C \<le> card \<F>"
and contg: "continuous_on (\<Union>\<F> - C) g"
and gim: "g ` (\<Union>\<F> - C) \<subseteq> T"
and gh: "\<And>x. x \<in> (\<Union>\<F> - C) \<inter> K \<Longrightarrow> g x = h x"
using insert.IH [OF pwF \<F> clo] by auto
obtain a f where "a \<notin> U"
and contf: "continuous_on (X - {a}) f"
and fim: "f ` (X - {a}) \<subseteq> T"
and fh: "(\<forall>x \<in> X \<inter> K. f x = h x)"
using insert.prems by (meson insertI1)
show ?case
proof (intro exI conjI)
show "finite (insert a C)"
by (simp add: C)
show "disjnt (insert a C) U"
using C \<open>a \<notin> U\<close> by simp
show "card (insert a C) \<le> card (insert X \<F>)"
by (simp add: C card_insert_if insert.hyps le_SucI)
have "closed (\<Union>\<F>)"
using clo insert.hyps by blast
have "continuous_on (X - insert a C \<union> (\<Union>\<F> - insert a C)) (\<lambda>x. if x \<in> X then f x else g x)"
apply (rule continuous_on_cases_local)
apply (simp_all add: closedin_closed)
using \<open>closed X\<close> apply blast
using \<open>closed (\<Union>\<F>)\<close> apply blast
using contf apply (force simp: elim: continuous_on_subset)
using contg apply (force simp: elim: continuous_on_subset)
using fh gh insert.hyps pwX by fastforce
then show "continuous_on (\<Union>(insert X \<F>) - insert a C) (\<lambda>a. if a \<in> X then f a else g a)"
by (blast intro: continuous_on_subset)
show "\<forall>x\<in>(\<Union>(insert X \<F>) - insert a C) \<inter> K. (if x \<in> X then f x else g x) = h x"
using gh by (auto simp: fh)
show "(\<lambda>a. if a \<in> X then f a else g a) ` (\<Union>(insert X \<F>) - insert a C) \<subseteq> T"
using fim gim by auto force
qed
qed
lemma extend_map_lemma_cofinite1:
assumes "finite \<F>"
and \<F>: "\<And>X. X \<in> \<F> \<Longrightarrow> \<exists>a g. a \<notin> U \<and> continuous_on (X - {a}) g \<and> g ` (X - {a}) \<subseteq> T \<and> (\<forall>x \<in> X \<inter> K. g x = h x)"
and clo: "\<And>X. X \<in> \<F> \<Longrightarrow> closed X"
and K: "\<And>X Y. \<lbrakk>X \<in> \<F>; Y \<in> \<F>; \<not> X \<subseteq> Y; \<not> Y \<subseteq> X\<rbrakk> \<Longrightarrow> X \<inter> Y \<subseteq> K"
obtains C g where "finite C" "disjnt C U" "card C \<le> card \<F>" "continuous_on (\<Union>\<F> - C) g"
"g ` (\<Union>\<F> - C) \<subseteq> T"
"\<And>x. x \<in> (\<Union>\<F> - C) \<inter> K \<Longrightarrow> g x = h x"
proof -
let ?\<F> = "{X \<in> \<F>. \<forall>Y\<in>\<F>. \<not> X \<subset> Y}"
have [simp]: "\<Union>?\<F> = \<Union>\<F>"
by (simp add: Union_maximal_sets assms)
have fin: "finite ?\<F>"
by (force intro: finite_subset [OF _ \<open>finite \<F>\<close>])
have pw: "pairwise (\<lambda> S T. S \<inter> T \<subseteq> K) ?\<F>"
by (simp add: pairwise_def) (metis K psubsetI)
have "card {X \<in> \<F>. \<forall>Y\<in>\<F>. \<not> X \<subset> Y} \<le> card \<F>"
by (simp add: \<open>finite \<F>\<close> card_mono)
moreover
obtain C g where "finite C \<and> disjnt C U \<and> card C \<le> card ?\<F> \<and>
continuous_on (\<Union>?\<F> - C) g \<and> g ` (\<Union>?\<F> - C) \<subseteq> T
\<and> (\<forall>x \<in> (\<Union>?\<F> - C) \<inter> K. g x = h x)"
apply (rule exE [OF extend_map_lemma_cofinite0 [OF fin pw, of U T h]])
apply (fastforce intro!: clo \<F>)+
done
ultimately show ?thesis
by (rule_tac C=C and g=g in that) auto
qed
lemma extend_map_lemma_cofinite:
assumes "finite \<F>" "\<G> \<subseteq> \<F>" and T: "convex T" "bounded T"
and poly: "\<And>X. X \<in> \<F> \<Longrightarrow> polytope X"
and contf: "continuous_on (\<Union>\<G>) f" and fim: "f ` (\<Union>\<G>) \<subseteq> rel_frontier T"
and face: "\<And>X Y. \<lbrakk>X \<in> \<F>; Y \<in> \<F>\<rbrakk> \<Longrightarrow> (X \<inter> Y) face_of X \<and> (X \<inter> Y) face_of Y"
and aff: "\<And>X. X \<in> \<F> - \<G> \<Longrightarrow> aff_dim X \<le> aff_dim T"
obtains C g where
"finite C" "disjnt C (\<Union>\<G>)" "card C \<le> card \<F>" "continuous_on (\<Union>\<F> - C) g"
"g ` (\<Union> \<F> - C) \<subseteq> rel_frontier T" "\<And>x. x \<in> \<Union>\<G> \<Longrightarrow> g x = f x"
proof -
define \<H> where "\<H> \<equiv> \<G> \<union> {D. \<exists>C \<in> \<F> - \<G>. D face_of C \<and> aff_dim D < aff_dim T}"
have "finite \<G>"
using assms finite_subset by blast
moreover have "finite (\<Union>{{D. D face_of C} |C. C \<in> \<F>})"
apply (rule finite_Union)
apply (simp add: \<open>finite \<F>\<close>)
using finite_polytope_faces poly by auto
ultimately have "finite \<H>"
apply (simp add: \<H>_def)
apply (rule finite_subset [of _ "\<Union> {{D. D face_of C} | C. C \<in> \<F>}"], auto)
done
have *: "\<And>X Y. \<lbrakk>X \<in> \<H>; Y \<in> \<H>\<rbrakk> \<Longrightarrow> X \<inter> Y face_of X \<and> X \<inter> Y face_of Y"
unfolding \<H>_def
apply (elim UnE bexE CollectE DiffE)
using subsetD [OF \<open>\<G> \<subseteq> \<F>\<close>] apply (simp_all add: face)
apply (meson subsetD [OF \<open>\<G> \<subseteq> \<F>\<close>] face face_of_Int_subface face_of_imp_subset face_of_refl poly polytope_imp_convex)+
done
obtain h where conth: "continuous_on (\<Union>\<H>) h" and him: "h ` (\<Union>\<H>) \<subseteq> rel_frontier T"
and hf: "\<And>x. x \<in> \<Union>\<G> \<Longrightarrow> h x = f x"
using \<open>finite \<H>\<close>
unfolding \<H>_def
apply (rule extend_map_lemma [OF _ Un_upper1 T _ _ _ contf fim])
using \<open>\<G> \<subseteq> \<F>\<close> face_of_polytope_polytope poly apply fastforce
using * apply (auto simp: \<H>_def)
done
have "bounded (\<Union>\<G>)"
using \<open>finite \<G>\<close> \<open>\<G> \<subseteq> \<F>\<close> poly polytope_imp_bounded by blast
then have "\<Union>\<G> \<noteq> UNIV"
by auto
then obtain a where a: "a \<notin> \<Union>\<G>"
by blast
have \<F>: "\<exists>a g. a \<notin> \<Union>\<G> \<and> continuous_on (D - {a}) g \<and>
g ` (D - {a}) \<subseteq> rel_frontier T \<and> (\<forall>x \<in> D \<inter> \<Union>\<H>. g x = h x)"
if "D \<in> \<F>" for D
proof (cases "D \<subseteq> \<Union>\<H>")
case True
then show ?thesis
apply (rule_tac x=a in exI)
apply (rule_tac x=h in exI)
using him apply (blast intro!: \<open>a \<notin> \<Union>\<G>\<close> continuous_on_subset [OF conth]) +
done
next
case False
note D_not_subset = False
show ?thesis
proof (cases "D \<in> \<G>")
case True
with D_not_subset show ?thesis
by (auto simp: \<H>_def)
next
case False
then have affD: "aff_dim D \<le> aff_dim T"
by (simp add: \<open>D \<in> \<F>\<close> aff)
show ?thesis
proof (cases "rel_interior D = {}")
case True
with \<open>D \<in> \<F>\<close> poly a show ?thesis
by (force simp: rel_interior_eq_empty polytope_imp_convex)
next
case False
then obtain b where brelD: "b \<in> rel_interior D"
by blast
have "polyhedron D"
by (simp add: poly polytope_imp_polyhedron that)
have "rel_frontier D retract_of affine hull D - {b}"
by (simp add: rel_frontier_retract_of_punctured_affine_hull poly polytope_imp_bounded polytope_imp_convex that brelD)
then obtain r where relfD: "rel_frontier D \<subseteq> affine hull D - {b}"
and contr: "continuous_on (affine hull D - {b}) r"
and rim: "r ` (affine hull D - {b}) \<subseteq> rel_frontier D"
and rid: "\<And>x. x \<in> rel_frontier D \<Longrightarrow> r x = x"
by (auto simp: retract_of_def retraction_def)
show ?thesis
proof (intro exI conjI ballI)
show "b \<notin> \<Union>\<G>"
proof clarify
fix E
assume "b \<in> E" "E \<in> \<G>"
then have "E \<inter> D face_of E \<and> E \<inter> D face_of D"
using \<open>\<G> \<subseteq> \<F>\<close> face that by auto
with face_of_subset_rel_frontier \<open>E \<in> \<G>\<close> \<open>b \<in> E\<close> brelD rel_interior_subset [of D]
D_not_subset rel_frontier_def \<H>_def
show False
by blast
qed
have "r ` (D - {b}) \<subseteq> r ` (affine hull D - {b})"
by (simp add: Diff_mono hull_subset image_mono)
also have "... \<subseteq> rel_frontier D"
by (rule rim)
also have "... \<subseteq> \<Union>{E. E face_of D \<and> aff_dim E < aff_dim T}"
using affD
by (force simp: rel_frontier_of_polyhedron [OF \<open>polyhedron D\<close>] facet_of_def)
also have "... \<subseteq> \<Union>(\<H>)"
using D_not_subset \<H>_def that by fastforce
finally have rsub: "r ` (D - {b}) \<subseteq> \<Union>(\<H>)" .
show "continuous_on (D - {b}) (h \<circ> r)"
apply (intro conjI \<open>b \<notin> \<Union>\<G>\<close> continuous_on_compose)
apply (rule continuous_on_subset [OF contr])
apply (simp add: Diff_mono hull_subset)
apply (rule continuous_on_subset [OF conth rsub])
done
show "(h \<circ> r) ` (D - {b}) \<subseteq> rel_frontier T"
using brelD him rsub by fastforce
show "(h \<circ> r) x = h x" if x: "x \<in> D \<inter> \<Union>\<H>" for x
proof -
consider A where "x \<in> D" "A \<in> \<G>" "x \<in> A"
| A B where "x \<in> D" "A face_of B" "B \<in> \<F>" "B \<notin> \<G>" "aff_dim A < aff_dim T" "x \<in> A"
using x by (auto simp: \<H>_def)
then have xrel: "x \<in> rel_frontier D"
proof cases
case 1 show ?thesis
proof (rule face_of_subset_rel_frontier [THEN subsetD])
show "D \<inter> A face_of D"
using \<open>A \<in> \<G>\<close> \<open>\<G> \<subseteq> \<F>\<close> face \<open>D \<in> \<F>\<close> by blast
show "D \<inter> A \<noteq> D"
using \<open>A \<in> \<G>\<close> D_not_subset \<H>_def by blast
qed (auto simp: 1)
next
case 2 show ?thesis
proof (rule face_of_subset_rel_frontier [THEN subsetD])
show "D \<inter> A face_of D"
apply (rule face_of_Int_subface [of D B _ A, THEN conjunct1])
apply (simp_all add: 2 \<open>D \<in> \<F>\<close> face)
apply (simp add: \<open>polyhedron D\<close> polyhedron_imp_convex face_of_refl)
done
show "D \<inter> A \<noteq> D"
using "2" D_not_subset \<H>_def by blast
qed (auto simp: 2)
qed
show ?thesis
by (simp add: rid xrel)
qed
qed
qed
qed
qed
have clo: "\<And>S. S \<in> \<F> \<Longrightarrow> closed S"
by (simp add: poly polytope_imp_closed)
obtain C g where "finite C" "disjnt C (\<Union>\<G>)" "card C \<le> card \<F>" "continuous_on (\<Union>\<F> - C) g"
"g ` (\<Union>\<F> - C) \<subseteq> rel_frontier T"
and gh: "\<And>x. x \<in> (\<Union>\<F> - C) \<inter> \<Union>\<H> \<Longrightarrow> g x = h x"
proof (rule extend_map_lemma_cofinite1 [OF \<open>finite \<F>\<close> \<F> clo])
show "X \<inter> Y \<subseteq> \<Union>\<H>" if XY: "X \<in> \<F>" "Y \<in> \<F>" and "\<not> X \<subseteq> Y" "\<not> Y \<subseteq> X" for X Y
proof (cases "X \<in> \<G>")
case True
then show ?thesis
by (auto simp: \<H>_def)
next
case False
have "X \<inter> Y \<noteq> X"
using \<open>\<not> X \<subseteq> Y\<close> by blast
with XY
show ?thesis
by (clarsimp simp: \<H>_def)
(metis Diff_iff Int_iff aff antisym_conv face face_of_aff_dim_lt face_of_refl
not_le poly polytope_imp_convex)
qed
qed (blast)+
with \<open>\<G> \<subseteq> \<F>\<close> show ?thesis
apply (rule_tac C=C and g=g in that)
apply (auto simp: disjnt_def hf [symmetric] \<H>_def intro!: gh)
done
qed
text\<open>The next two proofs are similar\<close>
theorem extend_map_cell_complex_to_sphere:
assumes "finite \<F>" and S: "S \<subseteq> \<Union>\<F>" "closed S" and T: "convex T" "bounded T"
and poly: "\<And>X. X \<in> \<F> \<Longrightarrow> polytope X"
and aff: "\<And>X. X \<in> \<F> \<Longrightarrow> aff_dim X < aff_dim T"
and face: "\<And>X Y. \<lbrakk>X \<in> \<F>; Y \<in> \<F>\<rbrakk> \<Longrightarrow> (X \<inter> Y) face_of X \<and> (X \<inter> Y) face_of Y"
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> rel_frontier T"
obtains g where "continuous_on (\<Union>\<F>) g"
"g ` (\<Union>\<F>) \<subseteq> rel_frontier T" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
proof -
obtain V g where "S \<subseteq> V" "open V" "continuous_on V g" and gim: "g ` V \<subseteq> rel_frontier T" and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
using neighbourhood_extension_into_ANR [OF contf fim _ \<open>closed S\<close>] ANR_rel_frontier_convex T by blast
have "compact S"
by (meson assms compact_Union poly polytope_imp_compact seq_compact_closed_subset seq_compact_eq_compact)
then obtain d where "d > 0" and d: "\<And>x y. \<lbrakk>x \<in> S; y \<in> - V\<rbrakk> \<Longrightarrow> d \<le> dist x y"
using separate_compact_closed [of S "-V"] \<open>open V\<close> \<open>S \<subseteq> V\<close> by force
obtain \<G> where "finite \<G>" "\<Union>\<G> = \<Union>\<F>"
and diaG: "\<And>X. X \<in> \<G> \<Longrightarrow> diameter X < d"
and polyG: "\<And>X. X \<in> \<G> \<Longrightarrow> polytope X"
and affG: "\<And>X. X \<in> \<G> \<Longrightarrow> aff_dim X \<le> aff_dim T - 1"
and faceG: "\<And>X Y. \<lbrakk>X \<in> \<G>; Y \<in> \<G>\<rbrakk> \<Longrightarrow> X \<inter> Y face_of X \<and> X \<inter> Y face_of Y"
proof (rule cell_complex_subdivision_exists [OF \<open>d>0\<close> \<open>finite \<F>\<close> poly _ face])
show "\<And>X. X \<in> \<F> \<Longrightarrow> aff_dim X \<le> aff_dim T - 1"
by (simp add: aff)
qed auto
obtain h where conth: "continuous_on (\<Union>\<G>) h" and him: "h ` \<Union>\<G> \<subseteq> rel_frontier T" and hg: "\<And>x. x \<in> \<Union>(\<G> \<inter> Pow V) \<Longrightarrow> h x = g x"
proof (rule extend_map_lemma [of \<G> "\<G> \<inter> Pow V" T g])
show "continuous_on (\<Union>(\<G> \<inter> Pow V)) g"
by (metis Union_Int_subset Union_Pow_eq \<open>continuous_on V g\<close> continuous_on_subset le_inf_iff)
qed (use \<open>finite \<G>\<close> T polyG affG faceG gim in fastforce)+
show ?thesis
proof
show "continuous_on (\<Union>\<F>) h"
using \<open>\<Union>\<G> = \<Union>\<F>\<close> conth by auto
show "h ` \<Union>\<F> \<subseteq> rel_frontier T"
using \<open>\<Union>\<G> = \<Union>\<F>\<close> him by auto
show "h x = f x" if "x \<in> S" for x
proof -
have "x \<in> \<Union>\<G>"
using \<open>\<Union>\<G> = \<Union>\<F>\<close> \<open>S \<subseteq> \<Union>\<F>\<close> that by auto
then obtain X where "x \<in> X" "X \<in> \<G>" by blast
then have "diameter X < d" "bounded X"
by (auto simp: diaG \<open>X \<in> \<G>\<close> polyG polytope_imp_bounded)
then have "X \<subseteq> V" using d [OF \<open>x \<in> S\<close>] diameter_bounded_bound [OF \<open>bounded X\<close> \<open>x \<in> X\<close>]
by fastforce
have "h x = g x"
apply (rule hg)
using \<open>X \<in> \<G>\<close> \<open>X \<subseteq> V\<close> \<open>x \<in> X\<close> by blast
also have "... = f x"
by (simp add: gf that)
finally show "h x = f x" .
qed
qed
qed
theorem extend_map_cell_complex_to_sphere_cofinite:
assumes "finite \<F>" and S: "S \<subseteq> \<Union>\<F>" "closed S" and T: "convex T" "bounded T"
and poly: "\<And>X. X \<in> \<F> \<Longrightarrow> polytope X"
and aff: "\<And>X. X \<in> \<F> \<Longrightarrow> aff_dim X \<le> aff_dim T"
and face: "\<And>X Y. \<lbrakk>X \<in> \<F>; Y \<in> \<F>\<rbrakk> \<Longrightarrow> (X \<inter> Y) face_of X \<and> (X \<inter> Y) face_of Y"
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> rel_frontier T"
obtains C g where "finite C" "disjnt C S" "continuous_on (\<Union>\<F> - C) g"
"g ` (\<Union>\<F> - C) \<subseteq> rel_frontier T" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
proof -
obtain V g where "S \<subseteq> V" "open V" "continuous_on V g" and gim: "g ` V \<subseteq> rel_frontier T" and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
using neighbourhood_extension_into_ANR [OF contf fim _ \<open>closed S\<close>] ANR_rel_frontier_convex T by blast
have "compact S"
by (meson assms compact_Union poly polytope_imp_compact seq_compact_closed_subset seq_compact_eq_compact)
then obtain d where "d > 0" and d: "\<And>x y. \<lbrakk>x \<in> S; y \<in> - V\<rbrakk> \<Longrightarrow> d \<le> dist x y"
using separate_compact_closed [of S "-V"] \<open>open V\<close> \<open>S \<subseteq> V\<close> by force
obtain \<G> where "finite \<G>" "\<Union>\<G> = \<Union>\<F>"
and diaG: "\<And>X. X \<in> \<G> \<Longrightarrow> diameter X < d"
and polyG: "\<And>X. X \<in> \<G> \<Longrightarrow> polytope X"
and affG: "\<And>X. X \<in> \<G> \<Longrightarrow> aff_dim X \<le> aff_dim T"
and faceG: "\<And>X Y. \<lbrakk>X \<in> \<G>; Y \<in> \<G>\<rbrakk> \<Longrightarrow> X \<inter> Y face_of X \<and> X \<inter> Y face_of Y"
by (rule cell_complex_subdivision_exists [OF \<open>d>0\<close> \<open>finite \<F>\<close> poly aff face]) auto
obtain C h where "finite C" and dis: "disjnt C (\<Union>(\<G> \<inter> Pow V))"
and card: "card C \<le> card \<G>" and conth: "continuous_on (\<Union>\<G> - C) h"
and him: "h ` (\<Union>\<G> - C) \<subseteq> rel_frontier T"
and hg: "\<And>x. x \<in> \<Union>(\<G> \<inter> Pow V) \<Longrightarrow> h x = g x"
proof (rule extend_map_lemma_cofinite [of \<G> "\<G> \<inter> Pow V" T g])
show "continuous_on (\<Union>(\<G> \<inter> Pow V)) g"
by (metis Union_Int_subset Union_Pow_eq \<open>continuous_on V g\<close> continuous_on_subset le_inf_iff)
show "g ` \<Union>(\<G> \<inter> Pow V) \<subseteq> rel_frontier T"
using gim by force
qed (auto intro: \<open>finite \<G>\<close> T polyG affG dest: faceG)
have Ssub: "S \<subseteq> \<Union>(\<G> \<inter> Pow V)"
proof
fix x
assume "x \<in> S"
then have "x \<in> \<Union>\<G>"
using \<open>\<Union>\<G> = \<Union>\<F>\<close> \<open>S \<subseteq> \<Union>\<F>\<close> by auto
then obtain X where "x \<in> X" "X \<in> \<G>" by blast
then have "diameter X < d" "bounded X"
by (auto simp: diaG \<open>X \<in> \<G>\<close> polyG polytope_imp_bounded)
then have "X \<subseteq> V" using d [OF \<open>x \<in> S\<close>] diameter_bounded_bound [OF \<open>bounded X\<close> \<open>x \<in> X\<close>]
by fastforce
then show "x \<in> \<Union>(\<G> \<inter> Pow V)"
using \<open>X \<in> \<G>\<close> \<open>x \<in> X\<close> by blast
qed
show ?thesis
proof
show "continuous_on (\<Union>\<F>-C) h"
using \<open>\<Union>\<G> = \<Union>\<F>\<close> conth by auto
show "h ` (\<Union>\<F> - C) \<subseteq> rel_frontier T"
using \<open>\<Union>\<G> = \<Union>\<F>\<close> him by auto
show "h x = f x" if "x \<in> S" for x
proof -
have "h x = g x"
apply (rule hg)
using Ssub that by blast
also have "... = f x"
by (simp add: gf that)
finally show "h x = f x" .
qed
show "disjnt C S"
using dis Ssub by (meson disjnt_iff subset_eq)
qed (intro \<open>finite C\<close>)
qed
subsection\<open> Special cases and corollaries involving spheres\<close>
lemma disjnt_Diff1: "X \<subseteq> Y' \<Longrightarrow> disjnt (X - Y) (X' - Y')"
by (auto simp: disjnt_def)
proposition extend_map_affine_to_sphere_cofinite_simple:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "compact S" "convex U" "bounded U"
and aff: "aff_dim T \<le> aff_dim U"
and "S \<subseteq> T" and contf: "continuous_on S f"
and fim: "f ` S \<subseteq> rel_frontier U"
obtains K g where "finite K" "K \<subseteq> T" "disjnt K S" "continuous_on (T - K) g"
"g ` (T - K) \<subseteq> rel_frontier U"
"\<And>x. x \<in> S \<Longrightarrow> g x = f x"
proof -
have "\<exists>K g. finite K \<and> disjnt K S \<and> continuous_on (T - K) g \<and>
g ` (T - K) \<subseteq> rel_frontier U \<and> (\<forall>x \<in> S. g x = f x)"
if "affine T" "S \<subseteq> T" and aff: "aff_dim T \<le> aff_dim U" for T
proof (cases "S = {}")
case True
show ?thesis
proof (cases "rel_frontier U = {}")
case True
with \<open>bounded U\<close> have "aff_dim U \<le> 0"
using affine_bounded_eq_lowdim rel_frontier_eq_empty by auto
with aff have "aff_dim T \<le> 0" by auto
then obtain a where "T \<subseteq> {a}"
using \<open>affine T\<close> affine_bounded_eq_lowdim affine_bounded_eq_trivial by auto
then show ?thesis
using \<open>S = {}\<close> fim
by (metis Diff_cancel contf disjnt_empty2 finite.emptyI finite_insert finite_subset)
next
case False
then obtain a where "a \<in> rel_frontier U"
by auto
then show ?thesis
using continuous_on_const [of _ a] \<open>S = {}\<close> by force
qed
next
case False
have "bounded S"
by (simp add: \<open>compact S\<close> compact_imp_bounded)
then obtain b where b: "S \<subseteq> cbox (-b) b"
using bounded_subset_cbox_symmetric by blast
define bbox where "bbox \<equiv> cbox (-(b+One)) (b+One)"
have "cbox (-b) b \<subseteq> bbox"
by (auto simp: bbox_def algebra_simps intro!: subset_box_imp)
with b \<open>S \<subseteq> T\<close> have "S \<subseteq> bbox \<inter> T"
by auto
then have Ssub: "S \<subseteq> \<Union>{bbox \<inter> T}"
by auto
then have "aff_dim (bbox \<inter> T) \<le> aff_dim U"
by (metis aff aff_dim_subset inf_commute inf_le1 order_trans)
obtain K g where K: "finite K" "disjnt K S"
and contg: "continuous_on (\<Union>{bbox \<inter> T} - K) g"
and gim: "g ` (\<Union>{bbox \<inter> T} - K) \<subseteq> rel_frontier U"
and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
proof (rule extend_map_cell_complex_to_sphere_cofinite
[OF _ Ssub _ \<open>convex U\<close> \<open>bounded U\<close> _ _ _ contf fim])
show "closed S"
using \<open>compact S\<close> compact_eq_bounded_closed by auto
show poly: "\<And>X. X \<in> {bbox \<inter> T} \<Longrightarrow> polytope X"
by (simp add: polytope_Int_polyhedron bbox_def polytope_interval affine_imp_polyhedron \<open>affine T\<close>)
show "\<And>X Y. \<lbrakk>X \<in> {bbox \<inter> T}; Y \<in> {bbox \<inter> T}\<rbrakk> \<Longrightarrow> X \<inter> Y face_of X \<and> X \<inter> Y face_of Y"
by (simp add:poly face_of_refl polytope_imp_convex)
show "\<And>X. X \<in> {bbox \<inter> T} \<Longrightarrow> aff_dim X \<le> aff_dim U"
by (simp add: \<open>aff_dim (bbox \<inter> T) \<le> aff_dim U\<close>)
qed auto
define fro where "fro \<equiv> \<lambda>d. frontier(cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One))"
obtain d where d12: "1/2 \<le> d" "d \<le> 1" and dd: "disjnt K (fro d)"
proof (rule disjoint_family_elem_disjnt [OF _ \<open>finite K\<close>])
show "infinite {1/2..1::real}"
by (simp add: infinite_Icc)
have dis1: "disjnt (fro x) (fro y)" if "x<y" for x y
by (auto simp: algebra_simps that subset_box_imp disjnt_Diff1 frontier_def fro_def)
then show "disjoint_family_on fro {1/2..1}"
by (auto simp: disjoint_family_on_def disjnt_def neq_iff)
qed auto
define c where "c \<equiv> b + d *\<^sub>R One"
have cbsub: "cbox (-b) b \<subseteq> box (-c) c" "cbox (-b) b \<subseteq> cbox (-c) c" "cbox (-c) c \<subseteq> bbox"
using d12 by (auto simp: algebra_simps subset_box_imp c_def bbox_def)
have clo_cbT: "closed (cbox (- c) c \<inter> T)"
by (simp add: affine_closed closed_Int closed_cbox \<open>affine T\<close>)
have cpT_ne: "cbox (- c) c \<inter> T \<noteq> {}"
using \<open>S \<noteq> {}\<close> b cbsub(2) \<open>S \<subseteq> T\<close> by fastforce
have "closest_point (cbox (- c) c \<inter> T) x \<notin> K" if "x \<in> T" "x \<notin> K" for x
proof (cases "x \<in> cbox (-c) c")
case True with that show ?thesis
by (simp add: closest_point_self)
next
case False
have int_ne: "interior (cbox (-c) c) \<inter> T \<noteq> {}"
using \<open>S \<noteq> {}\<close> \<open>S \<subseteq> T\<close> b \<open>cbox (- b) b \<subseteq> box (- c) c\<close> by force
have "convex T"
by (meson \<open>affine T\<close> affine_imp_convex)
then have "x \<in> affine hull (cbox (- c) c \<inter> T)"
by (metis Int_commute Int_iff \<open>S \<noteq> {}\<close> \<open>S \<subseteq> T\<close> cbsub(1) \<open>x \<in> T\<close> affine_hull_convex_Int_nonempty_interior all_not_in_conv b hull_inc inf.orderE interior_cbox)
then have "x \<in> affine hull (cbox (- c) c \<inter> T) - rel_interior (cbox (- c) c \<inter> T)"
by (meson DiffI False Int_iff rel_interior_subset subsetCE)
then have "closest_point (cbox (- c) c \<inter> T) x \<in> rel_frontier (cbox (- c) c \<inter> T)"
by (rule closest_point_in_rel_frontier [OF clo_cbT cpT_ne])
moreover have "(rel_frontier (cbox (- c) c \<inter> T)) \<subseteq> fro d"
apply (subst convex_affine_rel_frontier_Int [OF _ \<open>affine T\<close> int_ne])
apply (auto simp: fro_def c_def)
done
ultimately show ?thesis
using dd by (force simp: disjnt_def)
qed
then have cpt_subset: "closest_point (cbox (- c) c \<inter> T) ` (T - K) \<subseteq> \<Union>{bbox \<inter> T} - K"
using closest_point_in_set [OF clo_cbT cpT_ne] cbsub(3) by force
show ?thesis
proof (intro conjI ballI exI)
have "continuous_on (T - K) (closest_point (cbox (- c) c \<inter> T))"
apply (rule continuous_on_closest_point)
using \<open>S \<noteq> {}\<close> cbsub(2) b that
by (auto simp: affine_imp_convex convex_Int affine_closed closed_Int closed_cbox \<open>affine T\<close>)
then show "continuous_on (T - K) (g \<circ> closest_point (cbox (- c) c \<inter> T))"
by (metis continuous_on_compose continuous_on_subset [OF contg cpt_subset])
have "(g \<circ> closest_point (cbox (- c) c \<inter> T)) ` (T - K) \<subseteq> g ` (\<Union>{bbox \<inter> T} - K)"
by (metis image_comp image_mono cpt_subset)
also have "... \<subseteq> rel_frontier U"
by (rule gim)
finally show "(g \<circ> closest_point (cbox (- c) c \<inter> T)) ` (T - K) \<subseteq> rel_frontier U" .
show "(g \<circ> closest_point (cbox (- c) c \<inter> T)) x = f x" if "x \<in> S" for x
proof -
have "(g \<circ> closest_point (cbox (- c) c \<inter> T)) x = g x"
unfolding o_def
by (metis IntI \<open>S \<subseteq> T\<close> b cbsub(2) closest_point_self subset_eq that)
also have "... = f x"
by (simp add: that gf)
finally show ?thesis .
qed
qed (auto simp: K)
qed
then obtain K g where "finite K" "disjnt K S"
and contg: "continuous_on (affine hull T - K) g"
and gim: "g ` (affine hull T - K) \<subseteq> rel_frontier U"
and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
by (metis aff affine_affine_hull aff_dim_affine_hull
order_trans [OF \<open>S \<subseteq> T\<close> hull_subset [of T affine]])
then obtain K g where "finite K" "disjnt K S"
and contg: "continuous_on (T - K) g"
and gim: "g ` (T - K) \<subseteq> rel_frontier U"
and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
by (rule_tac K=K and g=g in that) (auto simp: hull_inc elim: continuous_on_subset)
then show ?thesis
by (rule_tac K="K \<inter> T" and g=g in that) (auto simp: disjnt_iff Diff_Int contg)
qed
subsection\<open>Extending maps to spheres\<close>
(*Up to extend_map_affine_to_sphere_cofinite_gen*)
lemma extend_map_affine_to_sphere1:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::topological_space"
assumes "finite K" "affine U" and contf: "continuous_on (U - K) f"
and fim: "f ` (U - K) \<subseteq> T"
and comps: "\<And>C. \<lbrakk>C \<in> components(U - S); C \<inter> K \<noteq> {}\<rbrakk> \<Longrightarrow> C \<inter> L \<noteq> {}"
and clo: "closedin (top_of_set U) S" and K: "disjnt K S" "K \<subseteq> U"
obtains g where "continuous_on (U - L) g" "g ` (U - L) \<subseteq> T" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
proof (cases "K = {}")
case True
then show ?thesis
by (metis Diff_empty Diff_subset contf fim continuous_on_subset image_subsetI rev_image_eqI subset_iff that)
next
case False
have "S \<subseteq> U"
using clo closedin_limpt by blast
then have "(U - S) \<inter> K \<noteq> {}"
by (metis Diff_triv False Int_Diff K disjnt_def inf.absorb_iff2 inf_commute)
then have "\<Union>(components (U - S)) \<inter> K \<noteq> {}"
using Union_components by simp
then obtain C0 where C0: "C0 \<in> components (U - S)" "C0 \<inter> K \<noteq> {}"
by blast
have "convex U"
by (simp add: affine_imp_convex \<open>affine U\<close>)
then have "locally connected U"
by (rule convex_imp_locally_connected)
have "\<exists>a g. a \<in> C \<and> a \<in> L \<and> continuous_on (S \<union> (C - {a})) g \<and>
g ` (S \<union> (C - {a})) \<subseteq> T \<and> (\<forall>x \<in> S. g x = f x)"
if C: "C \<in> components (U - S)" and CK: "C \<inter> K \<noteq> {}" for C
proof -
have "C \<subseteq> U-S" "C \<inter> L \<noteq> {}"
by (simp_all add: in_components_subset comps that)
then obtain a where a: "a \<in> C" "a \<in> L" by auto
have opeUC: "openin (top_of_set U) C"
proof (rule openin_trans)
show "openin (top_of_set (U-S)) C"
by (simp add: \<open>locally connected U\<close> clo locally_diff_closed openin_components_locally_connected [OF _ C])
show "openin (top_of_set U) (U - S)"
by (simp add: clo openin_diff)
qed
then obtain d where "C \<subseteq> U" "0 < d" and d: "cball a d \<inter> U \<subseteq> C"
using openin_contains_cball by (metis \<open>a \<in> C\<close>)
then have "ball a d \<inter> U \<subseteq> C"
by auto
obtain h k where homhk: "homeomorphism (S \<union> C) (S \<union> C) h k"
and subC: "{x. (\<not> (h x = x \<and> k x = x))} \<subseteq> C"
and bou: "bounded {x. (\<not> (h x = x \<and> k x = x))}"
and hin: "\<And>x. x \<in> C \<inter> K \<Longrightarrow> h x \<in> ball a d \<inter> U"
proof (rule homeomorphism_grouping_points_exists_gen [of C "ball a d \<inter> U" "C \<inter> K" "S \<union> C"])
show "openin (top_of_set C) (ball a d \<inter> U)"
by (metis open_ball \<open>C \<subseteq> U\<close> \<open>ball a d \<inter> U \<subseteq> C\<close> inf.absorb_iff2 inf.orderE inf_assoc open_openin openin_subtopology)
show "openin (top_of_set (affine hull C)) C"
by (metis \<open>a \<in> C\<close> \<open>openin (top_of_set U) C\<close> affine_hull_eq affine_hull_openin all_not_in_conv \<open>affine U\<close>)
show "ball a d \<inter> U \<noteq> {}"
using \<open>0 < d\<close> \<open>C \<subseteq> U\<close> \<open>a \<in> C\<close> by force
show "finite (C \<inter> K)"
by (simp add: \<open>finite K\<close>)
show "S \<union> C \<subseteq> affine hull C"
by (metis \<open>C \<subseteq> U\<close> \<open>S \<subseteq> U\<close> \<open>a \<in> C\<close> opeUC affine_hull_eq affine_hull_openin all_not_in_conv assms(2) sup.bounded_iff)
show "connected C"
by (metis C in_components_connected)
qed auto
have a_BU: "a \<in> ball a d \<inter> U"
using \<open>0 < d\<close> \<open>C \<subseteq> U\<close> \<open>a \<in> C\<close> by auto
have "rel_frontier (cball a d \<inter> U) retract_of (affine hull (cball a d \<inter> U) - {a})"
apply (rule rel_frontier_retract_of_punctured_affine_hull)
apply (auto simp: \<open>convex U\<close> convex_Int)
by (metis \<open>affine U\<close> convex_cball empty_iff interior_cball a_BU rel_interior_convex_Int_affine)
moreover have "rel_frontier (cball a d \<inter> U) = frontier (cball a d) \<inter> U"
apply (rule convex_affine_rel_frontier_Int)
using a_BU by (force simp: \<open>affine U\<close>)+
moreover have "affine hull (cball a d \<inter> U) = U"
by (metis \<open>convex U\<close> a_BU affine_hull_convex_Int_nonempty_interior affine_hull_eq \<open>affine U\<close> equals0D inf.commute interior_cball)
ultimately have "frontier (cball a d) \<inter> U retract_of (U - {a})"
by metis
then obtain r where contr: "continuous_on (U - {a}) r"
and rim: "r ` (U - {a}) \<subseteq> sphere a d" "r ` (U - {a}) \<subseteq> U"
and req: "\<And>x. x \<in> sphere a d \<inter> U \<Longrightarrow> r x = x"
using \<open>affine U\<close> by (auto simp: retract_of_def retraction_def hull_same)
define j where "j \<equiv> \<lambda>x. if x \<in> ball a d then r x else x"
have kj: "\<And>x. x \<in> S \<Longrightarrow> k (j x) = x"
using \<open>C \<subseteq> U - S\<close> \<open>S \<subseteq> U\<close> \<open>ball a d \<inter> U \<subseteq> C\<close> j_def subC by auto
have Uaeq: "U - {a} = (cball a d - {a}) \<inter> U \<union> (U - ball a d)"
using \<open>0 < d\<close> by auto
have jim: "j ` (S \<union> (C - {a})) \<subseteq> (S \<union> C) - ball a d"
proof clarify
fix y assume "y \<in> S \<union> (C - {a})"
then have "y \<in> U - {a}"
using \<open>C \<subseteq> U - S\<close> \<open>S \<subseteq> U\<close> \<open>a \<in> C\<close> by auto
then have "r y \<in> sphere a d"
using rim by auto
then show "j y \<in> S \<union> C - ball a d"
apply (simp add: j_def)
using \<open>r y \<in> sphere a d\<close> \<open>y \<in> U - {a}\<close> \<open>y \<in> S \<union> (C - {a})\<close> d rim by fastforce
qed
have contj: "continuous_on (U - {a}) j"
unfolding j_def Uaeq
proof (intro continuous_on_cases_local continuous_on_id, simp_all add: req closedin_closed Uaeq [symmetric])
show "\<exists>T. closed T \<and> (cball a d - {a}) \<inter> U = (U - {a}) \<inter> T"
apply (rule_tac x="(cball a d) \<inter> U" in exI)
using affine_closed \<open>affine U\<close> by blast
show "\<exists>T. closed T \<and> U - ball a d = (U - {a}) \<inter> T"
apply (rule_tac x="U - ball a d" in exI)
using \<open>0 < d\<close> by (force simp: affine_closed \<open>affine U\<close> closed_Diff)
show "continuous_on ((cball a d - {a}) \<inter> U) r"
by (force intro: continuous_on_subset [OF contr])
qed
have fT: "x \<in> U - K \<Longrightarrow> f x \<in> T" for x
using fim by blast
show ?thesis
proof (intro conjI exI)
show "continuous_on (S \<union> (C - {a})) (f \<circ> k \<circ> j)"
proof (intro continuous_on_compose)
show "continuous_on (S \<union> (C - {a})) j"
apply (rule continuous_on_subset [OF contj])
using \<open>C \<subseteq> U - S\<close> \<open>S \<subseteq> U\<close> \<open>a \<in> C\<close> by force
show "continuous_on (j ` (S \<union> (C - {a}))) k"
apply (rule continuous_on_subset [OF homeomorphism_cont2 [OF homhk]])
using jim \<open>C \<subseteq> U - S\<close> \<open>S \<subseteq> U\<close> \<open>ball a d \<inter> U \<subseteq> C\<close> j_def by fastforce
show "continuous_on (k ` j ` (S \<union> (C - {a}))) f"
proof (clarify intro!: continuous_on_subset [OF contf])
fix y assume "y \<in> S \<union> (C - {a})"
have ky: "k y \<in> S \<union> C"
using homeomorphism_image2 [OF homhk] \<open>y \<in> S \<union> (C - {a})\<close> by blast
have jy: "j y \<in> S \<union> C - ball a d"
using Un_iff \<open>y \<in> S \<union> (C - {a})\<close> jim by auto
show "k (j y) \<in> U - K"
apply safe
using \<open>C \<subseteq> U\<close> \<open>S \<subseteq> U\<close> homeomorphism_image2 [OF homhk] jy apply blast
by (metis DiffD1 DiffD2 Int_iff Un_iff \<open>disjnt K S\<close> disjnt_def empty_iff hin homeomorphism_apply2 homeomorphism_image2 homhk imageI jy)
qed
qed
have ST: "\<And>x. x \<in> S \<Longrightarrow> (f \<circ> k \<circ> j) x \<in> T"
apply (simp add: kj)
apply (metis DiffI \<open>S \<subseteq> U\<close> \<open>disjnt K S\<close> subsetD disjnt_iff fim image_subset_iff)
done
moreover have "(f \<circ> k \<circ> j) x \<in> T" if "x \<in> C" "x \<noteq> a" "x \<notin> S" for x
proof -
have rx: "r x \<in> sphere a d"
using \<open>C \<subseteq> U\<close> rim that by fastforce
have jj: "j x \<in> S \<union> C - ball a d"
using jim that by blast
have "k (j x) = j x \<longrightarrow> k (j x) \<in> C \<or> j x \<in> C"
by (metis Diff_iff Int_iff Un_iff \<open>S \<subseteq> U\<close> subsetD d j_def jj rx sphere_cball that(1))
then have "k (j x) \<in> C"
using homeomorphism_apply2 [OF homhk, of "j x"] \<open>C \<subseteq> U\<close> \<open>S \<subseteq> U\<close> a rx
by (metis (mono_tags, lifting) Diff_iff subsetD jj mem_Collect_eq subC)
with jj \<open>C \<subseteq> U\<close> show ?thesis
apply safe
using ST j_def apply fastforce
apply (auto simp: not_less intro!: fT)
by (metis DiffD1 DiffD2 Int_iff hin homeomorphism_apply2 [OF homhk] jj)
qed
ultimately show "(f \<circ> k \<circ> j) ` (S \<union> (C - {a})) \<subseteq> T"
by force
show "\<forall>x\<in>S. (f \<circ> k \<circ> j) x = f x" using kj by simp
qed (auto simp: a)
qed
then obtain a h where
ah: "\<And>C. \<lbrakk>C \<in> components (U - S); C \<inter> K \<noteq> {}\<rbrakk>
\<Longrightarrow> a C \<in> C \<and> a C \<in> L \<and> continuous_on (S \<union> (C - {a C})) (h C) \<and>
h C ` (S \<union> (C - {a C})) \<subseteq> T \<and> (\<forall>x \<in> S. h C x = f x)"
using that by metis
define F where "F \<equiv> {C \<in> components (U - S). C \<inter> K \<noteq> {}}"
define G where "G \<equiv> {C \<in> components (U - S). C \<inter> K = {}}"
define UF where "UF \<equiv> (\<Union>C\<in>F. C - {a C})"
have "C0 \<in> F"
by (auto simp: F_def C0)
have "finite F"
proof (subst finite_image_iff [of "\<lambda>C. C \<inter> K" F, symmetric])
show "inj_on (\<lambda>C. C \<inter> K) F"
unfolding F_def inj_on_def
using components_nonoverlap by blast
show "finite ((\<lambda>C. C \<inter> K) ` F)"
unfolding F_def
by (rule finite_subset [of _ "Pow K"]) (auto simp: \<open>finite K\<close>)
qed
obtain g where contg: "continuous_on (S \<union> UF) g"
and gh: "\<And>x i. \<lbrakk>i \<in> F; x \<in> (S \<union> UF) \<inter> (S \<union> (i - {a i}))\<rbrakk>
\<Longrightarrow> g x = h i x"
proof (rule pasting_lemma_exists_closed [OF \<open>finite F\<close>])
let ?X = "top_of_set (S \<union> UF)"
show "topspace ?X \<subseteq> (\<Union>C\<in>F. S \<union> (C - {a C}))"
using \<open>C0 \<in> F\<close> by (force simp: UF_def)
show "closedin (top_of_set (S \<union> UF)) (S \<union> (C - {a C}))"
if "C \<in> F" for C
proof (rule closedin_closed_subset [of U "S \<union> C"])
show "closedin (top_of_set U) (S \<union> C)"
apply (rule closedin_Un_complement_component [OF \<open>locally connected U\<close> clo])
using F_def that by blast
next
have "x = a C'" if "C' \<in> F" "x \<in> C'" "x \<notin> U" for x C'
proof -
have "\<forall>A. x \<in> \<Union>A \<or> C' \<notin> A"
using \<open>x \<in> C'\<close> by blast
with that show "x = a C'"
by (metis (lifting) DiffD1 F_def Union_components mem_Collect_eq)
qed
then show "S \<union> UF \<subseteq> U"
using \<open>S \<subseteq> U\<close> by (force simp: UF_def)
next
show "S \<union> (C - {a C}) = (S \<union> C) \<inter> (S \<union> UF)"
using F_def UF_def components_nonoverlap that by auto
qed
show "continuous_map (subtopology ?X (S \<union> (C' - {a C'}))) euclidean (h C')" if "C' \<in> F" for C'
proof -
have C': "C' \<in> components (U - S)" "C' \<inter> K \<noteq> {}"
using F_def that by blast+
show ?thesis
using ah [OF C'] by (auto simp: F_def subtopology_subtopology intro: continuous_on_subset)
qed
show "\<And>i j x. \<lbrakk>i \<in> F; j \<in> F;
x \<in> topspace ?X \<inter> (S \<union> (i - {a i})) \<inter> (S \<union> (j - {a j}))\<rbrakk>
\<Longrightarrow> h i x = h j x"
using components_eq by (fastforce simp: components_eq F_def ah)
qed auto
have SU': "S \<union> \<Union>G \<union> (S \<union> UF) \<subseteq> U"
using \<open>S \<subseteq> U\<close> in_components_subset by (auto simp: F_def G_def UF_def)
have clo1: "closedin (top_of_set (S \<union> \<Union>G \<union> (S \<union> UF))) (S \<union> \<Union>G)"
proof (rule closedin_closed_subset [OF _ SU'])
have *: "\<And>C. C \<in> F \<Longrightarrow> openin (top_of_set U) C"
unfolding F_def
by clarify (metis (no_types, lifting) \<open>locally connected U\<close> clo closedin_def locally_diff_closed openin_components_locally_connected openin_trans topspace_euclidean_subtopology)
show "closedin (top_of_set U) (U - UF)"
unfolding UF_def
by (force intro: openin_delete *)
show "S \<union> \<Union>G = (U - UF) \<inter> (S \<union> \<Union>G \<union> (S \<union> UF))"
using \<open>S \<subseteq> U\<close> apply (auto simp: F_def G_def UF_def)
apply (metis Diff_iff UnionI Union_components)
apply (metis DiffD1 UnionI Union_components)
by (metis (no_types, lifting) IntI components_nonoverlap empty_iff)
qed
have clo2: "closedin (top_of_set (S \<union> \<Union>G \<union> (S \<union> UF))) (S \<union> UF)"
proof (rule closedin_closed_subset [OF _ SU'])
show "closedin (top_of_set U) (\<Union>C\<in>F. S \<union> C)"
apply (rule closedin_Union)
apply (simp add: \<open>finite F\<close>)
using F_def \<open>locally connected U\<close> clo closedin_Un_complement_component by blast
show "S \<union> UF = (\<Union>C\<in>F. S \<union> C) \<inter> (S \<union> \<Union>G \<union> (S \<union> UF))"
using \<open>S \<subseteq> U\<close> apply (auto simp: F_def G_def UF_def)
using C0 apply blast
by (metis components_nonoverlap disjnt_def disjnt_iff)
qed
have SUG: "S \<union> \<Union>G \<subseteq> U - K"
using \<open>S \<subseteq> U\<close> K apply (auto simp: G_def disjnt_iff)
by (meson Diff_iff subsetD in_components_subset)
then have contf': "continuous_on (S \<union> \<Union>G) f"
by (rule continuous_on_subset [OF contf])
have contg': "continuous_on (S \<union> UF) g"
apply (rule continuous_on_subset [OF contg])
using \<open>S \<subseteq> U\<close> by (auto simp: F_def G_def)
have "\<And>x. \<lbrakk>S \<subseteq> U; x \<in> S\<rbrakk> \<Longrightarrow> f x = g x"
by (subst gh) (auto simp: ah C0 intro: \<open>C0 \<in> F\<close>)
then have f_eq_g: "\<And>x. x \<in> S \<union> UF \<and> x \<in> S \<union> \<Union>G \<Longrightarrow> f x = g x"
using \<open>S \<subseteq> U\<close> apply (auto simp: F_def G_def UF_def dest: in_components_subset)
using components_eq by blast
have cont: "continuous_on (S \<union> \<Union>G \<union> (S \<union> UF)) (\<lambda>x. if x \<in> S \<union> \<Union>G then f x else g x)"
by (blast intro: continuous_on_cases_local [OF clo1 clo2 contf' contg' f_eq_g, of "\<lambda>x. x \<in> S \<union> \<Union>G"])
show ?thesis
proof
have UF: "\<Union>F - L \<subseteq> UF"
unfolding F_def UF_def using ah by blast
have "U - S - L = \<Union>(components (U - S)) - L"
by simp
also have "... = \<Union>F \<union> \<Union>G - L"
unfolding F_def G_def by blast
also have "... \<subseteq> UF \<union> \<Union>G"
using UF by blast
finally have "U - L \<subseteq> S \<union> \<Union>G \<union> (S \<union> UF)"
by blast
then show "continuous_on (U - L) (\<lambda>x. if x \<in> S \<union> \<Union>G then f x else g x)"
by (rule continuous_on_subset [OF cont])
have "((U - L) \<inter> {x. x \<notin> S \<and> (\<forall>xa\<in>G. x \<notin> xa)}) \<subseteq> ((U - L) \<inter> (-S \<inter> UF))"
using \<open>U - L \<subseteq> S \<union> \<Union>G \<union> (S \<union> UF)\<close> by auto
moreover have "g ` ((U - L) \<inter> (-S \<inter> UF)) \<subseteq> T"
proof -
have "g x \<in> T" if "x \<in> U" "x \<notin> L" "x \<notin> S" "C \<in> F" "x \<in> C" "x \<noteq> a C" for x C
proof (subst gh)
show "x \<in> (S \<union> UF) \<inter> (S \<union> (C - {a C}))"
using that by (auto simp: UF_def)
show "h C x \<in> T"
using ah that by (fastforce simp add: F_def)
qed (rule that)
then show ?thesis
by (force simp: UF_def)
qed
ultimately have "g ` ((U - L) \<inter> {x. x \<notin> S \<and> (\<forall>xa\<in>G. x \<notin> xa)}) \<subseteq> T"
using image_mono order_trans by blast
moreover have "f ` ((U - L) \<inter> (S \<union> \<Union>G)) \<subseteq> T"
using fim SUG by blast
ultimately show "(\<lambda>x. if x \<in> S \<union> \<Union>G then f x else g x) ` (U - L) \<subseteq> T"
by force
show "\<And>x. x \<in> S \<Longrightarrow> (if x \<in> S \<union> \<Union>G then f x else g x) = f x"
by (simp add: F_def G_def)
qed
qed
lemma extend_map_affine_to_sphere2:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "compact S" "convex U" "bounded U" "affine T" "S \<subseteq> T"
and affTU: "aff_dim T \<le> aff_dim U"
and contf: "continuous_on S f"
and fim: "f ` S \<subseteq> rel_frontier U"
and ovlap: "\<And>C. C \<in> components(T - S) \<Longrightarrow> C \<inter> L \<noteq> {}"
obtains K g where "finite K" "K \<subseteq> L" "K \<subseteq> T" "disjnt K S"
"continuous_on (T - K) g" "g ` (T - K) \<subseteq> rel_frontier U"
"\<And>x. x \<in> S \<Longrightarrow> g x = f x"
proof -
obtain K g where K: "finite K" "K \<subseteq> T" "disjnt K S"
and contg: "continuous_on (T - K) g"
and gim: "g ` (T - K) \<subseteq> rel_frontier U"
and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
using assms extend_map_affine_to_sphere_cofinite_simple by metis
have "(\<exists>y C. C \<in> components (T - S) \<and> x \<in> C \<and> y \<in> C \<and> y \<in> L)" if "x \<in> K" for x
proof -
have "x \<in> T-S"
using \<open>K \<subseteq> T\<close> \<open>disjnt K S\<close> disjnt_def that by fastforce
then obtain C where "C \<in> components(T - S)" "x \<in> C"
by (metis UnionE Union_components)
with ovlap [of C] show ?thesis
by blast
qed
then obtain \<xi> where \<xi>: "\<And>x. x \<in> K \<Longrightarrow> \<exists>C. C \<in> components (T - S) \<and> x \<in> C \<and> \<xi> x \<in> C \<and> \<xi> x \<in> L"
by metis
obtain h where conth: "continuous_on (T - \<xi> ` K) h"
and him: "h ` (T - \<xi> ` K) \<subseteq> rel_frontier U"
and hg: "\<And>x. x \<in> S \<Longrightarrow> h x = g x"
proof (rule extend_map_affine_to_sphere1 [OF \<open>finite K\<close> \<open>affine T\<close> contg gim, of S "\<xi> ` K"])
show cloTS: "closedin (top_of_set T) S"
by (simp add: \<open>compact S\<close> \<open>S \<subseteq> T\<close> closed_subset compact_imp_closed)
show "\<And>C. \<lbrakk>C \<in> components (T - S); C \<inter> K \<noteq> {}\<rbrakk> \<Longrightarrow> C \<inter> \<xi> ` K \<noteq> {}"
using \<xi> components_eq by blast
qed (use K in auto)
show ?thesis
proof
show *: "\<xi> ` K \<subseteq> L"
using \<xi> by blast
show "finite (\<xi> ` K)"
by (simp add: K)
show "\<xi> ` K \<subseteq> T"
by clarify (meson \<xi> Diff_iff contra_subsetD in_components_subset)
show "continuous_on (T - \<xi> ` K) h"
by (rule conth)
show "disjnt (\<xi> ` K) S"
using K
apply (auto simp: disjnt_def)
by (metis \<xi> DiffD2 UnionI Union_components)
qed (simp_all add: him hg gf)
qed
proposition extend_map_affine_to_sphere_cofinite_gen:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes SUT: "compact S" "convex U" "bounded U" "affine T" "S \<subseteq> T"
and aff: "aff_dim T \<le> aff_dim U"
and contf: "continuous_on S f"
and fim: "f ` S \<subseteq> rel_frontier U"
and dis: "\<And>C. \<lbrakk>C \<in> components(T - S); bounded C\<rbrakk> \<Longrightarrow> C \<inter> L \<noteq> {}"
obtains K g where "finite K" "K \<subseteq> L" "K \<subseteq> T" "disjnt K S" "continuous_on (T - K) g"
"g ` (T - K) \<subseteq> rel_frontier U"
"\<And>x. x \<in> S \<Longrightarrow> g x = f x"
proof (cases "S = {}")
case True
show ?thesis
proof (cases "rel_frontier U = {}")
case True
with aff have "aff_dim T \<le> 0"
apply (simp add: rel_frontier_eq_empty)
using affine_bounded_eq_lowdim \<open>bounded U\<close> order_trans by auto
with aff_dim_geq [of T] consider "aff_dim T = -1" | "aff_dim T = 0"
by linarith
then show ?thesis
proof cases
assume "aff_dim T = -1"
then have "T = {}"
by (simp add: aff_dim_empty)
then show ?thesis
by (rule_tac K="{}" in that) auto
next
assume "aff_dim T = 0"
then obtain a where "T = {a}"
using aff_dim_eq_0 by blast
then have "a \<in> L"
using dis [of "{a}"] \<open>S = {}\<close> by (auto simp: in_components_self)
with \<open>S = {}\<close> \<open>T = {a}\<close> show ?thesis
by (rule_tac K="{a}" and g=f in that) auto
qed
next
case False
then obtain y where "y \<in> rel_frontier U"
by auto
with \<open>S = {}\<close> show ?thesis
by (rule_tac K="{}" and g="\<lambda>x. y" in that) (auto)
qed
next
case False
have "bounded S"
by (simp add: assms compact_imp_bounded)
then obtain b where b: "S \<subseteq> cbox (-b) b"
using bounded_subset_cbox_symmetric by blast
define LU where "LU \<equiv> L \<union> (\<Union> {C \<in> components (T - S). \<not>bounded C} - cbox (-(b+One)) (b+One))"
obtain K g where "finite K" "K \<subseteq> LU" "K \<subseteq> T" "disjnt K S"
and contg: "continuous_on (T - K) g"
and gim: "g ` (T - K) \<subseteq> rel_frontier U"
and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
proof (rule extend_map_affine_to_sphere2 [OF SUT aff contf fim])
show "C \<inter> LU \<noteq> {}" if "C \<in> components (T - S)" for C
proof (cases "bounded C")
case True
with dis that show ?thesis
unfolding LU_def by fastforce
next
case False
then have "\<not> bounded (\<Union>{C \<in> components (T - S). \<not> bounded C})"
by (metis (no_types, lifting) Sup_upper bounded_subset mem_Collect_eq that)
then show ?thesis
apply (clarsimp simp: LU_def Int_Un_distrib Diff_Int_distrib Int_UN_distrib)
by (metis (no_types, lifting) False Sup_upper bounded_cbox bounded_subset inf.orderE mem_Collect_eq that)
qed
qed blast
have *: False if "x \<in> cbox (- b - m *\<^sub>R One) (b + m *\<^sub>R One)"
"x \<notin> box (- b - n *\<^sub>R One) (b + n *\<^sub>R One)"
"0 \<le> m" "m < n" "n \<le> 1" for m n x
using that by (auto simp: mem_box algebra_simps)
have "disjoint_family_on (\<lambda>d. frontier (cbox (- b - d *\<^sub>R One) (b + d *\<^sub>R One))) {1 / 2..1}"
by (auto simp: disjoint_family_on_def neq_iff frontier_def dest: *)
then obtain d where d12: "1/2 \<le> d" "d \<le> 1"
and ddis: "disjnt K (frontier (cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One)))"
using disjoint_family_elem_disjnt [of "{1/2..1::real}" K "\<lambda>d. frontier (cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One))"]
by (auto simp: \<open>finite K\<close>)
define c where "c \<equiv> b + d *\<^sub>R One"
have cbsub: "cbox (-b) b \<subseteq> box (-c) c"
"cbox (-b) b \<subseteq> cbox (-c) c"
"cbox (-c) c \<subseteq> cbox (-(b+One)) (b+One)"
using d12 by (simp_all add: subset_box c_def inner_diff_left inner_left_distrib)
have clo_cT: "closed (cbox (- c) c \<inter> T)"
using affine_closed \<open>affine T\<close> by blast
have cT_ne: "cbox (- c) c \<inter> T \<noteq> {}"
using \<open>S \<noteq> {}\<close> \<open>S \<subseteq> T\<close> b cbsub by fastforce
have S_sub_cc: "S \<subseteq> cbox (- c) c"
using \<open>cbox (- b) b \<subseteq> cbox (- c) c\<close> b by auto
show ?thesis
proof
show "finite (K \<inter> cbox (-(b+One)) (b+One))"
using \<open>finite K\<close> by blast
show "K \<inter> cbox (- (b + One)) (b + One) \<subseteq> L"
using \<open>K \<subseteq> LU\<close> by (auto simp: LU_def)
show "K \<inter> cbox (- (b + One)) (b + One) \<subseteq> T"
using \<open>K \<subseteq> T\<close> by auto
show "disjnt (K \<inter> cbox (- (b + One)) (b + One)) S"
using \<open>disjnt K S\<close> by (simp add: disjnt_def disjoint_eq_subset_Compl inf.coboundedI1)
have cloTK: "closest_point (cbox (- c) c \<inter> T) x \<in> T - K"
if "x \<in> T" and Knot: "x \<in> K \<longrightarrow> x \<notin> cbox (- b - One) (b + One)" for x
proof (cases "x \<in> cbox (- c) c")
case True
with \<open>x \<in> T\<close> show ?thesis
using cbsub(3) Knot by (force simp: closest_point_self)
next
case False
have clo_in_rf: "closest_point (cbox (- c) c \<inter> T) x \<in> rel_frontier (cbox (- c) c \<inter> T)"
proof (intro closest_point_in_rel_frontier [OF clo_cT cT_ne] DiffI notI)
have "T \<inter> interior (cbox (- c) c) \<noteq> {}"
using \<open>S \<noteq> {}\<close> \<open>S \<subseteq> T\<close> b cbsub(1) by fastforce
then show "x \<in> affine hull (cbox (- c) c \<inter> T)"
by (simp add: Int_commute affine_hull_affine_Int_nonempty_interior \<open>affine T\<close> hull_inc that(1))
next
show "False" if "x \<in> rel_interior (cbox (- c) c \<inter> T)"
proof -
have "interior (cbox (- c) c) \<inter> T \<noteq> {}"
using \<open>S \<noteq> {}\<close> \<open>S \<subseteq> T\<close> b cbsub(1) by fastforce
then have "affine hull (T \<inter> cbox (- c) c) = T"
using affine_hull_convex_Int_nonempty_interior [of T "cbox (- c) c"]
by (simp add: affine_imp_convex \<open>affine T\<close> inf_commute)
then show ?thesis
by (meson subsetD le_inf_iff rel_interior_subset that False)
qed
qed
have "closest_point (cbox (- c) c \<inter> T) x \<notin> K"
proof
assume inK: "closest_point (cbox (- c) c \<inter> T) x \<in> K"
have "\<And>x. x \<in> K \<Longrightarrow> x \<notin> frontier (cbox (- (b + d *\<^sub>R One)) (b + d *\<^sub>R One))"
by (metis ddis disjnt_iff)
then show False
by (metis DiffI Int_iff \<open>affine T\<close> cT_ne c_def clo_cT clo_in_rf closest_point_in_set
convex_affine_rel_frontier_Int convex_box(1) empty_iff frontier_cbox inK interior_cbox)
qed
then show ?thesis
using cT_ne clo_cT closest_point_in_set by blast
qed
show "continuous_on (T - K \<inter> cbox (- (b + One)) (b + One)) (g \<circ> closest_point (cbox (-c) c \<inter> T))"
apply (intro continuous_on_compose continuous_on_closest_point continuous_on_subset [OF contg])
apply (simp_all add: clo_cT affine_imp_convex \<open>affine T\<close> convex_Int cT_ne)
using cloTK by blast
have "g (closest_point (cbox (- c) c \<inter> T) x) \<in> rel_frontier U"
if "x \<in> T" "x \<in> K \<longrightarrow> x \<notin> cbox (- b - One) (b + One)" for x
apply (rule gim [THEN subsetD])
using that cloTK by blast
then show "(g \<circ> closest_point (cbox (- c) c \<inter> T)) ` (T - K \<inter> cbox (- (b + One)) (b + One))
\<subseteq> rel_frontier U"
by force
show "\<And>x. x \<in> S \<Longrightarrow> (g \<circ> closest_point (cbox (- c) c \<inter> T)) x = f x"
by simp (metis (mono_tags, lifting) IntI \<open>S \<subseteq> T\<close> cT_ne clo_cT closest_point_refl gf subsetD S_sub_cc)
qed
qed
corollary extend_map_affine_to_sphere_cofinite:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes SUT: "compact S" "affine T" "S \<subseteq> T"
and aff: "aff_dim T \<le> DIM('b)" and "0 \<le> r"
and contf: "continuous_on S f"
and fim: "f ` S \<subseteq> sphere a r"
and dis: "\<And>C. \<lbrakk>C \<in> components(T - S); bounded C\<rbrakk> \<Longrightarrow> C \<inter> L \<noteq> {}"
obtains K g where "finite K" "K \<subseteq> L" "K \<subseteq> T" "disjnt K S" "continuous_on (T - K) g"
"g ` (T - K) \<subseteq> sphere a r" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
proof (cases "r = 0")
case True
with fim show ?thesis
by (rule_tac K="{}" and g = "\<lambda>x. a" in that) (auto)
next
case False
with assms have "0 < r" by auto
then have "aff_dim T \<le> aff_dim (cball a r)"
by (simp add: aff aff_dim_cball)
then show ?thesis
apply (rule extend_map_affine_to_sphere_cofinite_gen
[OF \<open>compact S\<close> convex_cball bounded_cball \<open>affine T\<close> \<open>S \<subseteq> T\<close> _ contf])
using fim apply (auto simp: assms False that dest: dis)
done
qed
corollary extend_map_UNIV_to_sphere_cofinite:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes aff: "DIM('a) \<le> DIM('b)" and "0 \<le> r"
and SUT: "compact S"
and contf: "continuous_on S f"
and fim: "f ` S \<subseteq> sphere a r"
and dis: "\<And>C. \<lbrakk>C \<in> components(- S); bounded C\<rbrakk> \<Longrightarrow> C \<inter> L \<noteq> {}"
obtains K g where "finite K" "K \<subseteq> L" "disjnt K S" "continuous_on (- K) g"
"g ` (- K) \<subseteq> sphere a r" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
apply (rule extend_map_affine_to_sphere_cofinite
[OF \<open>compact S\<close> affine_UNIV subset_UNIV _ \<open>0 \<le> r\<close> contf fim dis])
apply (auto simp: assms that Compl_eq_Diff_UNIV [symmetric])
done
corollary extend_map_UNIV_to_sphere_no_bounded_component:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes aff: "DIM('a) \<le> DIM('b)" and "0 \<le> r"
and SUT: "compact S"
and contf: "continuous_on S f"
and fim: "f ` S \<subseteq> sphere a r"
and dis: "\<And>C. C \<in> components(- S) \<Longrightarrow> \<not> bounded C"
obtains g where "continuous_on UNIV g" "g ` UNIV \<subseteq> sphere a r" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
apply (rule extend_map_UNIV_to_sphere_cofinite [OF aff \<open>0 \<le> r\<close> \<open>compact S\<close> contf fim, of "{}"])
apply (auto simp: that dest: dis)
done
theorem Borsuk_separation_theorem_gen:
fixes S :: "'a::euclidean_space set"
assumes "compact S"
shows "(\<forall>c \<in> components(- S). \<not>bounded c) \<longleftrightarrow>
(\<forall>f. continuous_on S f \<and> f ` S \<subseteq> sphere (0::'a) 1
\<longrightarrow> (\<exists>c. homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) f (\<lambda>x. c)))"
(is "?lhs = ?rhs")
proof
assume L [rule_format]: ?lhs
show ?rhs
proof clarify
fix f :: "'a \<Rightarrow> 'a"
assume contf: "continuous_on S f" and fim: "f ` S \<subseteq> sphere 0 1"
obtain g where contg: "continuous_on UNIV g" and gim: "range g \<subseteq> sphere 0 1"
and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
by (rule extend_map_UNIV_to_sphere_no_bounded_component [OF _ _ \<open>compact S\<close> contf fim L]) auto
then obtain c where c: "homotopic_with_canon (\<lambda>h. True) UNIV (sphere 0 1) g (\<lambda>x. c)"
using contractible_UNIV nullhomotopic_from_contractible by blast
then show "\<exists>c. homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) f (\<lambda>x. c)"
by (metis assms compact_imp_closed contf contg contractible_empty fim gf gim nullhomotopic_from_contractible nullhomotopic_into_sphere_extension)
qed
next
assume R [rule_format]: ?rhs
show ?lhs
unfolding components_def
proof clarify
fix a
assume "a \<notin> S" and a: "bounded (connected_component_set (- S) a)"
have cont: "continuous_on S (\<lambda>x. inverse(norm(x - a)) *\<^sub>R (x - a))"
apply (intro continuous_intros)
using \<open>a \<notin> S\<close> by auto
have im: "(\<lambda>x. inverse(norm(x - a)) *\<^sub>R (x - a)) ` S \<subseteq> sphere 0 1"
by clarsimp (metis \<open>a \<notin> S\<close> eq_iff_diff_eq_0 left_inverse norm_eq_zero)
show False
using R cont im Borsuk_map_essential_bounded_component [OF \<open>compact S\<close> \<open>a \<notin> S\<close>] a by blast
qed
qed
corollary Borsuk_separation_theorem:
fixes S :: "'a::euclidean_space set"
assumes "compact S" and 2: "2 \<le> DIM('a)"
shows "connected(- S) \<longleftrightarrow>
(\<forall>f. continuous_on S f \<and> f ` S \<subseteq> sphere (0::'a) 1
\<longrightarrow> (\<exists>c. homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) f (\<lambda>x. c)))"
(is "?lhs = ?rhs")
proof
assume L: ?lhs
show ?rhs
proof (cases "S = {}")
case True
then show ?thesis by auto
next
case False
then have "(\<forall>c\<in>components (- S). \<not> bounded c)"
by (metis L assms(1) bounded_empty cobounded_imp_unbounded compact_imp_bounded in_components_maximal order_refl)
then show ?thesis
by (simp add: Borsuk_separation_theorem_gen [OF \<open>compact S\<close>])
qed
next
assume R: ?rhs
then show ?lhs
apply (simp add: Borsuk_separation_theorem_gen [OF \<open>compact S\<close>, symmetric])
apply (auto simp: components_def connected_iff_eq_connected_component_set)
using connected_component_in apply fastforce
using cobounded_unique_unbounded_component [OF _ 2, of "-S"] \<open>compact S\<close> compact_eq_bounded_closed by fastforce
qed
lemma homotopy_eqv_separation:
fixes S :: "'a::euclidean_space set" and T :: "'a set"
assumes "S homotopy_eqv T" and "compact S" and "compact T"
shows "connected(- S) \<longleftrightarrow> connected(- T)"
proof -
consider "DIM('a) = 1" | "2 \<le> DIM('a)"
by (metis DIM_ge_Suc0 One_nat_def Suc_1 dual_order.antisym not_less_eq_eq)
then show ?thesis
proof cases
case 1
then show ?thesis
using bounded_connected_Compl_1 compact_imp_bounded homotopy_eqv_empty1 homotopy_eqv_empty2 assms by metis
next
case 2
with assms show ?thesis
by (simp add: Borsuk_separation_theorem homotopy_eqv_cohomotopic_triviality_null)
qed
qed
proposition Jordan_Brouwer_separation:
fixes S :: "'a::euclidean_space set" and a::'a
assumes hom: "S homeomorphic sphere a r" and "0 < r"
shows "\<not> connected(- S)"
proof -
have "- sphere a r \<inter> ball a r \<noteq> {}"
using \<open>0 < r\<close> by (simp add: Int_absorb1 subset_eq)
moreover
have eq: "- sphere a r - ball a r = - cball a r"
by auto
have "- cball a r \<noteq> {}"
proof -
have "frontier (cball a r) \<noteq> {}"
using \<open>0 < r\<close> by auto
then show ?thesis
by (metis frontier_complement frontier_empty)
qed
with eq have "- sphere a r - ball a r \<noteq> {}"
by auto
moreover
have "connected (- S) = connected (- sphere a r)"
proof (rule homotopy_eqv_separation)
show "S homotopy_eqv sphere a r"
using hom homeomorphic_imp_homotopy_eqv by blast
show "compact (sphere a r)"
by simp
then show " compact S"
using hom homeomorphic_compactness by blast
qed
ultimately show ?thesis
using connected_Int_frontier [of "- sphere a r" "ball a r"] by (auto simp: \<open>0 < r\<close>)
qed
proposition Jordan_Brouwer_frontier:
fixes S :: "'a::euclidean_space set" and a::'a
assumes S: "S homeomorphic sphere a r" and T: "T \<in> components(- S)" and 2: "2 \<le> DIM('a)"
shows "frontier T = S"
proof (cases r rule: linorder_cases)
assume "r < 0"
with S T show ?thesis by auto
next
assume "r = 0"
with S T card_eq_SucD obtain b where "S = {b}"
by (auto simp: homeomorphic_finite [of "{a}" S])
have "components (- {b}) = { -{b}}"
using T \<open>S = {b}\<close> by (auto simp: components_eq_sing_iff connected_punctured_universe 2)
with T show ?thesis
by (metis \<open>S = {b}\<close> cball_trivial frontier_cball frontier_complement singletonD sphere_trivial)
next
assume "r > 0"
have "compact S"
using homeomorphic_compactness compact_sphere S by blast
show ?thesis
proof (rule frontier_minimal_separating_closed)
show "closed S"
using \<open>compact S\<close> compact_eq_bounded_closed by blast
show "\<not> connected (- S)"
using Jordan_Brouwer_separation S \<open>0 < r\<close> by blast
obtain f g where hom: "homeomorphism S (sphere a r) f g"
using S by (auto simp: homeomorphic_def)
show "connected (- T)" if "closed T" "T \<subset> S" for T
proof -
have "f ` T \<subseteq> sphere a r"
using \<open>T \<subset> S\<close> hom homeomorphism_image1 by blast
moreover have "f ` T \<noteq> sphere a r"
using \<open>T \<subset> S\<close> hom
by (metis homeomorphism_image2 homeomorphism_of_subsets order_refl psubsetE)
ultimately have "f ` T \<subset> sphere a r" by blast
then have "connected (- f ` T)"
by (rule psubset_sphere_Compl_connected [OF _ \<open>0 < r\<close> 2])
moreover have "compact T"
using \<open>compact S\<close> bounded_subset compact_eq_bounded_closed that by blast
moreover then have "compact (f ` T)"
by (meson compact_continuous_image continuous_on_subset hom homeomorphism_def psubsetE \<open>T \<subset> S\<close>)
moreover have "T homotopy_eqv f ` T"
by (meson \<open>f ` T \<subseteq> sphere a r\<close> dual_order.strict_implies_order hom homeomorphic_def homeomorphic_imp_homotopy_eqv homeomorphism_of_subsets \<open>T \<subset> S\<close>)
ultimately show ?thesis
using homotopy_eqv_separation [of T "f`T"] by blast
qed
qed (rule T)
qed
proposition Jordan_Brouwer_nonseparation:
fixes S :: "'a::euclidean_space set" and a::'a
assumes S: "S homeomorphic sphere a r" and "T \<subset> S" and 2: "2 \<le> DIM('a)"
shows "connected(- T)"
proof -
have *: "connected(C \<union> (S - T))" if "C \<in> components(- S)" for C
proof (rule connected_intermediate_closure)
show "connected C"
using in_components_connected that by auto
have "S = frontier C"
using "2" Jordan_Brouwer_frontier S that by blast
with closure_subset show "C \<union> (S - T) \<subseteq> closure C"
by (auto simp: frontier_def)
qed auto
have "components(- S) \<noteq> {}"
by (metis S bounded_empty cobounded_imp_unbounded compact_eq_bounded_closed compact_sphere
components_eq_empty homeomorphic_compactness)
then have "- T = (\<Union>C \<in> components(- S). C \<union> (S - T))"
using Union_components [of "-S"] \<open>T \<subset> S\<close> by auto
then show ?thesis
apply (rule ssubst)
apply (rule connected_Union)
using \<open>T \<subset> S\<close> apply (auto simp: *)
done
qed
subsection\<open> Invariance of domain and corollaries\<close>
lemma invariance_of_domain_ball:
fixes f :: "'a \<Rightarrow> 'a::euclidean_space"
assumes contf: "continuous_on (cball a r) f" and "0 < r"
and inj: "inj_on f (cball a r)"
shows "open(f ` ball a r)"
proof (cases "DIM('a) = 1")
case True
obtain h::"'a\<Rightarrow>real" and k
where "linear h" "linear k" "h ` UNIV = UNIV" "k ` UNIV = UNIV"
"\<And>x. norm(h x) = norm x" "\<And>x. norm(k x) = norm x"
"\<And>x. k(h x) = x" "\<And>x. h(k x) = x"
apply (rule isomorphisms_UNIV_UNIV [where 'M='a and 'N=real])
using True
apply force
by (metis UNIV_I UNIV_eq_I imageI)
have cont: "continuous_on S h" "continuous_on T k" for S T
by (simp_all add: \<open>linear h\<close> \<open>linear k\<close> linear_continuous_on linear_linear)
have "continuous_on (h ` cball a r) (h \<circ> f \<circ> k)"
apply (intro continuous_on_compose cont continuous_on_subset [OF contf])
apply (auto simp: \<open>\<And>x. k (h x) = x\<close>)
done
moreover have "is_interval (h ` cball a r)"
by (simp add: is_interval_connected_1 \<open>linear h\<close> linear_continuous_on linear_linear connected_continuous_image)
moreover have "inj_on (h \<circ> f \<circ> k) (h ` cball a r)"
using inj by (simp add: inj_on_def) (metis \<open>\<And>x. k (h x) = x\<close>)
ultimately have *: "\<And>T. \<lbrakk>open T; T \<subseteq> h ` cball a r\<rbrakk> \<Longrightarrow> open ((h \<circ> f \<circ> k) ` T)"
using injective_eq_1d_open_map_UNIV by blast
have "open ((h \<circ> f \<circ> k) ` (h ` ball a r))"
by (rule *) (auto simp: \<open>linear h\<close> \<open>range h = UNIV\<close> open_surjective_linear_image)
then have "open ((h \<circ> f) ` ball a r)"
by (simp add: image_comp \<open>\<And>x. k (h x) = x\<close> cong: image_cong)
then show ?thesis
apply (simp only: image_comp [symmetric])
apply (metis open_bijective_linear_image_eq \<open>linear h\<close> \<open>\<And>x. k (h x) = x\<close> \<open>range h = UNIV\<close> bijI inj_on_def)
done
next
case False
then have 2: "DIM('a) \<ge> 2"
by (metis DIM_ge_Suc0 One_nat_def Suc_1 antisym not_less_eq_eq)
have fimsub: "f ` ball a r \<subseteq> - f ` sphere a r"
using inj by clarsimp (metis inj_onD less_eq_real_def mem_cball order_less_irrefl)
have hom: "f ` sphere a r homeomorphic sphere a r"
by (meson compact_sphere contf continuous_on_subset homeomorphic_compact homeomorphic_sym inj inj_on_subset sphere_cball)
then have nconn: "\<not> connected (- f ` sphere a r)"
by (rule Jordan_Brouwer_separation) (auto simp: \<open>0 < r\<close>)
obtain C where C: "C \<in> components (- f ` sphere a r)" and "bounded C"
apply (rule cobounded_has_bounded_component [OF _ nconn])
apply (simp_all add: 2)
by (meson compact_imp_bounded compact_continuous_image_eq compact_sphere contf inj sphere_cball)
moreover have "f ` (ball a r) = C"
proof
have "C \<noteq> {}"
by (rule in_components_nonempty [OF C])
show "C \<subseteq> f ` ball a r"
proof (rule ccontr)
assume nonsub: "\<not> C \<subseteq> f ` ball a r"
have "- f ` cball a r \<subseteq> C"
proof (rule components_maximal [OF C])
have "f ` cball a r homeomorphic cball a r"
using compact_cball contf homeomorphic_compact homeomorphic_sym inj by blast
then show "connected (- f ` cball a r)"
by (auto intro: connected_complement_homeomorphic_convex_compact 2)
show "- f ` cball a r \<subseteq> - f ` sphere a r"
by auto
then show "C \<inter> - f ` cball a r \<noteq> {}"
using \<open>C \<noteq> {}\<close> in_components_subset [OF C] nonsub
using image_iff by fastforce
qed
then have "bounded (- f ` cball a r)"
using bounded_subset \<open>bounded C\<close> by auto
then have "\<not> bounded (f ` cball a r)"
using cobounded_imp_unbounded by blast
then show "False"
using compact_continuous_image [OF contf] compact_cball compact_imp_bounded by blast
qed
with \<open>C \<noteq> {}\<close> have "C \<inter> f ` ball a r \<noteq> {}"
by (simp add: inf.absorb_iff1)
then show "f ` ball a r \<subseteq> C"
by (metis components_maximal [OF C _ fimsub] connected_continuous_image ball_subset_cball connected_ball contf continuous_on_subset)
qed
moreover have "open (- f ` sphere a r)"
using hom compact_eq_bounded_closed compact_sphere homeomorphic_compactness by blast
ultimately show ?thesis
using open_components by blast
qed
text\<open>Proved by L. E. J. Brouwer (1912)\<close>
theorem invariance_of_domain:
fixes f :: "'a \<Rightarrow> 'a::euclidean_space"
assumes "continuous_on S f" "open S" "inj_on f S"
shows "open(f ` S)"
unfolding open_subopen [of "f`S"]
proof clarify
fix a
assume "a \<in> S"
obtain \<delta> where "\<delta> > 0" and \<delta>: "cball a \<delta> \<subseteq> S"
using \<open>open S\<close> \<open>a \<in> S\<close> open_contains_cball_eq by blast
show "\<exists>T. open T \<and> f a \<in> T \<and> T \<subseteq> f ` S"
proof (intro exI conjI)
show "open (f ` (ball a \<delta>))"
by (meson \<delta> \<open>0 < \<delta>\<close> assms continuous_on_subset inj_on_subset invariance_of_domain_ball)
show "f a \<in> f ` ball a \<delta>"
by (simp add: \<open>0 < \<delta>\<close>)
show "f ` ball a \<delta> \<subseteq> f ` S"
using \<delta> ball_subset_cball by blast
qed
qed
lemma inv_of_domain_ss0:
fixes f :: "'a \<Rightarrow> 'a::euclidean_space"
assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f ` U \<subseteq> S"
and "subspace S" and dimS: "dim S = DIM('b::euclidean_space)"
and ope: "openin (top_of_set S) U"
shows "openin (top_of_set S) (f ` U)"
proof -
have "U \<subseteq> S"
using ope openin_imp_subset by blast
have "(UNIV::'b set) homeomorphic S"
by (simp add: \<open>subspace S\<close> dimS homeomorphic_subspaces)
then obtain h k where homhk: "homeomorphism (UNIV::'b set) S h k"
using homeomorphic_def by blast
have homkh: "homeomorphism S (k ` S) k h"
using homhk homeomorphism_image2 homeomorphism_sym by fastforce
have "open ((k \<circ> f \<circ> h) ` k ` U)"
proof (rule invariance_of_domain)
show "continuous_on (k ` U) (k \<circ> f \<circ> h)"
proof (intro continuous_intros)
show "continuous_on (k ` U) h"
by (meson continuous_on_subset [OF homeomorphism_cont1 [OF homhk]] top_greatest)
show "continuous_on (h ` k ` U) f"
apply (rule continuous_on_subset [OF contf], clarify)
apply (metis homhk homeomorphism_def ope openin_imp_subset rev_subsetD)
done
show "continuous_on (f ` h ` k ` U) k"
apply (rule continuous_on_subset [OF homeomorphism_cont2 [OF homhk]])
using fim homhk homeomorphism_apply2 ope openin_subset by fastforce
qed
have ope_iff: "\<And>T. open T \<longleftrightarrow> openin (top_of_set (k ` S)) T"
using homhk homeomorphism_image2 open_openin by fastforce
show "open (k ` U)"
by (simp add: ope_iff homeomorphism_imp_open_map [OF homkh ope])
show "inj_on (k \<circ> f \<circ> h) (k ` U)"
apply (clarsimp simp: inj_on_def)
by (metis subsetD fim homeomorphism_apply2 [OF homhk] image_subset_iff inj_on_eq_iff injf \<open>U \<subseteq> S\<close>)
qed
moreover
have eq: "f ` U = h ` (k \<circ> f \<circ> h \<circ> k) ` U"
unfolding image_comp [symmetric] using \<open>U \<subseteq> S\<close> fim
by (metis homeomorphism_image2 homeomorphism_of_subsets homkh subset_image_iff)
ultimately show ?thesis
by (metis (no_types, hide_lams) homeomorphism_imp_open_map homhk image_comp open_openin subtopology_UNIV)
qed
lemma inv_of_domain_ss1:
fixes f :: "'a \<Rightarrow> 'a::euclidean_space"
assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f ` U \<subseteq> S"
and "subspace S"
and ope: "openin (top_of_set S) U"
shows "openin (top_of_set S) (f ` U)"
proof -
define S' where "S' \<equiv> {y. \<forall>x \<in> S. orthogonal x y}"
have "subspace S'"
by (simp add: S'_def subspace_orthogonal_to_vectors)
define g where "g \<equiv> \<lambda>z::'a*'a. ((f \<circ> fst)z, snd z)"
have "openin (top_of_set (S \<times> S')) (g ` (U \<times> S'))"
proof (rule inv_of_domain_ss0)
show "continuous_on (U \<times> S') g"
apply (simp add: g_def)
apply (intro continuous_intros continuous_on_compose2 [OF contf continuous_on_fst], auto)
done
show "g ` (U \<times> S') \<subseteq> S \<times> S'"
using fim by (auto simp: g_def)
show "inj_on g (U \<times> S')"
using injf by (auto simp: g_def inj_on_def)
show "subspace (S \<times> S')"
by (simp add: \<open>subspace S'\<close> \<open>subspace S\<close> subspace_Times)
show "openin (top_of_set (S \<times> S')) (U \<times> S')"
by (simp add: openin_Times [OF ope])
have "dim (S \<times> S') = dim S + dim S'"
by (simp add: \<open>subspace S'\<close> \<open>subspace S\<close> dim_Times)
also have "... = DIM('a)"
using dim_subspace_orthogonal_to_vectors [OF \<open>subspace S\<close> subspace_UNIV]
by (simp add: add.commute S'_def)
finally show "dim (S \<times> S') = DIM('a)" .
qed
moreover have "g ` (U \<times> S') = f ` U \<times> S'"
by (auto simp: g_def image_iff)
moreover have "0 \<in> S'"
using \<open>subspace S'\<close> subspace_affine by blast
ultimately show ?thesis
by (auto simp: openin_Times_eq)
qed
corollary invariance_of_domain_subspaces:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes ope: "openin (top_of_set U) S"
and "subspace U" "subspace V" and VU: "dim V \<le> dim U"
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V"
and injf: "inj_on f S"
shows "openin (top_of_set V) (f ` S)"
proof -
obtain V' where "subspace V'" "V' \<subseteq> U" "dim V' = dim V"
using choose_subspace_of_subspace [OF VU]
by (metis span_eq_iff \<open>subspace U\<close>)
then have "V homeomorphic V'"
by (simp add: \<open>subspace V\<close> homeomorphic_subspaces)
then obtain h k where homhk: "homeomorphism V V' h k"
using homeomorphic_def by blast
have eq: "f ` S = k ` (h \<circ> f) ` S"
proof -
have "k ` h ` f ` S = f ` S"
by (meson fim homeomorphism_def homeomorphism_of_subsets homhk subset_refl)
then show ?thesis
by (simp add: image_comp)
qed
show ?thesis
unfolding eq
proof (rule homeomorphism_imp_open_map)
show homkh: "homeomorphism V' V k h"
by (simp add: homeomorphism_symD homhk)
have hfV': "(h \<circ> f) ` S \<subseteq> V'"
using fim homeomorphism_image1 homhk by fastforce
moreover have "openin (top_of_set U) ((h \<circ> f) ` S)"
proof (rule inv_of_domain_ss1)
show "continuous_on S (h \<circ> f)"
by (meson contf continuous_on_compose continuous_on_subset fim homeomorphism_cont1 homhk)
show "inj_on (h \<circ> f) S"
apply (clarsimp simp: inj_on_def)
by (metis fim homeomorphism_apply2 [OF homkh] image_subset_iff inj_onD injf)
show "(h \<circ> f) ` S \<subseteq> U"
using \<open>V' \<subseteq> U\<close> hfV' by auto
qed (auto simp: assms)
ultimately show "openin (top_of_set V') ((h \<circ> f) ` S)"
using openin_subset_trans \<open>V' \<subseteq> U\<close> by force
qed
qed
corollary invariance_of_dimension_subspaces:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes ope: "openin (top_of_set U) S"
and "subspace U" "subspace V"
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V"
and injf: "inj_on f S" and "S \<noteq> {}"
shows "dim U \<le> dim V"
proof -
have "False" if "dim V < dim U"
proof -
obtain T where "subspace T" "T \<subseteq> U" "dim T = dim V"
using choose_subspace_of_subspace [of "dim V" U]
by (metis \<open>dim V < dim U\<close> assms(2) order.strict_implies_order span_eq_iff)
then have "V homeomorphic T"
by (simp add: \<open>subspace V\<close> homeomorphic_subspaces)
then obtain h k where homhk: "homeomorphism V T h k"
using homeomorphic_def by blast
have "continuous_on S (h \<circ> f)"
by (meson contf continuous_on_compose continuous_on_subset fim homeomorphism_cont1 homhk)
moreover have "(h \<circ> f) ` S \<subseteq> U"
using \<open>T \<subseteq> U\<close> fim homeomorphism_image1 homhk by fastforce
moreover have "inj_on (h \<circ> f) S"
apply (clarsimp simp: inj_on_def)
by (metis fim homeomorphism_apply1 homhk image_subset_iff inj_onD injf)
ultimately have ope_hf: "openin (top_of_set U) ((h \<circ> f) ` S)"
using invariance_of_domain_subspaces [OF ope \<open>subspace U\<close> \<open>subspace U\<close>] by blast
have "(h \<circ> f) ` S \<subseteq> T"
using fim homeomorphism_image1 homhk by fastforce
then have "dim ((h \<circ> f) ` S) \<le> dim T"
by (rule dim_subset)
also have "dim ((h \<circ> f) ` S) = dim U"
using \<open>S \<noteq> {}\<close> \<open>subspace U\<close>
by (blast intro: dim_openin ope_hf)
finally show False
using \<open>dim V < dim U\<close> \<open>dim T = dim V\<close> by simp
qed
then show ?thesis
using not_less by blast
qed
corollary invariance_of_domain_affine_sets:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes ope: "openin (top_of_set U) S"
and aff: "affine U" "affine V" "aff_dim V \<le> aff_dim U"
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V"
and injf: "inj_on f S"
shows "openin (top_of_set V) (f ` S)"
proof (cases "S = {}")
case True
then show ?thesis by auto
next
case False
obtain a b where "a \<in> S" "a \<in> U" "b \<in> V"
using False fim ope openin_contains_cball by fastforce
have "openin (top_of_set ((+) (- b) ` V)) (((+) (- b) \<circ> f \<circ> (+) a) ` (+) (- a) ` S)"
proof (rule invariance_of_domain_subspaces)
show "openin (top_of_set ((+) (- a) ` U)) ((+) (- a) ` S)"
by (metis ope homeomorphism_imp_open_map homeomorphism_translation translation_galois)
show "subspace ((+) (- a) ` U)"
by (simp add: \<open>a \<in> U\<close> affine_diffs_subspace_subtract \<open>affine U\<close> cong: image_cong_simp)
show "subspace ((+) (- b) ` V)"
by (simp add: \<open>b \<in> V\<close> affine_diffs_subspace_subtract \<open>affine V\<close> cong: image_cong_simp)
show "dim ((+) (- b) ` V) \<le> dim ((+) (- a) ` U)"
by (metis \<open>a \<in> U\<close> \<open>b \<in> V\<close> aff_dim_eq_dim affine_hull_eq aff of_nat_le_iff)
show "continuous_on ((+) (- a) ` S) ((+) (- b) \<circ> f \<circ> (+) a)"
by (metis contf continuous_on_compose homeomorphism_cont2 homeomorphism_translation translation_galois)
show "((+) (- b) \<circ> f \<circ> (+) a) ` (+) (- a) ` S \<subseteq> (+) (- b) ` V"
using fim by auto
show "inj_on ((+) (- b) \<circ> f \<circ> (+) a) ((+) (- a) ` S)"
by (auto simp: inj_on_def) (meson inj_onD injf)
qed
then show ?thesis
by (metis (no_types, lifting) homeomorphism_imp_open_map homeomorphism_translation image_comp translation_galois)
qed
corollary invariance_of_dimension_affine_sets:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes ope: "openin (top_of_set U) S"
and aff: "affine U" "affine V"
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V"
and injf: "inj_on f S" and "S \<noteq> {}"
shows "aff_dim U \<le> aff_dim V"
proof -
obtain a b where "a \<in> S" "a \<in> U" "b \<in> V"
using \<open>S \<noteq> {}\<close> fim ope openin_contains_cball by fastforce
have "dim ((+) (- a) ` U) \<le> dim ((+) (- b) ` V)"
proof (rule invariance_of_dimension_subspaces)
show "openin (top_of_set ((+) (- a) ` U)) ((+) (- a) ` S)"
by (metis ope homeomorphism_imp_open_map homeomorphism_translation translation_galois)
show "subspace ((+) (- a) ` U)"
by (simp add: \<open>a \<in> U\<close> affine_diffs_subspace_subtract \<open>affine U\<close> cong: image_cong_simp)
show "subspace ((+) (- b) ` V)"
by (simp add: \<open>b \<in> V\<close> affine_diffs_subspace_subtract \<open>affine V\<close> cong: image_cong_simp)
show "continuous_on ((+) (- a) ` S) ((+) (- b) \<circ> f \<circ> (+) a)"
by (metis contf continuous_on_compose homeomorphism_cont2 homeomorphism_translation translation_galois)
show "((+) (- b) \<circ> f \<circ> (+) a) ` (+) (- a) ` S \<subseteq> (+) (- b) ` V"
using fim by auto
show "inj_on ((+) (- b) \<circ> f \<circ> (+) a) ((+) (- a) ` S)"
by (auto simp: inj_on_def) (meson inj_onD injf)
qed (use \<open>S \<noteq> {}\<close> in auto)
then show ?thesis
by (metis \<open>a \<in> U\<close> \<open>b \<in> V\<close> aff_dim_eq_dim affine_hull_eq aff of_nat_le_iff)
qed
corollary invariance_of_dimension:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes contf: "continuous_on S f" and "open S"
and injf: "inj_on f S" and "S \<noteq> {}"
shows "DIM('a) \<le> DIM('b)"
using invariance_of_dimension_subspaces [of UNIV S UNIV f] assms
by auto
corollary continuous_injective_image_subspace_dim_le:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "subspace S" "subspace T"
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> T"
and injf: "inj_on f S"
shows "dim S \<le> dim T"
apply (rule invariance_of_dimension_subspaces [of S S _ f])
using assms by (auto simp: subspace_affine)
lemma invariance_of_dimension_convex_domain:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "convex S"
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> affine hull T"
and injf: "inj_on f S"
shows "aff_dim S \<le> aff_dim T"
proof (cases "S = {}")
case True
then show ?thesis by (simp add: aff_dim_geq)
next
case False
have "aff_dim (affine hull S) \<le> aff_dim (affine hull T)"
proof (rule invariance_of_dimension_affine_sets)
show "openin (top_of_set (affine hull S)) (rel_interior S)"
by (simp add: openin_rel_interior)
show "continuous_on (rel_interior S) f"
using contf continuous_on_subset rel_interior_subset by blast
show "f ` rel_interior S \<subseteq> affine hull T"
using fim rel_interior_subset by blast
show "inj_on f (rel_interior S)"
using inj_on_subset injf rel_interior_subset by blast
show "rel_interior S \<noteq> {}"
by (simp add: False \<open>convex S\<close> rel_interior_eq_empty)
qed auto
then show ?thesis
by simp
qed
lemma homeomorphic_convex_sets_le:
assumes "convex S" "S homeomorphic T"
shows "aff_dim S \<le> aff_dim T"
proof -
obtain h k where homhk: "homeomorphism S T h k"
using homeomorphic_def assms by blast
show ?thesis
proof (rule invariance_of_dimension_convex_domain [OF \<open>convex S\<close>])
show "continuous_on S h"
using homeomorphism_def homhk by blast
show "h ` S \<subseteq> affine hull T"
by (metis homeomorphism_def homhk hull_subset)
show "inj_on h S"
by (meson homeomorphism_apply1 homhk inj_on_inverseI)
qed
qed
lemma homeomorphic_convex_sets:
assumes "convex S" "convex T" "S homeomorphic T"
shows "aff_dim S = aff_dim T"
by (meson assms dual_order.antisym homeomorphic_convex_sets_le homeomorphic_sym)
lemma homeomorphic_convex_compact_sets_eq:
assumes "convex S" "compact S" "convex T" "compact T"
shows "S homeomorphic T \<longleftrightarrow> aff_dim S = aff_dim T"
by (meson assms homeomorphic_convex_compact_sets homeomorphic_convex_sets)
lemma invariance_of_domain_gen:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "open S" "continuous_on S f" "inj_on f S" "DIM('b) \<le> DIM('a)"
shows "open(f ` S)"
using invariance_of_domain_subspaces [of UNIV S UNIV f] assms by auto
lemma injective_into_1d_imp_open_map_UNIV:
fixes f :: "'a::euclidean_space \<Rightarrow> real"
assumes "open T" "continuous_on S f" "inj_on f S" "T \<subseteq> S"
shows "open (f ` T)"
apply (rule invariance_of_domain_gen [OF \<open>open T\<close>])
using assms apply (auto simp: elim: continuous_on_subset subset_inj_on)
done
lemma continuous_on_inverse_open:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "open S" "continuous_on S f" "DIM('b) \<le> DIM('a)" and gf: "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x"
shows "continuous_on (f ` S) g"
proof (clarsimp simp add: continuous_openin_preimage_eq)
fix T :: "'a set"
assume "open T"
have eq: "f ` S \<inter> g -` T = f ` (S \<inter> T)"
by (auto simp: gf)
show "openin (top_of_set (f ` S)) (f ` S \<inter> g -` T)"
apply (subst eq)
apply (rule open_openin_trans)
apply (rule invariance_of_domain_gen)
using assms
apply auto
using inj_on_inverseI apply auto[1]
by (metis \<open>open T\<close> continuous_on_subset inj_onI inj_on_subset invariance_of_domain_gen openin_open openin_open_eq)
qed
lemma invariance_of_domain_homeomorphism:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "open S" "continuous_on S f" "DIM('b) \<le> DIM('a)" "inj_on f S"
obtains g where "homeomorphism S (f ` S) f g"
proof
show "homeomorphism S (f ` S) f (inv_into S f)"
by (simp add: assms continuous_on_inverse_open homeomorphism_def)
qed
corollary invariance_of_domain_homeomorphic:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "open S" "continuous_on S f" "DIM('b) \<le> DIM('a)" "inj_on f S"
shows "S homeomorphic (f ` S)"
using invariance_of_domain_homeomorphism [OF assms]
by (meson homeomorphic_def)
lemma continuous_image_subset_interior:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "continuous_on S f" "inj_on f S" "DIM('b) \<le> DIM('a)"
shows "f ` (interior S) \<subseteq> interior(f ` S)"
apply (rule interior_maximal)
apply (simp add: image_mono interior_subset)
apply (rule invariance_of_domain_gen)
using assms
apply (auto simp: subset_inj_on interior_subset continuous_on_subset)
done
lemma homeomorphic_interiors_same_dimension:
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
assumes "S homeomorphic T" and dimeq: "DIM('a) = DIM('b)"
shows "(interior S) homeomorphic (interior T)"
using assms [unfolded homeomorphic_minimal]
unfolding homeomorphic_def
proof (clarify elim!: ex_forward)
fix f g
assume S: "\<forall>x\<in>S. f x \<in> T \<and> g (f x) = x" and T: "\<forall>y\<in>T. g y \<in> S \<and> f (g y) = y"
and contf: "continuous_on S f" and contg: "continuous_on T g"
then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T"
by (auto simp: inj_on_def intro: rev_image_eqI) metis+
have fim: "f ` interior S \<subseteq> interior T"
using continuous_image_subset_interior [OF contf \<open>inj_on f S\<close>] dimeq fST by simp
have gim: "g ` interior T \<subseteq> interior S"
using continuous_image_subset_interior [OF contg \<open>inj_on g T\<close>] dimeq gTS by simp
show "homeomorphism (interior S) (interior T) f g"
unfolding homeomorphism_def
proof (intro conjI ballI)
show "\<And>x. x \<in> interior S \<Longrightarrow> g (f x) = x"
by (meson \<open>\<forall>x\<in>S. f x \<in> T \<and> g (f x) = x\<close> subsetD interior_subset)
have "interior T \<subseteq> f ` interior S"
proof
fix x assume "x \<in> interior T"
then have "g x \<in> interior S"
using gim by blast
then show "x \<in> f ` interior S"
by (metis T \<open>x \<in> interior T\<close> image_iff interior_subset subsetCE)
qed
then show "f ` interior S = interior T"
using fim by blast
show "continuous_on (interior S) f"
by (metis interior_subset continuous_on_subset contf)
show "\<And>y. y \<in> interior T \<Longrightarrow> f (g y) = y"
by (meson T subsetD interior_subset)
have "interior S \<subseteq> g ` interior T"
proof
fix x assume "x \<in> interior S"
then have "f x \<in> interior T"
using fim by blast
then show "x \<in> g ` interior T"
by (metis S \<open>x \<in> interior S\<close> image_iff interior_subset subsetCE)
qed
then show "g ` interior T = interior S"
using gim by blast
show "continuous_on (interior T) g"
by (metis interior_subset continuous_on_subset contg)
qed
qed
lemma homeomorphic_open_imp_same_dimension:
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
assumes "S homeomorphic T" "open S" "S \<noteq> {}" "open T" "T \<noteq> {}"
shows "DIM('a) = DIM('b)"
using assms
apply (simp add: homeomorphic_minimal)
apply (rule order_antisym; metis inj_onI invariance_of_dimension)
done
proposition homeomorphic_interiors:
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
assumes "S homeomorphic T" "interior S = {} \<longleftrightarrow> interior T = {}"
shows "(interior S) homeomorphic (interior T)"
proof (cases "interior T = {}")
case True
with assms show ?thesis by auto
next
case False
then have "DIM('a) = DIM('b)"
using assms
apply (simp add: homeomorphic_minimal)
apply (rule order_antisym; metis continuous_on_subset inj_onI inj_on_subset interior_subset invariance_of_dimension open_interior)
done
then show ?thesis
by (rule homeomorphic_interiors_same_dimension [OF \<open>S homeomorphic T\<close>])
qed
lemma homeomorphic_frontiers_same_dimension:
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
assumes "S homeomorphic T" "closed S" "closed T" and dimeq: "DIM('a) = DIM('b)"
shows "(frontier S) homeomorphic (frontier T)"
using assms [unfolded homeomorphic_minimal]
unfolding homeomorphic_def
proof (clarify elim!: ex_forward)
fix f g
assume S: "\<forall>x\<in>S. f x \<in> T \<and> g (f x) = x" and T: "\<forall>y\<in>T. g y \<in> S \<and> f (g y) = y"
and contf: "continuous_on S f" and contg: "continuous_on T g"
then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T"
by (auto simp: inj_on_def intro: rev_image_eqI) metis+
have "g ` interior T \<subseteq> interior S"
using continuous_image_subset_interior [OF contg \<open>inj_on g T\<close>] dimeq gTS by simp
then have fim: "f ` frontier S \<subseteq> frontier T"
apply (simp add: frontier_def)
using continuous_image_subset_interior assms(2) assms(3) S by auto
have "f ` interior S \<subseteq> interior T"
using continuous_image_subset_interior [OF contf \<open>inj_on f S\<close>] dimeq fST by simp
then have gim: "g ` frontier T \<subseteq> frontier S"
apply (simp add: frontier_def)
using continuous_image_subset_interior T assms(2) assms(3) by auto
show "homeomorphism (frontier S) (frontier T) f g"
unfolding homeomorphism_def
proof (intro conjI ballI)
show gf: "\<And>x. x \<in> frontier S \<Longrightarrow> g (f x) = x"
by (simp add: S assms(2) frontier_def)
show fg: "\<And>y. y \<in> frontier T \<Longrightarrow> f (g y) = y"
by (simp add: T assms(3) frontier_def)
have "frontier T \<subseteq> f ` frontier S"
proof
fix x assume "x \<in> frontier T"
then have "g x \<in> frontier S"
using gim by blast
then show "x \<in> f ` frontier S"
by (metis fg \<open>x \<in> frontier T\<close> imageI)
qed
then show "f ` frontier S = frontier T"
using fim by blast
show "continuous_on (frontier S) f"
by (metis Diff_subset assms(2) closure_eq contf continuous_on_subset frontier_def)
have "frontier S \<subseteq> g ` frontier T"
proof
fix x assume "x \<in> frontier S"
then have "f x \<in> frontier T"
using fim by blast
then show "x \<in> g ` frontier T"
by (metis gf \<open>x \<in> frontier S\<close> imageI)
qed
then show "g ` frontier T = frontier S"
using gim by blast
show "continuous_on (frontier T) g"
by (metis Diff_subset assms(3) closure_closed contg continuous_on_subset frontier_def)
qed
qed
lemma homeomorphic_frontiers:
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
assumes "S homeomorphic T" "closed S" "closed T"
"interior S = {} \<longleftrightarrow> interior T = {}"
shows "(frontier S) homeomorphic (frontier T)"
proof (cases "interior T = {}")
case True
then show ?thesis
by (metis Diff_empty assms closure_eq frontier_def)
next
case False
show ?thesis
apply (rule homeomorphic_frontiers_same_dimension)
apply (simp_all add: assms)
using False assms homeomorphic_interiors homeomorphic_open_imp_same_dimension by blast
qed
lemma continuous_image_subset_rel_interior:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \<subseteq> T"
and TS: "aff_dim T \<le> aff_dim S"
shows "f ` (rel_interior S) \<subseteq> rel_interior(f ` S)"
proof (rule rel_interior_maximal)
show "f ` rel_interior S \<subseteq> f ` S"
by(simp add: image_mono rel_interior_subset)
show "openin (top_of_set (affine hull f ` S)) (f ` rel_interior S)"
proof (rule invariance_of_domain_affine_sets)
show "openin (top_of_set (affine hull S)) (rel_interior S)"
by (simp add: openin_rel_interior)
show "aff_dim (affine hull f ` S) \<le> aff_dim (affine hull S)"
by (metis aff_dim_affine_hull aff_dim_subset fim TS order_trans)
show "f ` rel_interior S \<subseteq> affine hull f ` S"
by (meson \<open>f ` rel_interior S \<subseteq> f ` S\<close> hull_subset order_trans)
show "continuous_on (rel_interior S) f"
using contf continuous_on_subset rel_interior_subset by blast
show "inj_on f (rel_interior S)"
using inj_on_subset injf rel_interior_subset by blast
qed auto
qed
lemma homeomorphic_rel_interiors_same_dimension:
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
assumes "S homeomorphic T" and aff: "aff_dim S = aff_dim T"
shows "(rel_interior S) homeomorphic (rel_interior T)"
using assms [unfolded homeomorphic_minimal]
unfolding homeomorphic_def
proof (clarify elim!: ex_forward)
fix f g
assume S: "\<forall>x\<in>S. f x \<in> T \<and> g (f x) = x" and T: "\<forall>y\<in>T. g y \<in> S \<and> f (g y) = y"
and contf: "continuous_on S f" and contg: "continuous_on T g"
then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T"
by (auto simp: inj_on_def intro: rev_image_eqI) metis+
have fim: "f ` rel_interior S \<subseteq> rel_interior T"
by (metis \<open>inj_on f S\<close> aff contf continuous_image_subset_rel_interior fST order_refl)
have gim: "g ` rel_interior T \<subseteq> rel_interior S"
by (metis \<open>inj_on g T\<close> aff contg continuous_image_subset_rel_interior gTS order_refl)
show "homeomorphism (rel_interior S) (rel_interior T) f g"
unfolding homeomorphism_def
proof (intro conjI ballI)
show gf: "\<And>x. x \<in> rel_interior S \<Longrightarrow> g (f x) = x"
using S rel_interior_subset by blast
show fg: "\<And>y. y \<in> rel_interior T \<Longrightarrow> f (g y) = y"
using T mem_rel_interior_ball by blast
have "rel_interior T \<subseteq> f ` rel_interior S"
proof
fix x assume "x \<in> rel_interior T"
then have "g x \<in> rel_interior S"
using gim by blast
then show "x \<in> f ` rel_interior S"
by (metis fg \<open>x \<in> rel_interior T\<close> imageI)
qed
moreover have "f ` rel_interior S \<subseteq> rel_interior T"
by (metis \<open>inj_on f S\<close> aff contf continuous_image_subset_rel_interior fST order_refl)
ultimately show "f ` rel_interior S = rel_interior T"
by blast
show "continuous_on (rel_interior S) f"
using contf continuous_on_subset rel_interior_subset by blast
have "rel_interior S \<subseteq> g ` rel_interior T"
proof
fix x assume "x \<in> rel_interior S"
then have "f x \<in> rel_interior T"
using fim by blast
then show "x \<in> g ` rel_interior T"
by (metis gf \<open>x \<in> rel_interior S\<close> imageI)
qed
then show "g ` rel_interior T = rel_interior S"
using gim by blast
show "continuous_on (rel_interior T) g"
using contg continuous_on_subset rel_interior_subset by blast
qed
qed
lemma homeomorphic_rel_interiors:
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
assumes "S homeomorphic T" "rel_interior S = {} \<longleftrightarrow> rel_interior T = {}"
shows "(rel_interior S) homeomorphic (rel_interior T)"
proof (cases "rel_interior T = {}")
case True
with assms show ?thesis by auto
next
case False
obtain f g
where S: "\<forall>x\<in>S. f x \<in> T \<and> g (f x) = x" and T: "\<forall>y\<in>T. g y \<in> S \<and> f (g y) = y"
and contf: "continuous_on S f" and contg: "continuous_on T g"
using assms [unfolded homeomorphic_minimal] by auto
have "aff_dim (affine hull S) \<le> aff_dim (affine hull T)"
apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior S" _ f])
apply (simp_all add: openin_rel_interior False assms)
using contf continuous_on_subset rel_interior_subset apply blast
apply (meson S hull_subset image_subsetI rel_interior_subset rev_subsetD)
apply (metis S inj_on_inverseI inj_on_subset rel_interior_subset)
done
moreover have "aff_dim (affine hull T) \<le> aff_dim (affine hull S)"
apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior T" _ g])
apply (simp_all add: openin_rel_interior False assms)
using contg continuous_on_subset rel_interior_subset apply blast
apply (meson T hull_subset image_subsetI rel_interior_subset rev_subsetD)
apply (metis T inj_on_inverseI inj_on_subset rel_interior_subset)
done
ultimately have "aff_dim S = aff_dim T" by force
then show ?thesis
by (rule homeomorphic_rel_interiors_same_dimension [OF \<open>S homeomorphic T\<close>])
qed
lemma homeomorphic_rel_boundaries_same_dimension:
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
assumes "S homeomorphic T" and aff: "aff_dim S = aff_dim T"
shows "(S - rel_interior S) homeomorphic (T - rel_interior T)"
using assms [unfolded homeomorphic_minimal]
unfolding homeomorphic_def
proof (clarify elim!: ex_forward)
fix f g
assume S: "\<forall>x\<in>S. f x \<in> T \<and> g (f x) = x" and T: "\<forall>y\<in>T. g y \<in> S \<and> f (g y) = y"
and contf: "continuous_on S f" and contg: "continuous_on T g"
then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T"
by (auto simp: inj_on_def intro: rev_image_eqI) metis+
have fim: "f ` rel_interior S \<subseteq> rel_interior T"
by (metis \<open>inj_on f S\<close> aff contf continuous_image_subset_rel_interior fST order_refl)
have gim: "g ` rel_interior T \<subseteq> rel_interior S"
by (metis \<open>inj_on g T\<close> aff contg continuous_image_subset_rel_interior gTS order_refl)
show "homeomorphism (S - rel_interior S) (T - rel_interior T) f g"
unfolding homeomorphism_def
proof (intro conjI ballI)
show gf: "\<And>x. x \<in> S - rel_interior S \<Longrightarrow> g (f x) = x"
using S rel_interior_subset by blast
show fg: "\<And>y. y \<in> T - rel_interior T \<Longrightarrow> f (g y) = y"
using T mem_rel_interior_ball by blast
show "f ` (S - rel_interior S) = T - rel_interior T"
using S fST fim gim by auto
show "continuous_on (S - rel_interior S) f"
using contf continuous_on_subset rel_interior_subset by blast
show "g ` (T - rel_interior T) = S - rel_interior S"
using T gTS gim fim by auto
show "continuous_on (T - rel_interior T) g"
using contg continuous_on_subset rel_interior_subset by blast
qed
qed
lemma homeomorphic_rel_boundaries:
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
assumes "S homeomorphic T" "rel_interior S = {} \<longleftrightarrow> rel_interior T = {}"
shows "(S - rel_interior S) homeomorphic (T - rel_interior T)"
proof (cases "rel_interior T = {}")
case True
with assms show ?thesis by auto
next
case False
obtain f g
where S: "\<forall>x\<in>S. f x \<in> T \<and> g (f x) = x" and T: "\<forall>y\<in>T. g y \<in> S \<and> f (g y) = y"
and contf: "continuous_on S f" and contg: "continuous_on T g"
using assms [unfolded homeomorphic_minimal] by auto
have "aff_dim (affine hull S) \<le> aff_dim (affine hull T)"
apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior S" _ f])
apply (simp_all add: openin_rel_interior False assms)
using contf continuous_on_subset rel_interior_subset apply blast
apply (meson S hull_subset image_subsetI rel_interior_subset rev_subsetD)
apply (metis S inj_on_inverseI inj_on_subset rel_interior_subset)
done
moreover have "aff_dim (affine hull T) \<le> aff_dim (affine hull S)"
apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior T" _ g])
apply (simp_all add: openin_rel_interior False assms)
using contg continuous_on_subset rel_interior_subset apply blast
apply (meson T hull_subset image_subsetI rel_interior_subset rev_subsetD)
apply (metis T inj_on_inverseI inj_on_subset rel_interior_subset)
done
ultimately have "aff_dim S = aff_dim T" by force
then show ?thesis
by (rule homeomorphic_rel_boundaries_same_dimension [OF \<open>S homeomorphic T\<close>])
qed
proposition uniformly_continuous_homeomorphism_UNIV_trivial:
fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
assumes contf: "uniformly_continuous_on S f" and hom: "homeomorphism S UNIV f g"
shows "S = UNIV"
proof (cases "S = {}")
case True
then show ?thesis
by (metis UNIV_I hom empty_iff homeomorphism_def image_eqI)
next
case False
have "inj g"
by (metis UNIV_I hom homeomorphism_apply2 injI)
then have "open (g ` UNIV)"
by (blast intro: invariance_of_domain hom homeomorphism_cont2)
then have "open S"
using hom homeomorphism_image2 by blast
moreover have "complete S"
unfolding complete_def
proof clarify
fix \<sigma>
assume \<sigma>: "\<forall>n. \<sigma> n \<in> S" and "Cauchy \<sigma>"
have "Cauchy (f o \<sigma>)"
using uniformly_continuous_imp_Cauchy_continuous \<open>Cauchy \<sigma>\<close> \<sigma> contf by blast
then obtain l where "(f \<circ> \<sigma>) \<longlonglongrightarrow> l"
by (auto simp: convergent_eq_Cauchy [symmetric])
show "\<exists>l\<in>S. \<sigma> \<longlonglongrightarrow> l"
proof
show "g l \<in> S"
using hom homeomorphism_image2 by blast
have "(g \<circ> (f \<circ> \<sigma>)) \<longlonglongrightarrow> g l"
by (meson UNIV_I \<open>(f \<circ> \<sigma>) \<longlonglongrightarrow> l\<close> continuous_on_sequentially hom homeomorphism_cont2)
then show "\<sigma> \<longlonglongrightarrow> g l"
proof -
have "\<forall>n. \<sigma> n = (g \<circ> (f \<circ> \<sigma>)) n"
by (metis (no_types) \<sigma> comp_eq_dest_lhs hom homeomorphism_apply1)
then show ?thesis
by (metis (no_types) LIMSEQ_iff \<open>(g \<circ> (f \<circ> \<sigma>)) \<longlonglongrightarrow> g l\<close>)
qed
qed
qed
then have "closed S"
by (simp add: complete_eq_closed)
ultimately show ?thesis
using clopen [of S] False by simp
qed
subsection\<open>Formulation of loop homotopy in terms of maps out of type complex\<close>
lemma homotopic_circlemaps_imp_homotopic_loops:
assumes "homotopic_with_canon (\<lambda>h. True) (sphere 0 1) S f g"
shows "homotopic_loops S (f \<circ> exp \<circ> (\<lambda>t. 2 * of_real pi * of_real t * \<i>))
(g \<circ> exp \<circ> (\<lambda>t. 2 * of_real pi * of_real t * \<i>))"
proof -
have "homotopic_with_canon (\<lambda>f. True) {z. cmod z = 1} S f g"
using assms by (auto simp: sphere_def)
moreover have "continuous_on {0..1} (exp \<circ> (\<lambda>t. 2 * of_real pi * of_real t * \<i>))"
by (intro continuous_intros)
moreover have "(exp \<circ> (\<lambda>t. 2 * of_real pi * of_real t * \<i>)) ` {0..1} \<subseteq> {z. cmod z = 1}"
by (auto simp: norm_mult)
ultimately
show ?thesis
apply (simp add: homotopic_loops_def comp_assoc)
apply (rule homotopic_with_compose_continuous_right)
apply (auto simp: pathstart_def pathfinish_def)
done
qed
lemma homotopic_loops_imp_homotopic_circlemaps:
assumes "homotopic_loops S p q"
shows "homotopic_with_canon (\<lambda>h. True) (sphere 0 1) S
(p \<circ> (\<lambda>z. (Arg2pi z / (2 * pi))))
(q \<circ> (\<lambda>z. (Arg2pi z / (2 * pi))))"
proof -
obtain h where conth: "continuous_on ({0..1::real} \<times> {0..1}) h"
and him: "h ` ({0..1} \<times> {0..1}) \<subseteq> S"
and h0: "(\<forall>x. h (0, x) = p x)"
and h1: "(\<forall>x. h (1, x) = q x)"
and h01: "(\<forall>t\<in>{0..1}. h (t, 1) = h (t, 0)) "
using assms
by (auto simp: homotopic_loops_def sphere_def homotopic_with_def pathstart_def pathfinish_def)
define j where "j \<equiv> \<lambda>z. if 0 \<le> Im (snd z)
then h (fst z, Arg2pi (snd z) / (2 * pi))
else h (fst z, 1 - Arg2pi (cnj (snd z)) / (2 * pi))"
have Arg2pi_eq: "1 - Arg2pi (cnj y) / (2 * pi) = Arg2pi y / (2 * pi) \<or> Arg2pi y = 0 \<and> Arg2pi (cnj y) = 0" if "cmod y = 1" for y
using that Arg2pi_eq_0_pi Arg2pi_eq_pi by (force simp: Arg2pi_cnj field_split_simps)
show ?thesis
proof (simp add: homotopic_with; intro conjI ballI exI)
show "continuous_on ({0..1} \<times> sphere 0 1) (\<lambda>w. h (fst w, Arg2pi (snd w) / (2 * pi)))"
proof (rule continuous_on_eq)
show j: "j x = h (fst x, Arg2pi (snd x) / (2 * pi))" if "x \<in> {0..1} \<times> sphere 0 1" for x
using Arg2pi_eq that h01 by (force simp: j_def)
have eq: "S = S \<inter> (UNIV \<times> {z. 0 \<le> Im z}) \<union> S \<inter> (UNIV \<times> {z. Im z \<le> 0})" for S :: "(real*complex)set"
by auto
have c1: "continuous_on ({0..1} \<times> sphere 0 1 \<inter> UNIV \<times> {z. 0 \<le> Im z}) (\<lambda>x. h (fst x, Arg2pi (snd x) / (2 * pi)))"
apply (intro continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF continuous_on_upperhalf_Arg2pi])
apply (auto simp: Arg2pi)
apply (meson Arg2pi_lt_2pi linear not_le)
done
have c2: "continuous_on ({0..1} \<times> sphere 0 1 \<inter> UNIV \<times> {z. Im z \<le> 0}) (\<lambda>x. h (fst x, 1 - Arg2pi (cnj (snd x)) / (2 * pi)))"
apply (intro continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF continuous_on_upperhalf_Arg2pi])
apply (auto simp: Arg2pi)
apply (meson Arg2pi_lt_2pi linear not_le)
done
show "continuous_on ({0..1} \<times> sphere 0 1) j"
apply (simp add: j_def)
apply (subst eq)
apply (rule continuous_on_cases_local)
apply (simp_all add: eq [symmetric] closedin_closed_Int closed_Times closed_halfspace_Im_le closed_halfspace_Im_ge c1 c2)
using Arg2pi_eq h01
by force
qed
have "(\<lambda>w. h (fst w, Arg2pi (snd w) / (2 * pi))) ` ({0..1} \<times> sphere 0 1) \<subseteq> h ` ({0..1} \<times> {0..1})"
by (auto simp: Arg2pi_ge_0 Arg2pi_lt_2pi less_imp_le)
also have "... \<subseteq> S"
using him by blast
finally show "(\<lambda>w. h (fst w, Arg2pi (snd w) / (2 * pi))) ` ({0..1} \<times> sphere 0 1) \<subseteq> S" .
qed (auto simp: h0 h1)
qed
lemma simply_connected_homotopic_loops:
"simply_connected S \<longleftrightarrow>
(\<forall>p q. homotopic_loops S p p \<and> homotopic_loops S q q \<longrightarrow> homotopic_loops S p q)"
unfolding simply_connected_def using homotopic_loops_refl by metis
lemma simply_connected_eq_homotopic_circlemaps1:
fixes f :: "complex \<Rightarrow> 'a::topological_space" and g :: "complex \<Rightarrow> 'a"
assumes S: "simply_connected S"
and contf: "continuous_on (sphere 0 1) f" and fim: "f ` (sphere 0 1) \<subseteq> S"
and contg: "continuous_on (sphere 0 1) g" and gim: "g ` (sphere 0 1) \<subseteq> S"
shows "homotopic_with_canon (\<lambda>h. True) (sphere 0 1) S f g"
proof -
have "homotopic_loops S (f \<circ> exp \<circ> (\<lambda>t. of_real(2 * pi * t) * \<i>)) (g \<circ> exp \<circ> (\<lambda>t. of_real(2 * pi * t) * \<i>))"
apply (rule S [unfolded simply_connected_homotopic_loops, rule_format])
apply (simp add: homotopic_circlemaps_imp_homotopic_loops contf fim contg gim)
done
then show ?thesis
apply (rule homotopic_with_eq [OF homotopic_loops_imp_homotopic_circlemaps])
apply (auto simp: o_def complex_norm_eq_1_exp mult.commute)
done
qed
lemma simply_connected_eq_homotopic_circlemaps2a:
fixes h :: "complex \<Rightarrow> 'a::topological_space"
assumes conth: "continuous_on (sphere 0 1) h" and him: "h ` (sphere 0 1) \<subseteq> S"
and hom: "\<And>f g::complex \<Rightarrow> 'a.
\<lbrakk>continuous_on (sphere 0 1) f; f ` (sphere 0 1) \<subseteq> S;
continuous_on (sphere 0 1) g; g ` (sphere 0 1) \<subseteq> S\<rbrakk>
\<Longrightarrow> homotopic_with_canon (\<lambda>h. True) (sphere 0 1) S f g"
shows "\<exists>a. homotopic_with_canon (\<lambda>h. True) (sphere 0 1) S h (\<lambda>x. a)"
apply (rule_tac x="h 1" in exI)
apply (rule hom)
using assms by (auto)
lemma simply_connected_eq_homotopic_circlemaps2b:
fixes S :: "'a::real_normed_vector set"
assumes "\<And>f g::complex \<Rightarrow> 'a.
\<lbrakk>continuous_on (sphere 0 1) f; f ` (sphere 0 1) \<subseteq> S;
continuous_on (sphere 0 1) g; g ` (sphere 0 1) \<subseteq> S\<rbrakk>
\<Longrightarrow> homotopic_with_canon (\<lambda>h. True) (sphere 0 1) S f g"
shows "path_connected S"
proof (clarsimp simp add: path_connected_eq_homotopic_points)
fix a b
assume "a \<in> S" "b \<in> S"
then show "homotopic_loops S (linepath a a) (linepath b b)"
using homotopic_circlemaps_imp_homotopic_loops [OF assms [of "\<lambda>x. a" "\<lambda>x. b"]]
by (auto simp: o_def linepath_def)
qed
lemma simply_connected_eq_homotopic_circlemaps3:
fixes h :: "complex \<Rightarrow> 'a::real_normed_vector"
assumes "path_connected S"
and hom: "\<And>f::complex \<Rightarrow> 'a.
\<lbrakk>continuous_on (sphere 0 1) f; f `(sphere 0 1) \<subseteq> S\<rbrakk>
\<Longrightarrow> \<exists>a. homotopic_with_canon (\<lambda>h. True) (sphere 0 1) S f (\<lambda>x. a)"
shows "simply_connected S"
proof (clarsimp simp add: simply_connected_eq_contractible_loop_some assms)
fix p
assume p: "path p" "path_image p \<subseteq> S" "pathfinish p = pathstart p"
then have "homotopic_loops S p p"
by (simp add: homotopic_loops_refl)
then obtain a where homp: "homotopic_with_canon (\<lambda>h. True) (sphere 0 1) S (p \<circ> (\<lambda>z. Arg2pi z / (2 * pi))) (\<lambda>x. a)"
by (metis homotopic_with_imp_subset2 homotopic_loops_imp_homotopic_circlemaps homotopic_with_imp_continuous hom)
show "\<exists>a. a \<in> S \<and> homotopic_loops S p (linepath a a)"
proof (intro exI conjI)
show "a \<in> S"
using homotopic_with_imp_subset2 [OF homp]
by (metis dist_0_norm image_subset_iff mem_sphere norm_one)
have teq: "\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk>
\<Longrightarrow> t = Arg2pi (exp (2 * of_real pi * of_real t * \<i>)) / (2 * pi) \<or> t=1 \<and> Arg2pi (exp (2 * of_real pi * of_real t * \<i>)) = 0"
apply (rule disjCI)
using Arg2pi_of_real [of 1] apply (auto simp: Arg2pi_exp)
done
have "homotopic_loops S p (p \<circ> (\<lambda>z. Arg2pi z / (2 * pi)) \<circ> exp \<circ> (\<lambda>t. 2 * complex_of_real pi * complex_of_real t * \<i>))"
apply (rule homotopic_loops_eq [OF p])
using p teq apply (fastforce simp: pathfinish_def pathstart_def)
done
then
show "homotopic_loops S p (linepath a a)"
by (simp add: linepath_refl homotopic_loops_trans [OF _ homotopic_circlemaps_imp_homotopic_loops [OF homp, simplified K_record_comp]])
qed
qed
proposition simply_connected_eq_homotopic_circlemaps:
fixes S :: "'a::real_normed_vector set"
shows "simply_connected S \<longleftrightarrow>
(\<forall>f g::complex \<Rightarrow> 'a.
continuous_on (sphere 0 1) f \<and> f ` (sphere 0 1) \<subseteq> S \<and>
continuous_on (sphere 0 1) g \<and> g ` (sphere 0 1) \<subseteq> S
\<longrightarrow> homotopic_with_canon (\<lambda>h. True) (sphere 0 1) S f g)"
apply (rule iffI)
apply (blast elim: dest: simply_connected_eq_homotopic_circlemaps1)
by (simp add: simply_connected_eq_homotopic_circlemaps2a simply_connected_eq_homotopic_circlemaps2b simply_connected_eq_homotopic_circlemaps3)
proposition simply_connected_eq_contractible_circlemap:
fixes S :: "'a::real_normed_vector set"
shows "simply_connected S \<longleftrightarrow>
path_connected S \<and>
(\<forall>f::complex \<Rightarrow> 'a.
continuous_on (sphere 0 1) f \<and> f `(sphere 0 1) \<subseteq> S
\<longrightarrow> (\<exists>a. homotopic_with_canon (\<lambda>h. True) (sphere 0 1) S f (\<lambda>x. a)))"
apply (rule iffI)
apply (simp add: simply_connected_eq_homotopic_circlemaps1 simply_connected_eq_homotopic_circlemaps2a simply_connected_eq_homotopic_circlemaps2b)
using simply_connected_eq_homotopic_circlemaps3 by blast
corollary homotopy_eqv_simple_connectedness:
fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set"
shows "S homotopy_eqv T \<Longrightarrow> simply_connected S \<longleftrightarrow> simply_connected T"
by (simp add: simply_connected_eq_homotopic_circlemaps homotopy_eqv_homotopic_triviality)
subsection\<open>Homeomorphism of simple closed curves to circles\<close>
proposition homeomorphic_simple_path_image_circle:
fixes a :: complex and \<gamma> :: "real \<Rightarrow> 'a::t2_space"
assumes "simple_path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>" and "0 < r"
shows "(path_image \<gamma>) homeomorphic sphere a r"
proof -
have "homotopic_loops (path_image \<gamma>) \<gamma> \<gamma>"
by (simp add: assms homotopic_loops_refl simple_path_imp_path)
then have hom: "homotopic_with_canon (\<lambda>h. True) (sphere 0 1) (path_image \<gamma>)
(\<gamma> \<circ> (\<lambda>z. Arg2pi z / (2*pi))) (\<gamma> \<circ> (\<lambda>z. Arg2pi z / (2*pi)))"
by (rule homotopic_loops_imp_homotopic_circlemaps)
have "\<exists>g. homeomorphism (sphere 0 1) (path_image \<gamma>) (\<gamma> \<circ> (\<lambda>z. Arg2pi z / (2*pi))) g"
proof (rule homeomorphism_compact)
show "continuous_on (sphere 0 1) (\<gamma> \<circ> (\<lambda>z. Arg2pi z / (2*pi)))"
using hom homotopic_with_imp_continuous by blast
show "inj_on (\<gamma> \<circ> (\<lambda>z. Arg2pi z / (2*pi))) (sphere 0 1)"
proof
fix x y
assume xy: "x \<in> sphere 0 1" "y \<in> sphere 0 1"
and eq: "(\<gamma> \<circ> (\<lambda>z. Arg2pi z / (2*pi))) x = (\<gamma> \<circ> (\<lambda>z. Arg2pi z / (2*pi))) y"
then have "(Arg2pi x / (2*pi)) = (Arg2pi y / (2*pi))"
proof -
have "(Arg2pi x / (2*pi)) \<in> {0..1}" "(Arg2pi y / (2*pi)) \<in> {0..1}"
using Arg2pi_ge_0 Arg2pi_lt_2pi dual_order.strict_iff_order by fastforce+
with eq show ?thesis
using \<open>simple_path \<gamma>\<close> Arg2pi_lt_2pi unfolding simple_path_def o_def
by (metis eq_divide_eq_1 not_less_iff_gr_or_eq)
qed
with xy show "x = y"
by (metis is_Arg_def Arg2pi Arg2pi_0 dist_0_norm divide_cancel_right dual_order.strict_iff_order mem_sphere)
qed
have "\<And>z. cmod z = 1 \<Longrightarrow> \<exists>x\<in>{0..1}. \<gamma> (Arg2pi z / (2*pi)) = \<gamma> x"
by (metis Arg2pi_ge_0 Arg2pi_lt_2pi atLeastAtMost_iff divide_less_eq_1 less_eq_real_def zero_less_mult_iff pi_gt_zero zero_le_divide_iff zero_less_numeral)
moreover have "\<exists>z\<in>sphere 0 1. \<gamma> x = \<gamma> (Arg2pi z / (2*pi))" if "0 \<le> x" "x \<le> 1" for x
proof (cases "x=1")
case True
with Arg2pi_of_real [of 1] loop show ?thesis
by (rule_tac x=1 in bexI) (auto simp: pathfinish_def pathstart_def \<open>0 \<le> x\<close>)
next
case False
then have *: "(Arg2pi (exp (\<i>*(2* of_real pi* of_real x))) / (2*pi)) = x"
using that by (auto simp: Arg2pi_exp field_split_simps)
show ?thesis
by (rule_tac x="exp(\<i> * of_real(2*pi*x))" in bexI) (auto simp: *)
qed
ultimately show "(\<gamma> \<circ> (\<lambda>z. Arg2pi z / (2*pi))) ` sphere 0 1 = path_image \<gamma>"
by (auto simp: path_image_def image_iff)
qed auto
then have "path_image \<gamma> homeomorphic sphere (0::complex) 1"
using homeomorphic_def homeomorphic_sym by blast
also have "... homeomorphic sphere a r"
by (simp add: assms homeomorphic_spheres)
finally show ?thesis .
qed
lemma homeomorphic_simple_path_images:
fixes \<gamma>1 :: "real \<Rightarrow> 'a::t2_space" and \<gamma>2 :: "real \<Rightarrow> 'b::t2_space"
assumes "simple_path \<gamma>1" and loop: "pathfinish \<gamma>1 = pathstart \<gamma>1"
assumes "simple_path \<gamma>2" and loop: "pathfinish \<gamma>2 = pathstart \<gamma>2"
shows "(path_image \<gamma>1) homeomorphic (path_image \<gamma>2)"
by (meson assms homeomorphic_simple_path_image_circle homeomorphic_sym homeomorphic_trans loop pi_gt_zero)
subsection\<open>Dimension-based conditions for various homeomorphisms\<close>
lemma homeomorphic_subspaces_eq:
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
assumes "subspace S" "subspace T"
shows "S homeomorphic T \<longleftrightarrow> dim S = dim T"
proof
assume "S homeomorphic T"
then obtain f g where hom: "homeomorphism S T f g"
using homeomorphic_def by blast
show "dim S = dim T"
proof (rule order_antisym)
show "dim S \<le> dim T"
by (metis assms dual_order.refl inj_onI homeomorphism_cont1 [OF hom] homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] continuous_injective_image_subspace_dim_le)
show "dim T \<le> dim S"
by (metis assms dual_order.refl inj_onI homeomorphism_cont2 [OF hom] homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] continuous_injective_image_subspace_dim_le)
qed
next
assume "dim S = dim T"
then show "S homeomorphic T"
by (simp add: assms homeomorphic_subspaces)
qed
lemma homeomorphic_affine_sets_eq:
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
assumes "affine S" "affine T"
shows "S homeomorphic T \<longleftrightarrow> aff_dim S = aff_dim T"
proof (cases "S = {} \<or> T = {}")
case True
then show ?thesis
using assms homeomorphic_affine_sets by force
next
case False
then obtain a b where "a \<in> S" "b \<in> T"
by blast
then have "subspace ((+) (- a) ` S)" "subspace ((+) (- b) ` T)"
using affine_diffs_subspace assms by blast+
then show ?thesis
by (metis affine_imp_convex assms homeomorphic_affine_sets homeomorphic_convex_sets)
qed
lemma homeomorphic_hyperplanes_eq:
fixes a :: "'a::euclidean_space" and c :: "'b::euclidean_space"
assumes "a \<noteq> 0" "c \<noteq> 0"
shows "({x. a \<bullet> x = b} homeomorphic {x. c \<bullet> x = d} \<longleftrightarrow> DIM('a) = DIM('b))"
apply (auto simp: homeomorphic_affine_sets_eq affine_hyperplane assms)
by (metis DIM_positive Suc_pred)
lemma homeomorphic_UNIV_UNIV:
shows "(UNIV::'a set) homeomorphic (UNIV::'b set) \<longleftrightarrow>
DIM('a::euclidean_space) = DIM('b::euclidean_space)"
by (simp add: homeomorphic_subspaces_eq)
lemma simply_connected_sphere_gen:
assumes "convex S" "bounded S" and 3: "3 \<le> aff_dim S"
shows "simply_connected(rel_frontier S)"
proof -
have pa: "path_connected (rel_frontier S)"
using assms by (simp add: path_connected_sphere_gen)
show ?thesis
proof (clarsimp simp add: simply_connected_eq_contractible_circlemap pa)
fix f
assume f: "continuous_on (sphere (0::complex) 1) f" "f ` sphere 0 1 \<subseteq> rel_frontier S"
have eq: "sphere (0::complex) 1 = rel_frontier(cball 0 1)"
by simp
have "convex (cball (0::complex) 1)"
by (rule convex_cball)
then obtain c where "homotopic_with_canon (\<lambda>z. True) (sphere (0::complex) 1) (rel_frontier S) f (\<lambda>x. c)"
apply (rule inessential_spheremap_lowdim_gen [OF _ bounded_cball \<open>convex S\<close> \<open>bounded S\<close>, where f=f])
using f 3
apply (auto simp: aff_dim_cball)
done
then show "\<exists>a. homotopic_with_canon (\<lambda>h. True) (sphere 0 1) (rel_frontier S) f (\<lambda>x. a)"
by blast
qed
qed
subsection\<open>more invariance of domain\<close>(*FIX ME title? *)
proposition invariance_of_domain_sphere_affine_set_gen:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \<subseteq> T"
and U: "bounded U" "convex U"
and "affine T" and affTU: "aff_dim T < aff_dim U"
and ope: "openin (top_of_set (rel_frontier U)) S"
shows "openin (top_of_set T) (f ` S)"
proof (cases "rel_frontier U = {}")
case True
then show ?thesis
using ope openin_subset by force
next
case False
obtain b c where b: "b \<in> rel_frontier U" and c: "c \<in> rel_frontier U" and "b \<noteq> c"
using \<open>bounded U\<close> rel_frontier_not_sing [of U] subset_singletonD False by fastforce
obtain V :: "'a set" where "affine V" and affV: "aff_dim V = aff_dim U - 1"
proof (rule choose_affine_subset [OF affine_UNIV])
show "- 1 \<le> aff_dim U - 1"
by (metis aff_dim_empty aff_dim_geq aff_dim_negative_iff affTU diff_0 diff_right_mono not_le)
show "aff_dim U - 1 \<le> aff_dim (UNIV::'a set)"
by (metis aff_dim_UNIV aff_dim_le_DIM le_cases not_le zle_diff1_eq)
qed auto
have SU: "S \<subseteq> rel_frontier U"
using ope openin_imp_subset by auto
have homb: "rel_frontier U - {b} homeomorphic V"
and homc: "rel_frontier U - {c} homeomorphic V"
using homeomorphic_punctured_sphere_affine_gen [of U _ V]
by (simp_all add: \<open>affine V\<close> affV U b c)
then obtain g h j k
where gh: "homeomorphism (rel_frontier U - {b}) V g h"
and jk: "homeomorphism (rel_frontier U - {c}) V j k"
by (auto simp: homeomorphic_def)
with SU have hgsub: "(h ` g ` (S - {b})) \<subseteq> S" and kjsub: "(k ` j ` (S - {c})) \<subseteq> S"
by (simp_all add: homeomorphism_def subset_eq)
have [simp]: "aff_dim T \<le> aff_dim V"
by (simp add: affTU affV)
have "openin (top_of_set T) ((f \<circ> h) ` g ` (S - {b}))"
proof (rule invariance_of_domain_affine_sets [OF _ \<open>affine V\<close>])
show "openin (top_of_set V) (g ` (S - {b}))"
apply (rule homeomorphism_imp_open_map [OF gh])
by (meson Diff_mono Diff_subset SU ope openin_delete openin_subset_trans order_refl)
show "continuous_on (g ` (S - {b})) (f \<circ> h)"
apply (rule continuous_on_compose)
apply (meson Diff_mono SU homeomorphism_def homeomorphism_of_subsets gh set_eq_subset)
using contf continuous_on_subset hgsub by blast
show "inj_on (f \<circ> h) (g ` (S - {b}))"
using kjsub
apply (clarsimp simp add: inj_on_def)
by (metis SU b homeomorphism_def inj_onD injf insert_Diff insert_iff gh rev_subsetD)
show "(f \<circ> h) ` g ` (S - {b}) \<subseteq> T"
by (metis fim image_comp image_mono hgsub subset_trans)
qed (auto simp: assms)
moreover
have "openin (top_of_set T) ((f \<circ> k) ` j ` (S - {c}))"
proof (rule invariance_of_domain_affine_sets [OF _ \<open>affine V\<close>])
show "openin (top_of_set V) (j ` (S - {c}))"
apply (rule homeomorphism_imp_open_map [OF jk])
by (meson Diff_mono Diff_subset SU ope openin_delete openin_subset_trans order_refl)
show "continuous_on (j ` (S - {c})) (f \<circ> k)"
apply (rule continuous_on_compose)
apply (meson Diff_mono SU homeomorphism_def homeomorphism_of_subsets jk set_eq_subset)
using contf continuous_on_subset kjsub by blast
show "inj_on (f \<circ> k) (j ` (S - {c}))"
using kjsub
apply (clarsimp simp add: inj_on_def)
by (metis SU c homeomorphism_def inj_onD injf insert_Diff insert_iff jk rev_subsetD)
show "(f \<circ> k) ` j ` (S - {c}) \<subseteq> T"
by (metis fim image_comp image_mono kjsub subset_trans)
qed (auto simp: assms)
ultimately have "openin (top_of_set T) ((f \<circ> h) ` g ` (S - {b}) \<union> ((f \<circ> k) ` j ` (S - {c})))"
by (rule openin_Un)
moreover have "(f \<circ> h) ` g ` (S - {b}) = f ` (S - {b})"
proof -
have "h ` g ` (S - {b}) = (S - {b})"
proof
show "h ` g ` (S - {b}) \<subseteq> S - {b}"
using homeomorphism_apply1 [OF gh] SU
by (fastforce simp add: image_iff image_subset_iff)
show "S - {b} \<subseteq> h ` g ` (S - {b})"
apply clarify
by (metis SU subsetD homeomorphism_apply1 [OF gh] image_iff member_remove remove_def)
qed
then show ?thesis
by (metis image_comp)
qed
moreover have "(f \<circ> k) ` j ` (S - {c}) = f ` (S - {c})"
proof -
have "k ` j ` (S - {c}) = (S - {c})"
proof
show "k ` j ` (S - {c}) \<subseteq> S - {c}"
using homeomorphism_apply1 [OF jk] SU
by (fastforce simp add: image_iff image_subset_iff)
show "S - {c} \<subseteq> k ` j ` (S - {c})"
apply clarify
by (metis SU subsetD homeomorphism_apply1 [OF jk] image_iff member_remove remove_def)
qed
then show ?thesis
by (metis image_comp)
qed
moreover have "f ` (S - {b}) \<union> f ` (S - {c}) = f ` (S)"
using \<open>b \<noteq> c\<close> by blast
ultimately show ?thesis
by simp
qed
lemma invariance_of_domain_sphere_affine_set:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \<subseteq> T"
and "r \<noteq> 0" "affine T" and affTU: "aff_dim T < DIM('a)"
and ope: "openin (top_of_set (sphere a r)) S"
shows "openin (top_of_set T) (f ` S)"
proof (cases "sphere a r = {}")
case True
then show ?thesis
using ope openin_subset by force
next
case False
show ?thesis
proof (rule invariance_of_domain_sphere_affine_set_gen [OF contf injf fim bounded_cball convex_cball \<open>affine T\<close>])
show "aff_dim T < aff_dim (cball a r)"
by (metis False affTU aff_dim_cball assms(4) linorder_cases sphere_empty)
show "openin (top_of_set (rel_frontier (cball a r))) S"
by (simp add: \<open>r \<noteq> 0\<close> ope)
qed
qed
lemma no_embedding_sphere_lowdim:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes contf: "continuous_on (sphere a r) f" and injf: "inj_on f (sphere a r)" and "r > 0"
shows "DIM('a) \<le> DIM('b)"
proof -
have "False" if "DIM('a) > DIM('b)"
proof -
have "compact (f ` sphere a r)"
using compact_continuous_image
by (simp add: compact_continuous_image contf)
then have "\<not> open (f ` sphere a r)"
using compact_open
by (metis assms(3) image_is_empty not_less_iff_gr_or_eq sphere_eq_empty)
then show False
using invariance_of_domain_sphere_affine_set [OF contf injf subset_UNIV] \<open>r > 0\<close>
by (metis aff_dim_UNIV affine_UNIV less_irrefl of_nat_less_iff open_openin openin_subtopology_self subtopology_UNIV that)
qed
then show ?thesis
using not_less by blast
qed
lemma simply_connected_sphere:
fixes a :: "'a::euclidean_space"
assumes "3 \<le> DIM('a)"
shows "simply_connected(sphere a r)"
proof (cases rule: linorder_cases [of r 0])
case less
then show ?thesis by simp
next
case equal
then show ?thesis by (auto simp: convex_imp_simply_connected)
next
case greater
then show ?thesis
using simply_connected_sphere_gen [of "cball a r"] assms
by (simp add: aff_dim_cball)
qed
lemma simply_connected_sphere_eq:
fixes a :: "'a::euclidean_space"
shows "simply_connected(sphere a r) \<longleftrightarrow> 3 \<le> DIM('a) \<or> r \<le> 0" (is "?lhs = ?rhs")
proof (cases "r \<le> 0")
case True
have "simply_connected (sphere a r)"
apply (rule convex_imp_simply_connected)
using True less_eq_real_def by auto
with True show ?thesis by auto
next
case False
show ?thesis
proof
assume L: ?lhs
have "False" if "DIM('a) = 1 \<or> DIM('a) = 2"
using that
proof
assume "DIM('a) = 1"
with L show False
using connected_sphere_eq simply_connected_imp_connected
by (metis False Suc_1 not_less_eq_eq order_refl)
next
assume "DIM('a) = 2"
then have "sphere a r homeomorphic sphere (0::complex) 1"
by (metis DIM_complex False homeomorphic_spheres_gen not_less zero_less_one)
then have "simply_connected(sphere (0::complex) 1)"
using L homeomorphic_simply_connected_eq by blast
then obtain a::complex where "homotopic_with_canon (\<lambda>h. True) (sphere 0 1) (sphere 0 1) id (\<lambda>x. a)"
apply (simp add: simply_connected_eq_contractible_circlemap)
by (metis continuous_on_id' id_apply image_id subset_refl)
then show False
using contractible_sphere contractible_def not_one_le_zero by blast
qed
with False show ?rhs
apply simp
by (metis DIM_ge_Suc0 le_antisym not_less_eq_eq numeral_2_eq_2 numeral_3_eq_3)
next
assume ?rhs
with False show ?lhs by (simp add: simply_connected_sphere)
qed
qed
lemma simply_connected_punctured_universe_eq:
fixes a :: "'a::euclidean_space"
shows "simply_connected(- {a}) \<longleftrightarrow> 3 \<le> DIM('a)"
proof -
have [simp]: "a \<in> rel_interior (cball a 1)"
by (simp add: rel_interior_nonempty_interior)
have [simp]: "affine hull cball a 1 - {a} = -{a}"
by (metis Compl_eq_Diff_UNIV aff_dim_cball aff_dim_lt_full not_less_iff_gr_or_eq zero_less_one)
have "simply_connected(- {a}) \<longleftrightarrow> simply_connected(sphere a 1)"
apply (rule sym)
apply (rule homotopy_eqv_simple_connectedness)
using homotopy_eqv_rel_frontier_punctured_affine_hull [of "cball a 1" a] apply auto
done
also have "... \<longleftrightarrow> 3 \<le> DIM('a)"
by (simp add: simply_connected_sphere_eq)
finally show ?thesis .
qed
lemma not_simply_connected_circle:
fixes a :: complex
shows "0 < r \<Longrightarrow> \<not> simply_connected(sphere a r)"
by (simp add: simply_connected_sphere_eq)
proposition simply_connected_punctured_convex:
fixes a :: "'a::euclidean_space"
assumes "convex S" and 3: "3 \<le> aff_dim S"
shows "simply_connected(S - {a})"
proof (cases "a \<in> rel_interior S")
case True
then obtain e where "a \<in> S" "0 < e" and e: "cball a e \<inter> affine hull S \<subseteq> S"
by (auto simp: rel_interior_cball)
have con: "convex (cball a e \<inter> affine hull S)"
by (simp add: convex_Int)
have bo: "bounded (cball a e \<inter> affine hull S)"
by (simp add: bounded_Int)
have "affine hull S \<inter> interior (cball a e) \<noteq> {}"
using \<open>0 < e\<close> \<open>a \<in> S\<close> hull_subset by fastforce
then have "3 \<le> aff_dim (affine hull S \<inter> cball a e)"
by (simp add: 3 aff_dim_convex_Int_nonempty_interior [OF convex_affine_hull])
also have "... = aff_dim (cball a e \<inter> affine hull S)"
by (simp add: Int_commute)
finally have "3 \<le> aff_dim (cball a e \<inter> affine hull S)" .
moreover have "rel_frontier (cball a e \<inter> affine hull S) homotopy_eqv S - {a}"
proof (rule homotopy_eqv_rel_frontier_punctured_convex)
show "a \<in> rel_interior (cball a e \<inter> affine hull S)"
by (meson IntI Int_mono \<open>a \<in> S\<close> \<open>0 < e\<close> e \<open>cball a e \<inter> affine hull S \<subseteq> S\<close> ball_subset_cball centre_in_cball dual_order.strict_implies_order hull_inc hull_mono mem_rel_interior_ball)
have "closed (cball a e \<inter> affine hull S)"
by blast
then show "rel_frontier (cball a e \<inter> affine hull S) \<subseteq> S"
apply (simp add: rel_frontier_def)
using e by blast
show "S \<subseteq> affine hull (cball a e \<inter> affine hull S)"
by (metis (no_types, lifting) IntI \<open>a \<in> S\<close> \<open>0 < e\<close> affine_hull_convex_Int_nonempty_interior centre_in_ball convex_affine_hull empty_iff hull_subset inf_commute interior_cball subsetCE subsetI)
qed (auto simp: assms con bo)
ultimately show ?thesis
using homotopy_eqv_simple_connectedness simply_connected_sphere_gen [OF con bo]
by blast
next
case False
show ?thesis
apply (rule contractible_imp_simply_connected)
apply (rule contractible_convex_tweak_boundary_points [OF \<open>convex S\<close>])
apply (simp add: False rel_interior_subset subset_Diff_insert)
by (meson Diff_subset closure_subset subset_trans)
qed
corollary simply_connected_punctured_universe:
fixes a :: "'a::euclidean_space"
assumes "3 \<le> DIM('a)"
shows "simply_connected(- {a})"
proof -
have [simp]: "affine hull cball a 1 = UNIV"
apply auto
by (metis UNIV_I aff_dim_cball aff_dim_lt_full zero_less_one not_less_iff_gr_or_eq)
have "simply_connected (rel_frontier (cball a 1)) = simply_connected (affine hull cball a 1 - {a})"
apply (rule homotopy_eqv_simple_connectedness)
apply (rule homotopy_eqv_rel_frontier_punctured_affine_hull)
apply (force simp: rel_interior_cball intro: homotopy_eqv_simple_connectedness homotopy_eqv_rel_frontier_punctured_affine_hull)+
done
then show ?thesis
using simply_connected_sphere [of a 1, OF assms] by (auto simp: Compl_eq_Diff_UNIV)
qed
subsection\<open>The power, squaring and exponential functions as covering maps\<close>
proposition covering_space_power_punctured_plane:
assumes "0 < n"
shows "covering_space (- {0}) (\<lambda>z::complex. z^n) (- {0})"
proof -
consider "n = 1" | "2 \<le> n" using assms by linarith
then obtain e where "0 < e"
and e: "\<And>w z. cmod(w - z) < e * cmod z \<Longrightarrow> (w^n = z^n \<longleftrightarrow> w = z)"
proof cases
assume "n = 1" then show ?thesis
by (rule_tac e=1 in that) auto
next
assume "2 \<le> n"
have eq_if_pow_eq:
"w = z" if lt: "cmod (w - z) < 2 * sin (pi / real n) * cmod z"
and eq: "w^n = z^n" for w z
proof (cases "z = 0")
case True with eq assms show ?thesis by (auto simp: power_0_left)
next
case False
then have "z \<noteq> 0" by auto
have "(w/z)^n = 1"
by (metis False divide_self_if eq power_divide power_one)
then obtain j where j: "w / z = exp (2 * of_real pi * \<i> * j / n)" and "j < n"
using Suc_leI assms \<open>2 \<le> n\<close> complex_roots_unity [THEN eqset_imp_iff, of n "w/z"]
by force
have "cmod (w/z - 1) < 2 * sin (pi / real n)"
using lt assms \<open>z \<noteq> 0\<close> by (simp add: field_split_simps norm_divide)
then have "cmod (exp (\<i> * of_real (2 * pi * j / n)) - 1) < 2 * sin (pi / real n)"
by (simp add: j field_simps)
then have "2 * \<bar>sin((2 * pi * j / n) / 2)\<bar> < 2 * sin (pi / real n)"
by (simp only: dist_exp_i_1)
then have sin_less: "sin((pi * j / n)) < sin (pi / real n)"
by (simp add: field_simps)
then have "w / z = 1"
proof (cases "j = 0")
case True then show ?thesis by (auto simp: j)
next
case False
then have "sin (pi / real n) \<le> sin((pi * j / n))"
proof (cases "j / n \<le> 1/2")
case True
show ?thesis
apply (rule sin_monotone_2pi_le)
using \<open>j \<noteq> 0 \<close> \<open>j < n\<close> True
apply (auto simp: field_simps intro: order_trans [of _ 0])
done
next
case False
then have seq: "sin(pi * j / n) = sin(pi * (n - j) / n)"
using \<open>j < n\<close> by (simp add: algebra_simps diff_divide_distrib of_nat_diff)
show ?thesis
apply (simp only: seq)
apply (rule sin_monotone_2pi_le)
using \<open>j < n\<close> False
apply (auto simp: field_simps intro: order_trans [of _ 0])
done
qed
with sin_less show ?thesis by force
qed
then show ?thesis by simp
qed
show ?thesis
apply (rule_tac e = "2 * sin(pi / n)" in that)
apply (force simp: \<open>2 \<le> n\<close> sin_pi_divide_n_gt_0)
apply (meson eq_if_pow_eq)
done
qed
have zn1: "continuous_on (- {0}) (\<lambda>z::complex. z^n)"
by (rule continuous_intros)+
have zn2: "(\<lambda>z::complex. z^n) ` (- {0}) = - {0}"
using assms by (auto simp: image_def elim: exists_complex_root_nonzero [where n = n])
have zn3: "\<exists>T. z^n \<in> T \<and> open T \<and> 0 \<notin> T \<and>
(\<exists>v. \<Union>v = -{0} \<inter> (\<lambda>z. z ^ n) -` T \<and>
(\<forall>u\<in>v. open u \<and> 0 \<notin> u) \<and>
pairwise disjnt v \<and>
(\<forall>u\<in>v. Ex (homeomorphism u T (\<lambda>z. z^n))))"
if "z \<noteq> 0" for z::complex
proof -
define d where "d \<equiv> min (1/2) (e/4) * norm z"
have "0 < d"
by (simp add: d_def \<open>0 < e\<close> \<open>z \<noteq> 0\<close>)
have iff_x_eq_y: "x^n = y^n \<longleftrightarrow> x = y"
if eq: "w^n = z^n" and x: "x \<in> ball w d" and y: "y \<in> ball w d" for w x y
proof -
have [simp]: "norm z = norm w" using that
by (simp add: assms power_eq_imp_eq_norm)
show ?thesis
proof (cases "w = 0")
case True with \<open>z \<noteq> 0\<close> assms eq
show ?thesis by (auto simp: power_0_left)
next
case False
have "cmod (x - y) < 2*d"
using x y
by (simp add: dist_norm [symmetric]) (metis dist_commute mult_2 dist_triangle_less_add)
also have "... \<le> 2 * e / 4 * norm w"
using \<open>e > 0\<close> by (simp add: d_def min_mult_distrib_right)
also have "... = e * (cmod w / 2)"
by simp
also have "... \<le> e * cmod y"
apply (rule mult_left_mono)
using \<open>e > 0\<close> y
apply (simp_all add: dist_norm d_def min_mult_distrib_right del: divide_const_simps)
apply (metis dist_0_norm dist_complex_def dist_triangle_half_l linorder_not_less order_less_irrefl)
done
finally have "cmod (x - y) < e * cmod y" .
then show ?thesis by (rule e)
qed
qed
then have inj: "inj_on (\<lambda>w. w^n) (ball z d)"
by (simp add: inj_on_def)
have cont: "continuous_on (ball z d) (\<lambda>w. w ^ n)"
by (intro continuous_intros)
have noncon: "\<not> (\<lambda>w::complex. w^n) constant_on UNIV"
by (metis UNIV_I assms constant_on_def power_one zero_neq_one zero_power)
have im_eq: "(\<lambda>w. w^n) ` ball z' d = (\<lambda>w. w^n) ` ball z d"
if z': "z'^n = z^n" for z'
proof -
have nz': "norm z' = norm z" using that assms power_eq_imp_eq_norm by blast
have "(w \<in> (\<lambda>w. w^n) ` ball z' d) = (w \<in> (\<lambda>w. w^n) ` ball z d)" for w
proof (cases "w=0")
case True with assms show ?thesis
by (simp add: image_def ball_def nz')
next
case False
have "z' \<noteq> 0" using \<open>z \<noteq> 0\<close> nz' by force
have [simp]: "(z*x / z')^n = x^n" if "x \<noteq> 0" for x
using z' that by (simp add: field_simps \<open>z \<noteq> 0\<close>)
have [simp]: "cmod (z - z * x / z') = cmod (z' - x)" if "x \<noteq> 0" for x
proof -
have "cmod (z - z * x / z') = cmod z * cmod (1 - x / z')"
by (metis (no_types) ab_semigroup_mult_class.mult_ac(1) divide_complex_def mult.right_neutral norm_mult right_diff_distrib')
also have "... = cmod z' * cmod (1 - x / z')"
by (simp add: nz')
also have "... = cmod (z' - x)"
by (simp add: \<open>z' \<noteq> 0\<close> diff_divide_eq_iff norm_divide)
finally show ?thesis .
qed
have [simp]: "(z'*x / z)^n = x^n" if "x \<noteq> 0" for x
using z' that by (simp add: field_simps \<open>z \<noteq> 0\<close>)
have [simp]: "cmod (z' - z' * x / z) = cmod (z - x)" if "x \<noteq> 0" for x
proof -
have "cmod (z * (1 - x * inverse z)) = cmod (z - x)"
by (metis \<open>z \<noteq> 0\<close> diff_divide_distrib divide_complex_def divide_self_if nonzero_eq_divide_eq semiring_normalization_rules(7))
then show ?thesis
by (metis (no_types) mult.assoc divide_complex_def mult.right_neutral norm_mult nz' right_diff_distrib')
qed
show ?thesis
unfolding image_def ball_def
apply safe
apply simp_all
apply (rule_tac x="z/z' * x" in exI)
using assms False apply (simp add: dist_norm)
apply (rule_tac x="z'/z * x" in exI)
using assms False apply (simp add: dist_norm)
done
qed
then show ?thesis by blast
qed
have ex_ball: "\<exists>B. (\<exists>z'. B = ball z' d \<and> z'^n = z^n) \<and> x \<in> B"
if "x \<noteq> 0" and eq: "x^n = w^n" and dzw: "dist z w < d" for x w
proof -
have "w \<noteq> 0" by (metis assms power_eq_0_iff that(1) that(2))
have [simp]: "cmod x = cmod w"
using assms power_eq_imp_eq_norm eq by blast
have [simp]: "cmod (x * z / w - x) = cmod (z - w)"
proof -
have "cmod (x * z / w - x) = cmod x * cmod (z / w - 1)"
by (metis (no_types) mult.right_neutral norm_mult right_diff_distrib' times_divide_eq_right)
also have "... = cmod w * cmod (z / w - 1)"
by simp
also have "... = cmod (z - w)"
by (simp add: \<open>w \<noteq> 0\<close> divide_diff_eq_iff nonzero_norm_divide)
finally show ?thesis .
qed
show ?thesis
apply (rule_tac x="ball (z / w * x) d" in exI)
using \<open>d > 0\<close> that
apply (simp add: ball_eq_ball_iff)
apply (simp add: \<open>z \<noteq> 0\<close> \<open>w \<noteq> 0\<close> field_simps)
apply (simp add: dist_norm)
done
qed
show ?thesis
proof (rule exI, intro conjI)
show "z ^ n \<in> (\<lambda>w. w ^ n) ` ball z d"
using \<open>d > 0\<close> by simp
show "open ((\<lambda>w. w ^ n) ` ball z d)"
by (rule invariance_of_domain [OF cont open_ball inj])
show "0 \<notin> (\<lambda>w. w ^ n) ` ball z d"
using \<open>z \<noteq> 0\<close> assms by (force simp: d_def)
show "\<exists>v. \<Union>v = - {0} \<inter> (\<lambda>z. z ^ n) -` (\<lambda>w. w ^ n) ` ball z d \<and>
(\<forall>u\<in>v. open u \<and> 0 \<notin> u) \<and>
disjoint v \<and>
(\<forall>u\<in>v. Ex (homeomorphism u ((\<lambda>w. w ^ n) ` ball z d) (\<lambda>z. z ^ n)))"
proof (rule exI, intro ballI conjI)
show "\<Union>{ball z' d |z'. z'^n = z^n} = - {0} \<inter> (\<lambda>z. z ^ n) -` (\<lambda>w. w ^ n) ` ball z d" (is "?l = ?r")
proof
show "?l \<subseteq> ?r"
apply auto
apply (simp add: assms d_def power_eq_imp_eq_norm that)
by (metis im_eq image_eqI mem_ball)
show "?r \<subseteq> ?l"
by auto (meson ex_ball)
qed
show "\<And>u. u \<in> {ball z' d |z'. z' ^ n = z ^ n} \<Longrightarrow> 0 \<notin> u"
by (force simp add: assms d_def power_eq_imp_eq_norm that)
show "disjoint {ball z' d |z'. z' ^ n = z ^ n}"
proof (clarsimp simp add: pairwise_def disjnt_iff)
fix \<xi> \<zeta> x
assume "\<xi>^n = z^n" "\<zeta>^n = z^n" "ball \<xi> d \<noteq> ball \<zeta> d"
and "dist \<xi> x < d" "dist \<zeta> x < d"
then have "dist \<xi> \<zeta> < d+d"
using dist_triangle_less_add by blast
then have "cmod (\<xi> - \<zeta>) < 2*d"
by (simp add: dist_norm)
also have "... \<le> e * cmod z"
using mult_right_mono \<open>0 < e\<close> that by (auto simp: d_def)
finally have "cmod (\<xi> - \<zeta>) < e * cmod z" .
with e have "\<xi> = \<zeta>"
by (metis \<open>\<xi>^n = z^n\<close> \<open>\<zeta>^n = z^n\<close> assms power_eq_imp_eq_norm)
then show "False"
using \<open>ball \<xi> d \<noteq> ball \<zeta> d\<close> by blast
qed
show "Ex (homeomorphism u ((\<lambda>w. w ^ n) ` ball z d) (\<lambda>z. z ^ n))"
if "u \<in> {ball z' d |z'. z' ^ n = z ^ n}" for u
proof (rule invariance_of_domain_homeomorphism [of "u" "\<lambda>z. z^n"])
show "open u"
using that by auto
show "continuous_on u (\<lambda>z. z ^ n)"
by (intro continuous_intros)
show "inj_on (\<lambda>z. z ^ n) u"
using that by (auto simp: iff_x_eq_y inj_on_def)
show "\<And>g. homeomorphism u ((\<lambda>z. z ^ n) ` u) (\<lambda>z. z ^ n) g \<Longrightarrow> Ex (homeomorphism u ((\<lambda>w. w ^ n) ` ball z d) (\<lambda>z. z ^ n))"
using im_eq that by clarify metis
qed auto
qed auto
qed
qed
show ?thesis
using assms
apply (simp add: covering_space_def zn1 zn2)
apply (subst zn2 [symmetric])
apply (simp add: openin_open_eq open_Compl)
apply (blast intro: zn3)
done
qed
corollary covering_space_square_punctured_plane:
"covering_space (- {0}) (\<lambda>z::complex. z^2) (- {0})"
by (simp add: covering_space_power_punctured_plane)
proposition covering_space_exp_punctured_plane:
"covering_space UNIV (\<lambda>z::complex. exp z) (- {0})"
proof (simp add: covering_space_def, intro conjI ballI)
show "continuous_on UNIV (\<lambda>z::complex. exp z)"
by (rule continuous_on_exp [OF continuous_on_id])
show "range exp = - {0::complex}"
by auto (metis exp_Ln range_eqI)
show "\<exists>T. z \<in> T \<and> openin (top_of_set (- {0})) T \<and>
(\<exists>v. \<Union>v = exp -` T \<and> (\<forall>u\<in>v. open u) \<and> disjoint v \<and>
(\<forall>u\<in>v. \<exists>q. homeomorphism u T exp q))"
if "z \<in> - {0::complex}" for z
proof -
have "z \<noteq> 0"
using that by auto
have inj_exp: "inj_on exp (ball (Ln z) 1)"
apply (rule inj_on_subset [OF inj_on_exp_pi [of "Ln z"]])
using pi_ge_two by (simp add: ball_subset_ball_iff)
define \<V> where "\<V> \<equiv> range (\<lambda>n. (\<lambda>x. x + of_real (2 * of_int n * pi) * \<i>) ` (ball(Ln z) 1))"
show ?thesis
proof (intro exI conjI)
show "z \<in> exp ` (ball(Ln z) 1)"
by (metis \<open>z \<noteq> 0\<close> centre_in_ball exp_Ln rev_image_eqI zero_less_one)
have "open (- {0::complex})"
by blast
moreover have "inj_on exp (ball (Ln z) 1)"
apply (rule inj_on_subset [OF inj_on_exp_pi [of "Ln z"]])
using pi_ge_two by (simp add: ball_subset_ball_iff)
ultimately show "openin (top_of_set (- {0})) (exp ` ball (Ln z) 1)"
by (auto simp: openin_open_eq invariance_of_domain continuous_on_exp [OF continuous_on_id])
show "\<Union>\<V> = exp -` exp ` ball (Ln z) 1"
by (force simp: \<V>_def Complex_Transcendental.exp_eq image_iff)
show "\<forall>V\<in>\<V>. open V"
by (auto simp: \<V>_def inj_on_def continuous_intros invariance_of_domain)
have xy: "2 \<le> cmod (2 * of_int x * of_real pi * \<i> - 2 * of_int y * of_real pi * \<i>)"
if "x < y" for x y
proof -
have "1 \<le> abs (x - y)"
using that by linarith
then have "1 \<le> cmod (of_int x - of_int y) * 1"
by (metis mult.right_neutral norm_of_int of_int_1_le_iff of_int_abs of_int_diff)
also have "... \<le> cmod (of_int x - of_int y) * of_real pi"
apply (rule mult_left_mono)
using pi_ge_two by auto
also have "... \<le> cmod ((of_int x - of_int y) * of_real pi * \<i>)"
by (simp add: norm_mult)
also have "... \<le> cmod (of_int x * of_real pi * \<i> - of_int y * of_real pi * \<i>)"
by (simp add: algebra_simps)
finally have "1 \<le> cmod (of_int x * of_real pi * \<i> - of_int y * of_real pi * \<i>)" .
then have "2 * 1 \<le> cmod (2 * (of_int x * of_real pi * \<i> - of_int y * of_real pi * \<i>))"
by (metis mult_le_cancel_left_pos norm_mult_numeral1 zero_less_numeral)
then show ?thesis
by (simp add: algebra_simps)
qed
show "disjoint \<V>"
apply (clarsimp simp add: \<V>_def pairwise_def disjnt_def add.commute [of _ "x*y" for x y]
- image_add_ball ball_eq_ball_iff)
+ ball_eq_ball_iff)
apply (rule disjoint_ballI)
apply (auto simp: dist_norm neq_iff)
by (metis norm_minus_commute xy)+
show "\<forall>u\<in>\<V>. \<exists>q. homeomorphism u (exp ` ball (Ln z) 1) exp q"
proof
fix u
assume "u \<in> \<V>"
then obtain n where n: "u = (\<lambda>x. x + of_real (2 * of_int n * pi) * \<i>) ` (ball(Ln z) 1)"
by (auto simp: \<V>_def)
have "compact (cball (Ln z) 1)"
by simp
moreover have "continuous_on (cball (Ln z) 1) exp"
by (rule continuous_on_exp [OF continuous_on_id])
moreover have "inj_on exp (cball (Ln z) 1)"
apply (rule inj_on_subset [OF inj_on_exp_pi [of "Ln z"]])
using pi_ge_two by (simp add: cball_subset_ball_iff)
ultimately obtain \<gamma> where hom: "homeomorphism (cball (Ln z) 1) (exp ` cball (Ln z) 1) exp \<gamma>"
using homeomorphism_compact by blast
have eq1: "exp ` u = exp ` ball (Ln z) 1"
unfolding n
apply (auto simp: algebra_simps)
apply (rename_tac w)
apply (rule_tac x = "w + \<i> * (of_int n * (of_real pi * 2))" in image_eqI)
apply (auto simp: image_iff)
done
have \<gamma>exp: "\<gamma> (exp x) + 2 * of_int n * of_real pi * \<i> = x" if "x \<in> u" for x
proof -
have "exp x = exp (x - 2 * of_int n * of_real pi * \<i>)"
by (simp add: exp_eq)
then have "\<gamma> (exp x) = \<gamma> (exp (x - 2 * of_int n * of_real pi * \<i>))"
by simp
also have "... = x - 2 * of_int n * of_real pi * \<i>"
apply (rule homeomorphism_apply1 [OF hom])
using \<open>x \<in> u\<close> by (auto simp: n)
finally show ?thesis
by simp
qed
have exp2n: "exp (\<gamma> (exp x) + 2 * of_int n * complex_of_real pi * \<i>) = exp x"
if "dist (Ln z) x < 1" for x
using that by (auto simp: exp_eq homeomorphism_apply1 [OF hom])
have cont: "continuous_on (exp ` ball (Ln z) 1) (\<lambda>x. \<gamma> x + 2 * of_int n * complex_of_real pi * \<i>)"
apply (intro continuous_intros)
apply (rule continuous_on_subset [OF homeomorphism_cont2 [OF hom]])
apply (force simp:)
done
show "\<exists>q. homeomorphism u (exp ` ball (Ln z) 1) exp q"
apply (rule_tac x="(\<lambda>x. x + of_real(2 * n * pi) * \<i>) \<circ> \<gamma>" in exI)
unfolding homeomorphism_def
apply (intro conjI ballI eq1 continuous_on_exp [OF continuous_on_id])
apply (auto simp: \<gamma>exp exp2n cont n)
apply (simp add: homeomorphism_apply1 [OF hom])
using hom homeomorphism_apply1 apply (force simp: image_iff)
done
qed
qed
qed
qed
subsection\<open>Hence the Borsukian results about mappings into circles\<close>(*FIX ME title *)
lemma inessential_eq_continuous_logarithm:
fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
shows "(\<exists>a. homotopic_with_canon (\<lambda>h. True) S (-{0}) f (\<lambda>t. a)) \<longleftrightarrow>
(\<exists>g. continuous_on S g \<and> (\<forall>x \<in> S. f x = exp(g x)))"
(is "?lhs \<longleftrightarrow> ?rhs")
proof
assume ?lhs thus ?rhs
by (metis covering_space_lift_inessential_function covering_space_exp_punctured_plane)
next
assume ?rhs
then obtain g where contg: "continuous_on S g" and f: "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"
by metis
obtain a where "homotopic_with_canon (\<lambda>h. True) S (- {of_real 0}) (exp \<circ> g) (\<lambda>x. a)"
proof (rule nullhomotopic_through_contractible [OF contg subset_UNIV _ _ contractible_UNIV])
show "continuous_on (UNIV::complex set) exp"
by (intro continuous_intros)
show "range exp \<subseteq> - {0}"
by auto
qed force
then have "homotopic_with_canon (\<lambda>h. True) S (- {0}) f (\<lambda>t. a)"
using f homotopic_with_eq by fastforce
then show ?lhs ..
qed
corollary inessential_imp_continuous_logarithm_circle:
fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
assumes "homotopic_with_canon (\<lambda>h. True) S (sphere 0 1) f (\<lambda>t. a)"
obtains g where "continuous_on S g" and "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"
proof -
have "homotopic_with_canon (\<lambda>h. True) S (- {0}) f (\<lambda>t. a)"
using assms homotopic_with_subset_right by fastforce
then show ?thesis
by (metis inessential_eq_continuous_logarithm that)
qed
lemma inessential_eq_continuous_logarithm_circle:
fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
shows "(\<exists>a. homotopic_with_canon (\<lambda>h. True) S (sphere 0 1) f (\<lambda>t. a)) \<longleftrightarrow>
(\<exists>g. continuous_on S g \<and> (\<forall>x \<in> S. f x = exp(\<i> * of_real(g x))))"
(is "?lhs \<longleftrightarrow> ?rhs")
proof
assume L: ?lhs
then obtain g where contg: "continuous_on S g" and g: "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"
using inessential_imp_continuous_logarithm_circle by blast
have "f ` S \<subseteq> sphere 0 1"
by (metis L homotopic_with_imp_subset1)
then have "\<And>x. x \<in> S \<Longrightarrow> Re (g x) = 0"
using g by auto
then show ?rhs
apply (rule_tac x="Im \<circ> g" in exI)
apply (intro conjI contg continuous_intros)
apply (auto simp: Euler g)
done
next
assume ?rhs
then obtain g where contg: "continuous_on S g" and g: "\<And>x. x \<in> S \<Longrightarrow> f x = exp(\<i>* of_real(g x))"
by metis
obtain a where "homotopic_with_canon (\<lambda>h. True) S (sphere 0 1) ((exp \<circ> (\<lambda>z. \<i>*z)) \<circ> (of_real \<circ> g)) (\<lambda>x. a)"
proof (rule nullhomotopic_through_contractible)
show "continuous_on S (complex_of_real \<circ> g)"
by (intro conjI contg continuous_intros)
show "(complex_of_real \<circ> g) ` S \<subseteq> \<real>"
by auto
show "continuous_on \<real> (exp \<circ> (*)\<i>)"
by (intro continuous_intros)
show "(exp \<circ> (*)\<i>) ` \<real> \<subseteq> sphere 0 1"
by (auto simp: complex_is_Real_iff)
qed (auto simp: convex_Reals convex_imp_contractible)
moreover have "\<And>x. x \<in> S \<Longrightarrow> (exp \<circ> (*)\<i> \<circ> (complex_of_real \<circ> g)) x = f x"
by (simp add: g)
ultimately have "homotopic_with_canon (\<lambda>h. True) S (sphere 0 1) f (\<lambda>t. a)"
using homotopic_with_eq by force
then show ?lhs ..
qed
proposition homotopic_with_sphere_times:
fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
assumes hom: "homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) f g" and conth: "continuous_on S h"
and hin: "\<And>x. x \<in> S \<Longrightarrow> h x \<in> sphere 0 1"
shows "homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x * h x) (\<lambda>x. g x * h x)"
proof -
obtain k where contk: "continuous_on ({0..1::real} \<times> S) k"
and kim: "k ` ({0..1} \<times> S) \<subseteq> sphere 0 1"
and k0: "\<And>x. k(0, x) = f x"
and k1: "\<And>x. k(1, x) = g x"
using hom by (auto simp: homotopic_with_def)
show ?thesis
apply (simp add: homotopic_with)
apply (rule_tac x="\<lambda>z. k z*(h \<circ> snd)z" in exI)
apply (intro conjI contk continuous_intros)
apply (simp add: conth)
using kim hin apply (force simp: norm_mult k0 k1)+
done
qed
proposition homotopic_circlemaps_divide:
fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
shows "homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) f g \<longleftrightarrow>
continuous_on S f \<and> f ` S \<subseteq> sphere 0 1 \<and>
continuous_on S g \<and> g ` S \<subseteq> sphere 0 1 \<and>
(\<exists>c. homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x / g x) (\<lambda>x. c))"
proof -
have "homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x / g x) (\<lambda>x. 1)"
if "homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x / g x) (\<lambda>x. c)" for c
proof -
have "S = {} \<or> path_component (sphere 0 1) 1 c"
using homotopic_with_imp_subset2 [OF that] path_connected_sphere [of "0::complex" 1]
by (auto simp: path_connected_component)
then have "homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. 1) (\<lambda>x. c)"
by (simp add: homotopic_constant_maps)
then show ?thesis
using homotopic_with_symD homotopic_with_trans that by blast
qed
then have *: "(\<exists>c. homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x / g x) (\<lambda>x. c)) \<longleftrightarrow>
homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x / g x) (\<lambda>x. 1)"
by auto
have "homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) f g \<longleftrightarrow>
continuous_on S f \<and> f ` S \<subseteq> sphere 0 1 \<and>
continuous_on S g \<and> g ` S \<subseteq> sphere 0 1 \<and>
homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x / g x) (\<lambda>x. 1)"
(is "?lhs \<longleftrightarrow> ?rhs")
proof
assume L: ?lhs
have geq1 [simp]: "\<And>x. x \<in> S \<Longrightarrow> cmod (g x) = 1"
using homotopic_with_imp_subset2 [OF L]
by (simp add: image_subset_iff)
have cont: "continuous_on S (inverse \<circ> g)"
apply (rule continuous_intros)
using homotopic_with_imp_continuous [OF L] apply blast
apply (rule continuous_on_subset [of "sphere 0 1", OF continuous_on_inverse])
apply (auto)
done
have "homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x / g x) (\<lambda>x. 1)"
using homotopic_with_sphere_times [OF L cont]
apply (rule homotopic_with_eq)
apply (auto simp: division_ring_class.divide_inverse norm_inverse)
by (metis geq1 norm_zero right_inverse zero_neq_one)
with L show ?rhs
by (auto simp: homotopic_with_imp_continuous dest: homotopic_with_imp_subset1 homotopic_with_imp_subset2)
next
assume ?rhs then show ?lhs
by (force simp: elim: homotopic_with_eq dest: homotopic_with_sphere_times [where h=g])
qed
then show ?thesis
by (simp add: *)
qed
subsection\<open>Upper and lower hemicontinuous functions\<close>
text\<open>And relation in the case of preimage map to open and closed maps, and fact that upper and lower
hemicontinuity together imply continuity in the sense of the Hausdorff metric (at points where the
function gives a bounded and nonempty set).\<close>
text\<open>Many similar proofs below.\<close>
lemma upper_hemicontinuous:
assumes "\<And>x. x \<in> S \<Longrightarrow> f x \<subseteq> T"
shows "((\<forall>U. openin (top_of_set T) U
\<longrightarrow> openin (top_of_set S) {x \<in> S. f x \<subseteq> U}) \<longleftrightarrow>
(\<forall>U. closedin (top_of_set T) U
\<longrightarrow> closedin (top_of_set S) {x \<in> S. f x \<inter> U \<noteq> {}}))"
(is "?lhs = ?rhs")
proof (intro iffI allI impI)
fix U
assume * [rule_format]: ?lhs and "closedin (top_of_set T) U"
then have "openin (top_of_set T) (T - U)"
by (simp add: openin_diff)
then have "openin (top_of_set S) {x \<in> S. f x \<subseteq> T - U}"
using * [of "T-U"] by blast
moreover have "S - {x \<in> S. f x \<subseteq> T - U} = {x \<in> S. f x \<inter> U \<noteq> {}}"
using assms by blast
ultimately show "closedin (top_of_set S) {x \<in> S. f x \<inter> U \<noteq> {}}"
by (simp add: openin_closedin_eq)
next
fix U
assume * [rule_format]: ?rhs and "openin (top_of_set T) U"
then have "closedin (top_of_set T) (T - U)"
by (simp add: closedin_diff)
then have "closedin (top_of_set S) {x \<in> S. f x \<inter> (T - U) \<noteq> {}}"
using * [of "T-U"] by blast
moreover have "{x \<in> S. f x \<inter> (T - U) \<noteq> {}} = S - {x \<in> S. f x \<subseteq> U}"
using assms by auto
ultimately show "openin (top_of_set S) {x \<in> S. f x \<subseteq> U}"
by (simp add: openin_closedin_eq)
qed
lemma lower_hemicontinuous:
assumes "\<And>x. x \<in> S \<Longrightarrow> f x \<subseteq> T"
shows "((\<forall>U. closedin (top_of_set T) U
\<longrightarrow> closedin (top_of_set S) {x \<in> S. f x \<subseteq> U}) \<longleftrightarrow>
(\<forall>U. openin (top_of_set T) U
\<longrightarrow> openin (top_of_set S) {x \<in> S. f x \<inter> U \<noteq> {}}))"
(is "?lhs = ?rhs")
proof (intro iffI allI impI)
fix U
assume * [rule_format]: ?lhs and "openin (top_of_set T) U"
then have "closedin (top_of_set T) (T - U)"
by (simp add: closedin_diff)
then have "closedin (top_of_set S) {x \<in> S. f x \<subseteq> T-U}"
using * [of "T-U"] by blast
moreover have "{x \<in> S. f x \<subseteq> T-U} = S - {x \<in> S. f x \<inter> U \<noteq> {}}"
using assms by auto
ultimately show "openin (top_of_set S) {x \<in> S. f x \<inter> U \<noteq> {}}"
by (simp add: openin_closedin_eq)
next
fix U
assume * [rule_format]: ?rhs and "closedin (top_of_set T) U"
then have "openin (top_of_set T) (T - U)"
by (simp add: openin_diff)
then have "openin (top_of_set S) {x \<in> S. f x \<inter> (T - U) \<noteq> {}}"
using * [of "T-U"] by blast
moreover have "S - {x \<in> S. f x \<inter> (T - U) \<noteq> {}} = {x \<in> S. f x \<subseteq> U}"
using assms by blast
ultimately show "closedin (top_of_set S) {x \<in> S. f x \<subseteq> U}"
by (simp add: openin_closedin_eq)
qed
lemma open_map_iff_lower_hemicontinuous_preimage:
assumes "f ` S \<subseteq> T"
shows "((\<forall>U. openin (top_of_set S) U
\<longrightarrow> openin (top_of_set T) (f ` U)) \<longleftrightarrow>
(\<forall>U. closedin (top_of_set S) U
\<longrightarrow> closedin (top_of_set T) {y \<in> T. {x. x \<in> S \<and> f x = y} \<subseteq> U}))"
(is "?lhs = ?rhs")
proof (intro iffI allI impI)
fix U
assume * [rule_format]: ?lhs and "closedin (top_of_set S) U"
then have "openin (top_of_set S) (S - U)"
by (simp add: openin_diff)
then have "openin (top_of_set T) (f ` (S - U))"
using * [of "S-U"] by blast
moreover have "T - (f ` (S - U)) = {y \<in> T. {x \<in> S. f x = y} \<subseteq> U}"
using assms by blast
ultimately show "closedin (top_of_set T) {y \<in> T. {x \<in> S. f x = y} \<subseteq> U}"
by (simp add: openin_closedin_eq)
next
fix U
assume * [rule_format]: ?rhs and opeSU: "openin (top_of_set S) U"
then have "closedin (top_of_set S) (S - U)"
by (simp add: closedin_diff)
then have "closedin (top_of_set T) {y \<in> T. {x \<in> S. f x = y} \<subseteq> S - U}"
using * [of "S-U"] by blast
moreover have "{y \<in> T. {x \<in> S. f x = y} \<subseteq> S - U} = T - (f ` U)"
using assms openin_imp_subset [OF opeSU] by auto
ultimately show "openin (top_of_set T) (f ` U)"
using assms openin_imp_subset [OF opeSU] by (force simp: openin_closedin_eq)
qed
lemma closed_map_iff_upper_hemicontinuous_preimage:
assumes "f ` S \<subseteq> T"
shows "((\<forall>U. closedin (top_of_set S) U
\<longrightarrow> closedin (top_of_set T) (f ` U)) \<longleftrightarrow>
(\<forall>U. openin (top_of_set S) U
\<longrightarrow> openin (top_of_set T) {y \<in> T. {x. x \<in> S \<and> f x = y} \<subseteq> U}))"
(is "?lhs = ?rhs")
proof (intro iffI allI impI)
fix U
assume * [rule_format]: ?lhs and opeSU: "openin (top_of_set S) U"
then have "closedin (top_of_set S) (S - U)"
by (simp add: closedin_diff)
then have "closedin (top_of_set T) (f ` (S - U))"
using * [of "S-U"] by blast
moreover have "f ` (S - U) = T - {y \<in> T. {x. x \<in> S \<and> f x = y} \<subseteq> U}"
using assms openin_imp_subset [OF opeSU] by auto
ultimately show "openin (top_of_set T) {y \<in> T. {x. x \<in> S \<and> f x = y} \<subseteq> U}"
using assms openin_imp_subset [OF opeSU] by (force simp: openin_closedin_eq)
next
fix U
assume * [rule_format]: ?rhs and cloSU: "closedin (top_of_set S) U"
then have "openin (top_of_set S) (S - U)"
by (simp add: openin_diff)
then have "openin (top_of_set T) {y \<in> T. {x \<in> S. f x = y} \<subseteq> S - U}"
using * [of "S-U"] by blast
moreover have "(f ` U) = T - {y \<in> T. {x \<in> S. f x = y} \<subseteq> S - U}"
using assms closedin_imp_subset [OF cloSU] by auto
ultimately show "closedin (top_of_set T) (f ` U)"
by (simp add: openin_closedin_eq)
qed
proposition upper_lower_hemicontinuous_explicit:
fixes T :: "('b::{real_normed_vector,heine_borel}) set"
assumes fST: "\<And>x. x \<in> S \<Longrightarrow> f x \<subseteq> T"
and ope: "\<And>U. openin (top_of_set T) U
\<Longrightarrow> openin (top_of_set S) {x \<in> S. f x \<subseteq> U}"
and clo: "\<And>U. closedin (top_of_set T) U
\<Longrightarrow> closedin (top_of_set S) {x \<in> S. f x \<subseteq> U}"
and "x \<in> S" "0 < e" and bofx: "bounded(f x)" and fx_ne: "f x \<noteq> {}"
obtains d where "0 < d"
"\<And>x'. \<lbrakk>x' \<in> S; dist x x' < d\<rbrakk>
\<Longrightarrow> (\<forall>y \<in> f x. \<exists>y'. y' \<in> f x' \<and> dist y y' < e) \<and>
(\<forall>y' \<in> f x'. \<exists>y. y \<in> f x \<and> dist y' y < e)"
proof -
have "openin (top_of_set T) (T \<inter> (\<Union>a\<in>f x. \<Union>b\<in>ball 0 e. {a + b}))"
by (auto simp: open_sums openin_open_Int)
with ope have "openin (top_of_set S)
{u \<in> S. f u \<subseteq> T \<inter> (\<Union>a\<in>f x. \<Union>b\<in>ball 0 e. {a + b})}" by blast
with \<open>0 < e\<close> \<open>x \<in> S\<close> obtain d1 where "d1 > 0" and
d1: "\<And>x'. \<lbrakk>x' \<in> S; dist x' x < d1\<rbrakk> \<Longrightarrow> f x' \<subseteq> T \<and> f x' \<subseteq> (\<Union>a \<in> f x. \<Union>b \<in> ball 0 e. {a + b})"
by (force simp: openin_euclidean_subtopology_iff dest: fST)
have oo: "\<And>U. openin (top_of_set T) U \<Longrightarrow>
openin (top_of_set S) {x \<in> S. f x \<inter> U \<noteq> {}}"
apply (rule lower_hemicontinuous [THEN iffD1, rule_format])
using fST clo by auto
have "compact (closure(f x))"
by (simp add: bofx)
moreover have "closure(f x) \<subseteq> (\<Union>a \<in> f x. ball a (e/2))"
using \<open>0 < e\<close> by (force simp: closure_approachable simp del: divide_const_simps)
ultimately obtain C where "C \<subseteq> f x" "finite C" "closure(f x) \<subseteq> (\<Union>a \<in> C. ball a (e/2))"
apply (rule compactE, force)
by (metis finite_subset_image)
then have fx_cover: "f x \<subseteq> (\<Union>a \<in> C. ball a (e/2))"
by (meson closure_subset order_trans)
with fx_ne have "C \<noteq> {}"
by blast
have xin: "x \<in> (\<Inter>a \<in> C. {x \<in> S. f x \<inter> T \<inter> ball a (e/2) \<noteq> {}})"
using \<open>x \<in> S\<close> \<open>0 < e\<close> fST \<open>C \<subseteq> f x\<close> by force
have "openin (top_of_set S) {x \<in> S. f x \<inter> (T \<inter> ball a (e/2)) \<noteq> {}}" for a
by (simp add: openin_open_Int oo)
then have "openin (top_of_set S) (\<Inter>a \<in> C. {x \<in> S. f x \<inter> T \<inter> ball a (e/2) \<noteq> {}})"
by (simp add: Int_assoc openin_INT2 [OF \<open>finite C\<close> \<open>C \<noteq> {}\<close>])
with xin obtain d2 where "d2>0"
and d2: "\<And>u v. \<lbrakk>u \<in> S; dist u x < d2; v \<in> C\<rbrakk> \<Longrightarrow> f u \<inter> T \<inter> ball v (e/2) \<noteq> {}"
unfolding openin_euclidean_subtopology_iff using xin by fastforce
show ?thesis
proof (intro that conjI ballI)
show "0 < min d1 d2"
using \<open>0 < d1\<close> \<open>0 < d2\<close> by linarith
next
fix x' y
assume "x' \<in> S" "dist x x' < min d1 d2" "y \<in> f x"
then have dd2: "dist x' x < d2"
by (auto simp: dist_commute)
obtain a where "a \<in> C" "y \<in> ball a (e/2)"
using fx_cover \<open>y \<in> f x\<close> by auto
then show "\<exists>y'. y' \<in> f x' \<and> dist y y' < e"
using d2 [OF \<open>x' \<in> S\<close> dd2] dist_triangle_half_r by fastforce
next
fix x' y'
assume "x' \<in> S" "dist x x' < min d1 d2" "y' \<in> f x'"
then have "dist x' x < d1"
by (auto simp: dist_commute)
then have "y' \<in> (\<Union>a\<in>f x. \<Union>b\<in>ball 0 e. {a + b})"
using d1 [OF \<open>x' \<in> S\<close>] \<open>y' \<in> f x'\<close> by force
then show "\<exists>y. y \<in> f x \<and> dist y' y < e"
apply auto
by (metis add_diff_cancel_left' dist_norm)
qed
qed
subsection\<open>Complex logs exist on various "well-behaved" sets\<close>
lemma continuous_logarithm_on_contractible:
fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
assumes "continuous_on S f" "contractible S" "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"
obtains g where "continuous_on S g" "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"
proof -
obtain c where hom: "homotopic_with_canon (\<lambda>h. True) S (-{0}) f (\<lambda>x. c)"
using nullhomotopic_from_contractible assms
by (metis imageE subset_Compl_singleton)
then show ?thesis
by (metis inessential_eq_continuous_logarithm that)
qed
lemma continuous_logarithm_on_simply_connected:
fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
assumes contf: "continuous_on S f" and S: "simply_connected S" "locally path_connected S"
and f: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"
obtains g where "continuous_on S g" "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"
using covering_space_lift [OF covering_space_exp_punctured_plane S contf]
by (metis (full_types) f imageE subset_Compl_singleton)
lemma continuous_logarithm_on_cball:
fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
assumes "continuous_on (cball a r) f" and "\<And>z. z \<in> cball a r \<Longrightarrow> f z \<noteq> 0"
obtains h where "continuous_on (cball a r) h" "\<And>z. z \<in> cball a r \<Longrightarrow> f z = exp(h z)"
using assms continuous_logarithm_on_contractible convex_imp_contractible by blast
lemma continuous_logarithm_on_ball:
fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
assumes "continuous_on (ball a r) f" and "\<And>z. z \<in> ball a r \<Longrightarrow> f z \<noteq> 0"
obtains h where "continuous_on (ball a r) h" "\<And>z. z \<in> ball a r \<Longrightarrow> f z = exp(h z)"
using assms continuous_logarithm_on_contractible convex_imp_contractible by blast
lemma continuous_sqrt_on_contractible:
fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
assumes "continuous_on S f" "contractible S"
and "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"
obtains g where "continuous_on S g" "\<And>x. x \<in> S \<Longrightarrow> f x = (g x) ^ 2"
proof -
obtain g where contg: "continuous_on S g" and feq: "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"
using continuous_logarithm_on_contractible [OF assms] by blast
show ?thesis
proof
show "continuous_on S (\<lambda>z. exp (g z / 2))"
by (rule continuous_on_compose2 [of UNIV exp]; intro continuous_intros contg subset_UNIV) auto
show "\<And>x. x \<in> S \<Longrightarrow> f x = (exp (g x / 2))\<^sup>2"
by (metis exp_double feq nonzero_mult_div_cancel_left times_divide_eq_right zero_neq_numeral)
qed
qed
lemma continuous_sqrt_on_simply_connected:
fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
assumes contf: "continuous_on S f" and S: "simply_connected S" "locally path_connected S"
and f: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"
obtains g where "continuous_on S g" "\<And>x. x \<in> S \<Longrightarrow> f x = (g x) ^ 2"
proof -
obtain g where contg: "continuous_on S g" and feq: "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"
using continuous_logarithm_on_simply_connected [OF assms] by blast
show ?thesis
proof
show "continuous_on S (\<lambda>z. exp (g z / 2))"
by (rule continuous_on_compose2 [of UNIV exp]; intro continuous_intros contg subset_UNIV) auto
show "\<And>x. x \<in> S \<Longrightarrow> f x = (exp (g x / 2))\<^sup>2"
by (metis exp_double feq nonzero_mult_div_cancel_left times_divide_eq_right zero_neq_numeral)
qed
qed
subsection\<open>Another simple case where sphere maps are nullhomotopic\<close>
lemma inessential_spheremap_2_aux:
fixes f :: "'a::euclidean_space \<Rightarrow> complex"
assumes 2: "2 < DIM('a)" and contf: "continuous_on (sphere a r) f"
and fim: "f `(sphere a r) \<subseteq> (sphere 0 1)"
obtains c where "homotopic_with_canon (\<lambda>z. True) (sphere a r) (sphere 0 1) f (\<lambda>x. c)"
proof -
obtain g where contg: "continuous_on (sphere a r) g"
and feq: "\<And>x. x \<in> sphere a r \<Longrightarrow> f x = exp(g x)"
proof (rule continuous_logarithm_on_simply_connected [OF contf])
show "simply_connected (sphere a r)"
using 2 by (simp add: simply_connected_sphere_eq)
show "locally path_connected (sphere a r)"
by (simp add: locally_path_connected_sphere)
show "\<And>z. z \<in> sphere a r \<Longrightarrow> f z \<noteq> 0"
using fim by force
qed auto
have "\<exists>g. continuous_on (sphere a r) g \<and> (\<forall>x\<in>sphere a r. f x = exp (\<i> * complex_of_real (g x)))"
proof (intro exI conjI)
show "continuous_on (sphere a r) (Im \<circ> g)"
by (intro contg continuous_intros continuous_on_compose)
show "\<forall>x\<in>sphere a r. f x = exp (\<i> * complex_of_real ((Im \<circ> g) x))"
using exp_eq_polar feq fim norm_exp_eq_Re by auto
qed
with inessential_eq_continuous_logarithm_circle that show ?thesis
by metis
qed
lemma inessential_spheremap_2:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes a2: "2 < DIM('a)" and b2: "DIM('b) = 2"
and contf: "continuous_on (sphere a r) f" and fim: "f `(sphere a r) \<subseteq> (sphere b s)"
obtains c where "homotopic_with_canon (\<lambda>z. True) (sphere a r) (sphere b s) f (\<lambda>x. c)"
proof (cases "s \<le> 0")
case True
then show ?thesis
using contf contractible_sphere fim nullhomotopic_into_contractible that by blast
next
case False
then have "sphere b s homeomorphic sphere (0::complex) 1"
using assms by (simp add: homeomorphic_spheres_gen)
then obtain h k where hk: "homeomorphism (sphere b s) (sphere (0::complex) 1) h k"
by (auto simp: homeomorphic_def)
then have conth: "continuous_on (sphere b s) h"
and contk: "continuous_on (sphere 0 1) k"
and him: "h ` sphere b s \<subseteq> sphere 0 1"
and kim: "k ` sphere 0 1 \<subseteq> sphere b s"
by (simp_all add: homeomorphism_def)
obtain c where "homotopic_with_canon (\<lambda>z. True) (sphere a r) (sphere 0 1) (h \<circ> f) (\<lambda>x. c)"
proof (rule inessential_spheremap_2_aux [OF a2])
show "continuous_on (sphere a r) (h \<circ> f)"
by (meson continuous_on_compose [OF contf] conth continuous_on_subset fim)
show "(h \<circ> f) ` sphere a r \<subseteq> sphere 0 1"
using fim him by force
qed auto
then have "homotopic_with_canon (\<lambda>f. True) (sphere a r) (sphere b s) (k \<circ> (h \<circ> f)) (k \<circ> (\<lambda>x. c))"
by (rule homotopic_compose_continuous_left [OF _ contk kim])
then have "homotopic_with_canon (\<lambda>z. True) (sphere a r) (sphere b s) f (\<lambda>x. k c)"
apply (rule homotopic_with_eq, auto)
by (metis fim hk homeomorphism_def image_subset_iff mem_sphere)
then show ?thesis
by (metis that)
qed
subsection\<open>Holomorphic logarithms and square roots\<close>
lemma contractible_imp_holomorphic_log:
assumes holf: "f holomorphic_on S"
and S: "contractible S"
and fnz: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"
obtains g where "g holomorphic_on S" "\<And>z. z \<in> S \<Longrightarrow> f z = exp(g z)"
proof -
have contf: "continuous_on S f"
by (simp add: holf holomorphic_on_imp_continuous_on)
obtain g where contg: "continuous_on S g" and feq: "\<And>x. x \<in> S \<Longrightarrow> f x = exp (g x)"
by (metis continuous_logarithm_on_contractible [OF contf S fnz])
have "g field_differentiable at z within S" if "f field_differentiable at z within S" "z \<in> S" for z
proof -
obtain f' where f': "((\<lambda>y. (f y - f z) / (y - z)) \<longlongrightarrow> f') (at z within S)"
using \<open>f field_differentiable at z within S\<close> by (auto simp: field_differentiable_def has_field_derivative_iff)
then have ee: "((\<lambda>x. (exp(g x) - exp(g z)) / (x - z)) \<longlongrightarrow> f') (at z within S)"
by (simp add: feq \<open>z \<in> S\<close> Lim_transform_within [OF _ zero_less_one])
have "(((\<lambda>y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \<circ> g) \<longlongrightarrow> exp (g z))
(at z within S)"
proof (rule tendsto_compose_at)
show "(g \<longlongrightarrow> g z) (at z within S)"
using contg continuous_on \<open>z \<in> S\<close> by blast
show "(\<lambda>y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \<midarrow>g z\<rightarrow> exp (g z)"
apply (subst Lim_at_zero)
apply (simp add: DERIV_D cong: if_cong Lim_cong_within)
done
qed auto
then have dd: "((\<lambda>x. if g x = g z then exp(g z) else (exp(g x) - exp(g z)) / (g x - g z)) \<longlongrightarrow> exp(g z)) (at z within S)"
by (simp add: o_def)
have "continuous (at z within S) g"
using contg continuous_on_eq_continuous_within \<open>z \<in> S\<close> by blast
then have "(\<forall>\<^sub>F x in at z within S. dist (g x) (g z) < 2*pi)"
by (simp add: continuous_within tendsto_iff)
then have "\<forall>\<^sub>F x in at z within S. exp (g x) = exp (g z) \<longrightarrow> g x \<noteq> g z \<longrightarrow> x = z"
apply (rule eventually_mono)
apply (auto simp: exp_eq dist_norm norm_mult)
done
then have "((\<lambda>y. (g y - g z) / (y - z)) \<longlongrightarrow> f' / exp (g z)) (at z within S)"
by (auto intro!: Lim_transform_eventually [OF tendsto_divide [OF ee dd]])
then show ?thesis
by (auto simp: field_differentiable_def has_field_derivative_iff)
qed
then have "g holomorphic_on S"
using holf holomorphic_on_def by auto
then show ?thesis
using feq that by auto
qed
(*Identical proofs*)
lemma simply_connected_imp_holomorphic_log:
assumes holf: "f holomorphic_on S"
and S: "simply_connected S" "locally path_connected S"
and fnz: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"
obtains g where "g holomorphic_on S" "\<And>z. z \<in> S \<Longrightarrow> f z = exp(g z)"
proof -
have contf: "continuous_on S f"
by (simp add: holf holomorphic_on_imp_continuous_on)
obtain g where contg: "continuous_on S g" and feq: "\<And>x. x \<in> S \<Longrightarrow> f x = exp (g x)"
by (metis continuous_logarithm_on_simply_connected [OF contf S fnz])
have "g field_differentiable at z within S" if "f field_differentiable at z within S" "z \<in> S" for z
proof -
obtain f' where f': "((\<lambda>y. (f y - f z) / (y - z)) \<longlongrightarrow> f') (at z within S)"
using \<open>f field_differentiable at z within S\<close> by (auto simp: field_differentiable_def has_field_derivative_iff)
then have ee: "((\<lambda>x. (exp(g x) - exp(g z)) / (x - z)) \<longlongrightarrow> f') (at z within S)"
by (simp add: feq \<open>z \<in> S\<close> Lim_transform_within [OF _ zero_less_one])
have "(((\<lambda>y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \<circ> g) \<longlongrightarrow> exp (g z))
(at z within S)"
proof (rule tendsto_compose_at)
show "(g \<longlongrightarrow> g z) (at z within S)"
using contg continuous_on \<open>z \<in> S\<close> by blast
show "(\<lambda>y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \<midarrow>g z\<rightarrow> exp (g z)"
apply (subst Lim_at_zero)
apply (simp add: DERIV_D cong: if_cong Lim_cong_within)
done
qed auto
then have dd: "((\<lambda>x. if g x = g z then exp(g z) else (exp(g x) - exp(g z)) / (g x - g z)) \<longlongrightarrow> exp(g z)) (at z within S)"
by (simp add: o_def)
have "continuous (at z within S) g"
using contg continuous_on_eq_continuous_within \<open>z \<in> S\<close> by blast
then have "(\<forall>\<^sub>F x in at z within S. dist (g x) (g z) < 2*pi)"
by (simp add: continuous_within tendsto_iff)
then have "\<forall>\<^sub>F x in at z within S. exp (g x) = exp (g z) \<longrightarrow> g x \<noteq> g z \<longrightarrow> x = z"
apply (rule eventually_mono)
apply (auto simp: exp_eq dist_norm norm_mult)
done
then have "((\<lambda>y. (g y - g z) / (y - z)) \<longlongrightarrow> f' / exp (g z)) (at z within S)"
by (auto intro!: Lim_transform_eventually [OF tendsto_divide [OF ee dd]])
then show ?thesis
by (auto simp: field_differentiable_def has_field_derivative_iff)
qed
then have "g holomorphic_on S"
using holf holomorphic_on_def by auto
then show ?thesis
using feq that by auto
qed
lemma contractible_imp_holomorphic_sqrt:
assumes holf: "f holomorphic_on S"
and S: "contractible S"
and fnz: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"
obtains g where "g holomorphic_on S" "\<And>z. z \<in> S \<Longrightarrow> f z = g z ^ 2"
proof -
obtain g where holg: "g holomorphic_on S" and feq: "\<And>z. z \<in> S \<Longrightarrow> f z = exp(g z)"
using contractible_imp_holomorphic_log [OF assms] by blast
show ?thesis
proof
show "exp \<circ> (\<lambda>z. z / 2) \<circ> g holomorphic_on S"
by (intro holomorphic_on_compose holg holomorphic_intros) auto
show "\<And>z. z \<in> S \<Longrightarrow> f z = ((exp \<circ> (\<lambda>z. z / 2) \<circ> g) z)\<^sup>2"
apply (auto simp: feq)
by (metis eq_divide_eq_numeral1(1) exp_double mult.commute zero_neq_numeral)
qed
qed
lemma simply_connected_imp_holomorphic_sqrt:
assumes holf: "f holomorphic_on S"
and S: "simply_connected S" "locally path_connected S"
and fnz: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"
obtains g where "g holomorphic_on S" "\<And>z. z \<in> S \<Longrightarrow> f z = g z ^ 2"
proof -
obtain g where holg: "g holomorphic_on S" and feq: "\<And>z. z \<in> S \<Longrightarrow> f z = exp(g z)"
using simply_connected_imp_holomorphic_log [OF assms] by blast
show ?thesis
proof
show "exp \<circ> (\<lambda>z. z / 2) \<circ> g holomorphic_on S"
by (intro holomorphic_on_compose holg holomorphic_intros) auto
show "\<And>z. z \<in> S \<Longrightarrow> f z = ((exp \<circ> (\<lambda>z. z / 2) \<circ> g) z)\<^sup>2"
apply (auto simp: feq)
by (metis eq_divide_eq_numeral1(1) exp_double mult.commute zero_neq_numeral)
qed
qed
text\<open> Related theorems about holomorphic inverse cosines.\<close>
lemma contractible_imp_holomorphic_arccos:
assumes holf: "f holomorphic_on S" and S: "contractible S"
and non1: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 1 \<and> f z \<noteq> -1"
obtains g where "g holomorphic_on S" "\<And>z. z \<in> S \<Longrightarrow> f z = cos(g z)"
proof -
have hol1f: "(\<lambda>z. 1 - f z ^ 2) holomorphic_on S"
by (intro holomorphic_intros holf)
obtain g where holg: "g holomorphic_on S" and eq: "\<And>z. z \<in> S \<Longrightarrow> 1 - (f z)\<^sup>2 = (g z)\<^sup>2"
using contractible_imp_holomorphic_sqrt [OF hol1f S]
by (metis eq_iff_diff_eq_0 non1 power2_eq_1_iff)
have holfg: "(\<lambda>z. f z + \<i>*g z) holomorphic_on S"
by (intro holf holg holomorphic_intros)
have "\<And>z. z \<in> S \<Longrightarrow> f z + \<i>*g z \<noteq> 0"
by (metis Arccos_body_lemma eq add.commute add.inverse_unique complex_i_mult_minus power2_csqrt power2_eq_iff)
then obtain h where holh: "h holomorphic_on S" and fgeq: "\<And>z. z \<in> S \<Longrightarrow> f z + \<i>*g z = exp (h z)"
using contractible_imp_holomorphic_log [OF holfg S] by metis
show ?thesis
proof
show "(\<lambda>z. -\<i>*h z) holomorphic_on S"
by (intro holh holomorphic_intros)
show "f z = cos (- \<i>*h z)" if "z \<in> S" for z
proof -
have "(f z + \<i>*g z)*(f z - \<i>*g z) = 1"
using that eq by (auto simp: algebra_simps power2_eq_square)
then have "f z - \<i>*g z = inverse (f z + \<i>*g z)"
using inverse_unique by force
also have "... = exp (- h z)"
by (simp add: exp_minus fgeq that)
finally have "f z = exp (- h z) + \<i>*g z"
by (simp add: diff_eq_eq)
then show ?thesis
apply (simp add: cos_exp_eq)
by (metis fgeq add.assoc mult_2_right that)
qed
qed
qed
lemma contractible_imp_holomorphic_arccos_bounded:
assumes holf: "f holomorphic_on S" and S: "contractible S" and "a \<in> S"
and non1: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 1 \<and> f z \<noteq> -1"
obtains g where "g holomorphic_on S" "norm(g a) \<le> pi + norm(f a)" "\<And>z. z \<in> S \<Longrightarrow> f z = cos(g z)"
proof -
obtain g where holg: "g holomorphic_on S" and feq: "\<And>z. z \<in> S \<Longrightarrow> f z = cos (g z)"
using contractible_imp_holomorphic_arccos [OF holf S non1] by blast
obtain b where "cos b = f a" "norm b \<le> pi + norm (f a)"
using cos_Arccos norm_Arccos_bounded by blast
then have "cos b = cos (g a)"
by (simp add: \<open>a \<in> S\<close> feq)
then consider n where "n \<in> \<int>" "b = g a + of_real(2*n*pi)" | n where "n \<in> \<int>" "b = -g a + of_real(2*n*pi)"
by (auto simp: complex_cos_eq)
then show ?thesis
proof cases
case 1
show ?thesis
proof
show "(\<lambda>z. g z + of_real(2*n*pi)) holomorphic_on S"
by (intro holomorphic_intros holg)
show "cmod (g a + of_real(2*n*pi)) \<le> pi + cmod (f a)"
using "1" \<open>cmod b \<le> pi + cmod (f a)\<close> by blast
show "\<And>z. z \<in> S \<Longrightarrow> f z = cos (g z + complex_of_real (2*n*pi))"
by (metis \<open>n \<in> \<int>\<close> complex_cos_eq feq)
qed
next
case 2
show ?thesis
proof
show "(\<lambda>z. -g z + of_real(2*n*pi)) holomorphic_on S"
by (intro holomorphic_intros holg)
show "cmod (-g a + of_real(2*n*pi)) \<le> pi + cmod (f a)"
using "2" \<open>cmod b \<le> pi + cmod (f a)\<close> by blast
show "\<And>z. z \<in> S \<Longrightarrow> f z = cos (-g z + complex_of_real (2*n*pi))"
by (metis \<open>n \<in> \<int>\<close> complex_cos_eq feq)
qed
qed
qed
subsection\<open>The "Borsukian" property of sets\<close>
text\<open>This doesn't have a standard name. Kuratowski uses ``contractible with respect to \<open>[S\<^sup>1]\<close>''
while Whyburn uses ``property b''. It's closely related to unicoherence.\<close>
definition\<^marker>\<open>tag important\<close> Borsukian where
"Borsukian S \<equiv>
\<forall>f. continuous_on S f \<and> f ` S \<subseteq> (- {0::complex})
\<longrightarrow> (\<exists>a. homotopic_with_canon (\<lambda>h. True) S (- {0}) f (\<lambda>x. a))"
lemma Borsukian_retraction_gen:
assumes "Borsukian S" "continuous_on S h" "h ` S = T"
"continuous_on T k" "k ` T \<subseteq> S" "\<And>y. y \<in> T \<Longrightarrow> h(k y) = y"
shows "Borsukian T"
proof -
interpret R: Retracts S h T k
using assms by (simp add: Retracts.intro)
show ?thesis
using assms
apply (simp add: Borsukian_def, clarify)
apply (rule R.cohomotopically_trivial_retraction_null_gen [OF TrueI TrueI refl, of "-{0}"], auto)
done
qed
lemma retract_of_Borsukian: "\<lbrakk>Borsukian T; S retract_of T\<rbrakk> \<Longrightarrow> Borsukian S"
apply (auto simp: retract_of_def retraction_def)
apply (erule (1) Borsukian_retraction_gen)
apply (meson retraction retraction_def)
apply (auto)
done
lemma homeomorphic_Borsukian: "\<lbrakk>Borsukian S; S homeomorphic T\<rbrakk> \<Longrightarrow> Borsukian T"
using Borsukian_retraction_gen order_refl
by (fastforce simp add: homeomorphism_def homeomorphic_def)
lemma homeomorphic_Borsukian_eq:
"S homeomorphic T \<Longrightarrow> Borsukian S \<longleftrightarrow> Borsukian T"
by (meson homeomorphic_Borsukian homeomorphic_sym)
lemma Borsukian_translation:
fixes S :: "'a::real_normed_vector set"
shows "Borsukian (image (\<lambda>x. a + x) S) \<longleftrightarrow> Borsukian S"
apply (rule homeomorphic_Borsukian_eq)
using homeomorphic_translation homeomorphic_sym by blast
lemma Borsukian_injective_linear_image:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "linear f" "inj f"
shows "Borsukian(f ` S) \<longleftrightarrow> Borsukian S"
apply (rule homeomorphic_Borsukian_eq)
using assms homeomorphic_sym linear_homeomorphic_image by blast
lemma homotopy_eqv_Borsukianness:
fixes S :: "'a::real_normed_vector set"
and T :: "'b::real_normed_vector set"
assumes "S homotopy_eqv T"
shows "(Borsukian S \<longleftrightarrow> Borsukian T)"
by (meson Borsukian_def assms homotopy_eqv_cohomotopic_triviality_null)
lemma Borsukian_alt:
fixes S :: "'a::real_normed_vector set"
shows
"Borsukian S \<longleftrightarrow>
(\<forall>f g. continuous_on S f \<and> f ` S \<subseteq> -{0} \<and>
continuous_on S g \<and> g ` S \<subseteq> -{0}
\<longrightarrow> homotopic_with_canon (\<lambda>h. True) S (- {0::complex}) f g)"
unfolding Borsukian_def homotopic_triviality
by (simp add: path_connected_punctured_universe)
lemma Borsukian_continuous_logarithm:
fixes S :: "'a::real_normed_vector set"
shows "Borsukian S \<longleftrightarrow>
(\<forall>f. continuous_on S f \<and> f ` S \<subseteq> (- {0::complex})
\<longrightarrow> (\<exists>g. continuous_on S g \<and> (\<forall>x \<in> S. f x = exp(g x))))"
by (simp add: Borsukian_def inessential_eq_continuous_logarithm)
lemma Borsukian_continuous_logarithm_circle:
fixes S :: "'a::real_normed_vector set"
shows "Borsukian S \<longleftrightarrow>
(\<forall>f. continuous_on S f \<and> f ` S \<subseteq> sphere (0::complex) 1
\<longrightarrow> (\<exists>g. continuous_on S g \<and> (\<forall>x \<in> S. f x = exp(g x))))"
(is "?lhs = ?rhs")
proof
assume ?lhs then show ?rhs
by (force simp: Borsukian_continuous_logarithm)
next
assume RHS [rule_format]: ?rhs
show ?lhs
proof (clarsimp simp: Borsukian_continuous_logarithm)
fix f :: "'a \<Rightarrow> complex"
assume contf: "continuous_on S f" and 0: "0 \<notin> f ` S"
then have "continuous_on S (\<lambda>x. f x / complex_of_real (cmod (f x)))"
by (intro continuous_intros) auto
moreover have "(\<lambda>x. f x / complex_of_real (cmod (f x))) ` S \<subseteq> sphere 0 1"
using 0 by (auto simp: norm_divide)
ultimately obtain g where contg: "continuous_on S g"
and fg: "\<forall>x \<in> S. f x / complex_of_real (cmod (f x)) = exp(g x)"
using RHS [of "\<lambda>x. f x / of_real(norm(f x))"] by auto
show "\<exists>g. continuous_on S g \<and> (\<forall>x\<in>S. f x = exp (g x))"
proof (intro exI ballI conjI)
show "continuous_on S (\<lambda>x. (Ln \<circ> of_real \<circ> norm \<circ> f)x + g x)"
by (intro continuous_intros contf contg conjI) (use "0" in auto)
show "f x = exp ((Ln \<circ> complex_of_real \<circ> cmod \<circ> f) x + g x)" if "x \<in> S" for x
using 0 that
apply (clarsimp simp: exp_add)
apply (subst exp_Ln, force)
by (metis eq_divide_eq exp_not_eq_zero fg mult.commute)
qed
qed
qed
lemma Borsukian_continuous_logarithm_circle_real:
fixes S :: "'a::real_normed_vector set"
shows "Borsukian S \<longleftrightarrow>
(\<forall>f. continuous_on S f \<and> f ` S \<subseteq> sphere (0::complex) 1
\<longrightarrow> (\<exists>g. continuous_on S (complex_of_real \<circ> g) \<and> (\<forall>x \<in> S. f x = exp(\<i> * of_real(g x)))))"
(is "?lhs = ?rhs")
proof
assume LHS: ?lhs
show ?rhs
proof (clarify)
fix f :: "'a \<Rightarrow> complex"
assume "continuous_on S f" and f01: "f ` S \<subseteq> sphere 0 1"
then obtain g where contg: "continuous_on S g" and "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"
using LHS by (auto simp: Borsukian_continuous_logarithm_circle)
then have "\<forall>x\<in>S. f x = exp (\<i> * complex_of_real ((Im \<circ> g) x))"
using f01 apply (simp add: image_iff subset_iff)
by (metis cis_conv_exp exp_eq_polar mult.left_neutral norm_exp_eq_Re of_real_1)
then show "\<exists>g. continuous_on S (complex_of_real \<circ> g) \<and> (\<forall>x\<in>S. f x = exp (\<i> * complex_of_real (g x)))"
by (rule_tac x="Im \<circ> g" in exI) (force intro: continuous_intros contg)
qed
next
assume RHS [rule_format]: ?rhs
show ?lhs
proof (clarsimp simp: Borsukian_continuous_logarithm_circle)
fix f :: "'a \<Rightarrow> complex"
assume "continuous_on S f" and f01: "f ` S \<subseteq> sphere 0 1"
then obtain g where contg: "continuous_on S (complex_of_real \<circ> g)" and "\<And>x. x \<in> S \<Longrightarrow> f x = exp(\<i> * of_real(g x))"
by (metis RHS)
then show "\<exists>g. continuous_on S g \<and> (\<forall>x\<in>S. f x = exp (g x))"
by (rule_tac x="\<lambda>x. \<i>* of_real(g x)" in exI) (auto simp: continuous_intros contg)
qed
qed
lemma Borsukian_circle:
fixes S :: "'a::real_normed_vector set"
shows "Borsukian S \<longleftrightarrow>
(\<forall>f. continuous_on S f \<and> f ` S \<subseteq> sphere (0::complex) 1
\<longrightarrow> (\<exists>a. homotopic_with_canon (\<lambda>h. True) S (sphere (0::complex) 1) f (\<lambda>x. a)))"
by (simp add: inessential_eq_continuous_logarithm_circle Borsukian_continuous_logarithm_circle_real)
lemma contractible_imp_Borsukian: "contractible S \<Longrightarrow> Borsukian S"
by (meson Borsukian_def nullhomotopic_from_contractible)
lemma simply_connected_imp_Borsukian:
fixes S :: "'a::real_normed_vector set"
shows "\<lbrakk>simply_connected S; locally path_connected S\<rbrakk> \<Longrightarrow> Borsukian S"
apply (simp add: Borsukian_continuous_logarithm)
by (metis (no_types, lifting) continuous_logarithm_on_simply_connected image_iff)
lemma starlike_imp_Borsukian:
fixes S :: "'a::real_normed_vector set"
shows "starlike S \<Longrightarrow> Borsukian S"
by (simp add: contractible_imp_Borsukian starlike_imp_contractible)
lemma Borsukian_empty: "Borsukian {}"
by (auto simp: contractible_imp_Borsukian)
lemma Borsukian_UNIV: "Borsukian (UNIV :: 'a::real_normed_vector set)"
by (auto simp: contractible_imp_Borsukian)
lemma convex_imp_Borsukian:
fixes S :: "'a::real_normed_vector set"
shows "convex S \<Longrightarrow> Borsukian S"
by (meson Borsukian_def convex_imp_contractible nullhomotopic_from_contractible)
proposition Borsukian_sphere:
fixes a :: "'a::euclidean_space"
shows "3 \<le> DIM('a) \<Longrightarrow> Borsukian (sphere a r)"
apply (rule simply_connected_imp_Borsukian)
using simply_connected_sphere apply blast
using ENR_imp_locally_path_connected ENR_sphere by blast
proposition Borsukian_open_Un:
fixes S :: "'a::real_normed_vector set"
assumes opeS: "openin (top_of_set (S \<union> T)) S"
and opeT: "openin (top_of_set (S \<union> T)) T"
and BS: "Borsukian S" and BT: "Borsukian T" and ST: "connected(S \<inter> T)"
shows "Borsukian(S \<union> T)"
proof (clarsimp simp add: Borsukian_continuous_logarithm)
fix f :: "'a \<Rightarrow> complex"
assume contf: "continuous_on (S \<union> T) f" and 0: "0 \<notin> f ` (S \<union> T)"
then have contfS: "continuous_on S f" and contfT: "continuous_on T f"
using continuous_on_subset by auto
have "\<lbrakk>continuous_on S f; f ` S \<subseteq> -{0}\<rbrakk> \<Longrightarrow> \<exists>g. continuous_on S g \<and> (\<forall>x \<in> S. f x = exp(g x))"
using BS by (auto simp: Borsukian_continuous_logarithm)
then obtain g where contg: "continuous_on S g" and fg: "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"
using "0" contfS by blast
have "\<lbrakk>continuous_on T f; f ` T \<subseteq> -{0}\<rbrakk> \<Longrightarrow> \<exists>g. continuous_on T g \<and> (\<forall>x \<in> T. f x = exp(g x))"
using BT by (auto simp: Borsukian_continuous_logarithm)
then obtain h where conth: "continuous_on T h" and fh: "\<And>x. x \<in> T \<Longrightarrow> f x = exp(h x)"
using "0" contfT by blast
show "\<exists>g. continuous_on (S \<union> T) g \<and> (\<forall>x\<in>S \<union> T. f x = exp (g x))"
proof (cases "S \<inter> T = {}")
case True
show ?thesis
proof (intro exI conjI)
show "continuous_on (S \<union> T) (\<lambda>x. if x \<in> S then g x else h x)"
apply (rule continuous_on_cases_local_open [OF opeS opeT contg conth])
using True by blast
show "\<forall>x\<in>S \<union> T. f x = exp (if x \<in> S then g x else h x)"
using fg fh by auto
qed
next
case False
have "(\<lambda>x. g x - h x) constant_on S \<inter> T"
proof (rule continuous_discrete_range_constant [OF ST])
show "continuous_on (S \<inter> T) (\<lambda>x. g x - h x)"
apply (intro continuous_intros)
apply (meson contg continuous_on_subset inf_le1)
by (meson conth continuous_on_subset inf_sup_ord(2))
show "\<exists>e>0. \<forall>y. y \<in> S \<inter> T \<and> g y - h y \<noteq> g x - h x \<longrightarrow> e \<le> cmod (g y - h y - (g x - h x))"
if "x \<in> S \<inter> T" for x
proof -
have "g y - g x = h y - h x"
if "y \<in> S" "y \<in> T" "cmod (g y - g x - (h y - h x)) < 2 * pi" for y
proof (rule exp_complex_eqI)
have "\<bar>Im (g y - g x) - Im (h y - h x)\<bar> \<le> cmod (g y - g x - (h y - h x))"
by (metis abs_Im_le_cmod minus_complex.simps(2))
then show "\<bar>Im (g y - g x) - Im (h y - h x)\<bar> < 2 * pi"
using that by linarith
have "exp (g x) = exp (h x)" "exp (g y) = exp (h y)"
using fg fh that \<open>x \<in> S \<inter> T\<close> by fastforce+
then show "exp (g y - g x) = exp (h y - h x)"
by (simp add: exp_diff)
qed
then show ?thesis
by (rule_tac x="2*pi" in exI) (fastforce simp add: algebra_simps)
qed
qed
then obtain a where a: "\<And>x. x \<in> S \<inter> T \<Longrightarrow> g x - h x = a"
by (auto simp: constant_on_def)
with False have "exp a = 1"
by (metis IntI disjoint_iff_not_equal divide_self_if exp_diff exp_not_eq_zero fg fh)
with a show ?thesis
apply (rule_tac x="\<lambda>x. if x \<in> S then g x else a + h x" in exI)
apply (intro continuous_on_cases_local_open opeS opeT contg conth continuous_intros conjI)
apply (auto simp: algebra_simps fg fh exp_add)
done
qed
qed
text\<open>The proof is a duplicate of that of \<open>Borsukian_open_Un\<close>.\<close>
lemma Borsukian_closed_Un:
fixes S :: "'a::real_normed_vector set"
assumes cloS: "closedin (top_of_set (S \<union> T)) S"
and cloT: "closedin (top_of_set (S \<union> T)) T"
and BS: "Borsukian S" and BT: "Borsukian T" and ST: "connected(S \<inter> T)"
shows "Borsukian(S \<union> T)"
proof (clarsimp simp add: Borsukian_continuous_logarithm)
fix f :: "'a \<Rightarrow> complex"
assume contf: "continuous_on (S \<union> T) f" and 0: "0 \<notin> f ` (S \<union> T)"
then have contfS: "continuous_on S f" and contfT: "continuous_on T f"
using continuous_on_subset by auto
have "\<lbrakk>continuous_on S f; f ` S \<subseteq> -{0}\<rbrakk> \<Longrightarrow> \<exists>g. continuous_on S g \<and> (\<forall>x \<in> S. f x = exp(g x))"
using BS by (auto simp: Borsukian_continuous_logarithm)
then obtain g where contg: "continuous_on S g" and fg: "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"
using "0" contfS by blast
have "\<lbrakk>continuous_on T f; f ` T \<subseteq> -{0}\<rbrakk> \<Longrightarrow> \<exists>g. continuous_on T g \<and> (\<forall>x \<in> T. f x = exp(g x))"
using BT by (auto simp: Borsukian_continuous_logarithm)
then obtain h where conth: "continuous_on T h" and fh: "\<And>x. x \<in> T \<Longrightarrow> f x = exp(h x)"
using "0" contfT by blast
show "\<exists>g. continuous_on (S \<union> T) g \<and> (\<forall>x\<in>S \<union> T. f x = exp (g x))"
proof (cases "S \<inter> T = {}")
case True
show ?thesis
proof (intro exI conjI)
show "continuous_on (S \<union> T) (\<lambda>x. if x \<in> S then g x else h x)"
apply (rule continuous_on_cases_local [OF cloS cloT contg conth])
using True by blast
show "\<forall>x\<in>S \<union> T. f x = exp (if x \<in> S then g x else h x)"
using fg fh by auto
qed
next
case False
have "(\<lambda>x. g x - h x) constant_on S \<inter> T"
proof (rule continuous_discrete_range_constant [OF ST])
show "continuous_on (S \<inter> T) (\<lambda>x. g x - h x)"
apply (intro continuous_intros)
apply (meson contg continuous_on_subset inf_le1)
by (meson conth continuous_on_subset inf_sup_ord(2))
show "\<exists>e>0. \<forall>y. y \<in> S \<inter> T \<and> g y - h y \<noteq> g x - h x \<longrightarrow> e \<le> cmod (g y - h y - (g x - h x))"
if "x \<in> S \<inter> T" for x
proof -
have "g y - g x = h y - h x"
if "y \<in> S" "y \<in> T" "cmod (g y - g x - (h y - h x)) < 2 * pi" for y
proof (rule exp_complex_eqI)
have "\<bar>Im (g y - g x) - Im (h y - h x)\<bar> \<le> cmod (g y - g x - (h y - h x))"
by (metis abs_Im_le_cmod minus_complex.simps(2))
then show "\<bar>Im (g y - g x) - Im (h y - h x)\<bar> < 2 * pi"
using that by linarith
have "exp (g x) = exp (h x)" "exp (g y) = exp (h y)"
using fg fh that \<open>x \<in> S \<inter> T\<close> by fastforce+
then show "exp (g y - g x) = exp (h y - h x)"
by (simp add: exp_diff)
qed
then show ?thesis
by (rule_tac x="2*pi" in exI) (fastforce simp add: algebra_simps)
qed
qed
then obtain a where a: "\<And>x. x \<in> S \<inter> T \<Longrightarrow> g x - h x = a"
by (auto simp: constant_on_def)
with False have "exp a = 1"
by (metis IntI disjoint_iff_not_equal divide_self_if exp_diff exp_not_eq_zero fg fh)
with a show ?thesis
apply (rule_tac x="\<lambda>x. if x \<in> S then g x else a + h x" in exI)
apply (intro continuous_on_cases_local cloS cloT contg conth continuous_intros conjI)
apply (auto simp: algebra_simps fg fh exp_add)
done
qed
qed
lemma Borsukian_separation_compact:
fixes S :: "complex set"
assumes "compact S"
shows "Borsukian S \<longleftrightarrow> connected(- S)"
by (simp add: Borsuk_separation_theorem Borsukian_circle assms)
lemma Borsukian_monotone_image_compact:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "Borsukian S" and contf: "continuous_on S f" and fim: "f ` S = T"
and "compact S" and conn: "\<And>y. y \<in> T \<Longrightarrow> connected {x. x \<in> S \<and> f x = y}"
shows "Borsukian T"
proof (clarsimp simp add: Borsukian_continuous_logarithm)
fix g :: "'b \<Rightarrow> complex"
assume contg: "continuous_on T g" and 0: "0 \<notin> g ` T"
have "continuous_on S (g \<circ> f)"
using contf contg continuous_on_compose fim by blast
moreover have "(g \<circ> f) ` S \<subseteq> -{0}"
using fim 0 by auto
ultimately obtain h where conth: "continuous_on S h" and gfh: "\<And>x. x \<in> S \<Longrightarrow> (g \<circ> f) x = exp(h x)"
using \<open>Borsukian S\<close> by (auto simp: Borsukian_continuous_logarithm)
have "\<And>y. \<exists>x. y \<in> T \<longrightarrow> x \<in> S \<and> f x = y"
using fim by auto
then obtain f' where f': "\<And>y. y \<in> T \<longrightarrow> f' y \<in> S \<and> f (f' y) = y"
by metis
have *: "(\<lambda>x. h x - h(f' y)) constant_on {x. x \<in> S \<and> f x = y}" if "y \<in> T" for y
proof (rule continuous_discrete_range_constant [OF conn [OF that], of "\<lambda>x. h x - h (f' y)"], simp_all add: algebra_simps)
show "continuous_on {x \<in> S. f x = y} (\<lambda>x. h x - h (f' y))"
by (intro continuous_intros continuous_on_subset [OF conth]) auto
show "\<exists>e>0. \<forall>u. u \<in> S \<and> f u = y \<and> h u \<noteq> h x \<longrightarrow> e \<le> cmod (h u - h x)"
if x: "x \<in> S \<and> f x = y" for x
proof -
have "h u = h x" if "u \<in> S" "f u = y" "cmod (h u - h x) < 2 * pi" for u
proof (rule exp_complex_eqI)
have "\<bar>Im (h u) - Im (h x)\<bar> \<le> cmod (h u - h x)"
by (metis abs_Im_le_cmod minus_complex.simps(2))
then show "\<bar>Im (h u) - Im (h x)\<bar> < 2 * pi"
using that by linarith
show "exp (h u) = exp (h x)"
by (simp add: gfh [symmetric] x that)
qed
then show ?thesis
by (rule_tac x="2*pi" in exI) (fastforce simp add: algebra_simps)
qed
qed
have "h x = h (f' (f x))" if "x \<in> S" for x
using * [of "f x"] fim that unfolding constant_on_def by clarsimp (metis f' imageI right_minus_eq)
moreover have "\<And>x. x \<in> T \<Longrightarrow> \<exists>u. u \<in> S \<and> x = f u \<and> h (f' x) = h u"
using f' by fastforce
ultimately
have eq: "((\<lambda>x. (x, (h \<circ> f') x)) ` T) =
{p. \<exists>x. x \<in> S \<and> (x, p) \<in> (S \<times> UNIV) \<inter> ((\<lambda>z. snd z - ((f \<circ> fst) z, (h \<circ> fst) z)) -` {0})}"
using fim by (auto simp: image_iff)
show "\<exists>h. continuous_on T h \<and> (\<forall>x\<in>T. g x = exp (h x))"
proof (intro exI conjI)
show "continuous_on T (h \<circ> f')"
proof (rule continuous_from_closed_graph [of "h ` S"])
show "compact (h ` S)"
by (simp add: \<open>compact S\<close> compact_continuous_image conth)
show "(h \<circ> f') ` T \<subseteq> h ` S"
by (auto simp: f')
show "closed ((\<lambda>x. (x, (h \<circ> f') x)) ` T)"
apply (subst eq)
apply (intro closed_compact_projection [OF \<open>compact S\<close>] continuous_closed_preimage
continuous_intros continuous_on_subset [OF contf] continuous_on_subset [OF conth])
apply (auto simp: \<open>compact S\<close> closed_Times compact_imp_closed)
done
qed
qed (use f' gfh in fastforce)
qed
lemma Borsukian_open_map_image_compact:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "Borsukian S" and contf: "continuous_on S f" and fim: "f ` S = T" and "compact S"
and ope: "\<And>U. openin (top_of_set S) U
\<Longrightarrow> openin (top_of_set T) (f ` U)"
shows "Borsukian T"
proof (clarsimp simp add: Borsukian_continuous_logarithm_circle_real)
fix g :: "'b \<Rightarrow> complex"
assume contg: "continuous_on T g" and gim: "g ` T \<subseteq> sphere 0 1"
have "continuous_on S (g \<circ> f)"
using contf contg continuous_on_compose fim by blast
moreover have "(g \<circ> f) ` S \<subseteq> sphere 0 1"
using fim gim by auto
ultimately obtain h where cont_cxh: "continuous_on S (complex_of_real \<circ> h)"
and gfh: "\<And>x. x \<in> S \<Longrightarrow> (g \<circ> f) x = exp(\<i> * of_real(h x))"
using \<open>Borsukian S\<close> Borsukian_continuous_logarithm_circle_real by metis
then have conth: "continuous_on S h"
by simp
have "\<exists>x. x \<in> S \<and> f x = y \<and> (\<forall>x' \<in> S. f x' = y \<longrightarrow> h x \<le> h x')" if "y \<in> T" for y
proof -
have 1: "compact (h ` {x \<in> S. f x = y})"
proof (rule compact_continuous_image)
show "continuous_on {x \<in> S. f x = y} h"
by (rule continuous_on_subset [OF conth]) auto
have "compact (S \<inter> f -` {y})"
by (rule proper_map_from_compact [OF contf _ \<open>compact S\<close>, of T]) (simp_all add: fim that)
then show "compact {x \<in> S. f x = y}"
by (auto simp: vimage_def Int_def)
qed
have 2: "h ` {x \<in> S. f x = y} \<noteq> {}"
using fim that by auto
have "\<exists>s \<in> h ` {x \<in> S. f x = y}. \<forall>t \<in> h ` {x \<in> S. f x = y}. s \<le> t"
using compact_attains_inf [OF 1 2] by blast
then show ?thesis by auto
qed
then obtain k where kTS: "\<And>y. y \<in> T \<Longrightarrow> k y \<in> S"
and fk: "\<And>y. y \<in> T \<Longrightarrow> f (k y) = y "
and hle: "\<And>x' y. \<lbrakk>y \<in> T; x' \<in> S; f x' = y\<rbrakk> \<Longrightarrow> h (k y) \<le> h x'"
by metis
have "continuous_on T (h \<circ> k)"
proof (clarsimp simp add: continuous_on_iff)
fix y and e::real
assume "y \<in> T" "0 < e"
moreover have "uniformly_continuous_on S (complex_of_real \<circ> h)"
using \<open>compact S\<close> cont_cxh compact_uniformly_continuous by blast
ultimately obtain d where "0 < d"
and d: "\<And>x x'. \<lbrakk>x\<in>S; x'\<in>S; dist x' x < d\<rbrakk> \<Longrightarrow> dist (h x') (h x) < e"
by (force simp: uniformly_continuous_on_def)
obtain \<delta> where "0 < \<delta>" and \<delta>:
"\<And>x'. \<lbrakk>x' \<in> T; dist y x' < \<delta>\<rbrakk>
\<Longrightarrow> (\<forall>v \<in> {z \<in> S. f z = y}. \<exists>v'. v' \<in> {z \<in> S. f z = x'} \<and> dist v v' < d) \<and>
(\<forall>v' \<in> {z \<in> S. f z = x'}. \<exists>v. v \<in> {z \<in> S. f z = y} \<and> dist v' v < d)"
proof (rule upper_lower_hemicontinuous_explicit [of T "\<lambda>y. {z \<in> S. f z = y}" S])
show "\<And>U. openin (top_of_set S) U
\<Longrightarrow> openin (top_of_set T) {x \<in> T. {z \<in> S. f z = x} \<subseteq> U}"
using closed_map_iff_upper_hemicontinuous_preimage [OF fim [THEN equalityD1]]
by (simp add: Abstract_Topology_2.continuous_imp_closed_map \<open>compact S\<close> contf fim)
show "\<And>U. closedin (top_of_set S) U \<Longrightarrow>
closedin (top_of_set T) {x \<in> T. {z \<in> S. f z = x} \<subseteq> U}"
using ope open_map_iff_lower_hemicontinuous_preimage [OF fim [THEN equalityD1]]
by meson
show "bounded {z \<in> S. f z = y}"
by (metis (no_types, lifting) compact_imp_bounded [OF \<open>compact S\<close>] bounded_subset mem_Collect_eq subsetI)
qed (use \<open>y \<in> T\<close> \<open>0 < d\<close> fk kTS in \<open>force+\<close>)
have "dist (h (k y')) (h (k y)) < e" if "y' \<in> T" "dist y y' < \<delta>" for y'
proof -
have k1: "k y \<in> S" "f (k y) = y" and k2: "k y' \<in> S" "f (k y') = y'"
by (auto simp: \<open>y \<in> T\<close> \<open>y' \<in> T\<close> kTS fk)
have 1: "\<And>v. \<lbrakk>v \<in> S; f v = y\<rbrakk> \<Longrightarrow> \<exists>v'. v' \<in> {z \<in> S. f z = y'} \<and> dist v v' < d"
and 2: "\<And>v'. \<lbrakk>v' \<in> S; f v' = y'\<rbrakk> \<Longrightarrow> \<exists>v. v \<in> {z \<in> S. f z = y} \<and> dist v' v < d"
using \<delta> [OF that] by auto
then obtain w' w where "w' \<in> S" "f w' = y'" "dist (k y) w' < d"
and "w \<in> S" "f w = y" "dist (k y') w < d"
using 1 [OF k1] 2 [OF k2] by auto
then show ?thesis
using d [of w "k y'"] d [of w' "k y"] k1 k2 \<open>y' \<in> T\<close> \<open>y \<in> T\<close> hle
by (fastforce simp: dist_norm abs_diff_less_iff algebra_simps)
qed
then show "\<exists>d>0. \<forall>x'\<in>T. dist x' y < d \<longrightarrow> dist (h (k x')) (h (k y)) < e"
using \<open>0 < \<delta>\<close> by (auto simp: dist_commute)
qed
then show "\<exists>h. continuous_on T h \<and> (\<forall>x\<in>T. g x = exp (\<i> * complex_of_real (h x)))"
using fk gfh kTS by force
qed
text\<open>If two points are separated by a closed set, there's a minimal one.\<close>
proposition closed_irreducible_separator:
fixes a :: "'a::real_normed_vector"
assumes "closed S" and ab: "\<not> connected_component (- S) a b"
obtains T where "T \<subseteq> S" "closed T" "T \<noteq> {}" "\<not> connected_component (- T) a b"
"\<And>U. U \<subset> T \<Longrightarrow> connected_component (- U) a b"
proof (cases "a \<in> S \<or> b \<in> S")
case True
then show ?thesis
proof
assume *: "a \<in> S"
show ?thesis
proof
show "{a} \<subseteq> S"
using * by blast
show "\<not> connected_component (- {a}) a b"
using connected_component_in by auto
show "\<And>U. U \<subset> {a} \<Longrightarrow> connected_component (- U) a b"
by (metis connected_component_UNIV UNIV_I compl_bot_eq connected_component_eq_eq less_le_not_le subset_singletonD)
qed auto
next
assume *: "b \<in> S"
show ?thesis
proof
show "{b} \<subseteq> S"
using * by blast
show "\<not> connected_component (- {b}) a b"
using connected_component_in by auto
show "\<And>U. U \<subset> {b} \<Longrightarrow> connected_component (- U) a b"
by (metis connected_component_UNIV UNIV_I compl_bot_eq connected_component_eq_eq less_le_not_le subset_singletonD)
qed auto
qed
next
case False
define A where "A \<equiv> connected_component_set (- S) a"
define B where "B \<equiv> connected_component_set (- (closure A)) b"
have "a \<in> A"
using False A_def by auto
have "b \<in> B"
unfolding A_def B_def closure_Un_frontier
using ab False \<open>closed S\<close> frontier_complement frontier_of_connected_component_subset frontier_subset_closed by force
have "frontier B \<subseteq> frontier (connected_component_set (- closure A) b)"
using B_def by blast
also have frsub: "... \<subseteq> frontier A"
proof -
have "\<And>A. closure (- closure (- A)) \<subseteq> closure A"
by (metis (no_types) closure_mono closure_subset compl_le_compl_iff double_compl)
then show ?thesis
by (metis (no_types) closure_closure double_compl frontier_closures frontier_of_connected_component_subset le_inf_iff subset_trans)
qed
finally have frBA: "frontier B \<subseteq> frontier A" .
show ?thesis
proof
show "frontier B \<subseteq> S"
proof -
have "frontier S \<subseteq> S"
by (simp add: \<open>closed S\<close> frontier_subset_closed)
then show ?thesis
using frsub frontier_complement frontier_of_connected_component_subset
unfolding A_def B_def by blast
qed
show "closed (frontier B)"
by simp
show "\<not> connected_component (- frontier B) a b"
unfolding connected_component_def
proof clarify
fix T
assume "connected T" and TB: "T \<subseteq> - frontier B" and "a \<in> T" and "b \<in> T"
have "a \<notin> B"
by (metis A_def B_def ComplD \<open>a \<in> A\<close> assms(1) closed_open connected_component_subset in_closure_connected_component subsetD)
have "T \<inter> B \<noteq> {}"
using \<open>b \<in> B\<close> \<open>b \<in> T\<close> by blast
moreover have "T - B \<noteq> {}"
using \<open>a \<notin> B\<close> \<open>a \<in> T\<close> by blast
ultimately show "False"
using connected_Int_frontier [of T B] TB \<open>connected T\<close> by blast
qed
moreover have "connected_component (- frontier B) a b" if "frontier B = {}"
apply (simp add: that)
using connected_component_eq_UNIV by blast
ultimately show "frontier B \<noteq> {}"
by blast
show "connected_component (- U) a b" if "U \<subset> frontier B" for U
proof -
obtain p where Usub: "U \<subseteq> frontier B" and p: "p \<in> frontier B" "p \<notin> U"
using \<open>U \<subset> frontier B\<close> by blast
show ?thesis
unfolding connected_component_def
proof (intro exI conjI)
have "connected ((insert p A) \<union> (insert p B))"
proof (rule connected_Un)
show "connected (insert p A)"
by (metis A_def IntD1 frBA \<open>p \<in> frontier B\<close> closure_insert closure_subset connected_connected_component connected_intermediate_closure frontier_closures insert_absorb subsetCE subset_insertI)
show "connected (insert p B)"
by (metis B_def IntD1 \<open>p \<in> frontier B\<close> closure_insert closure_subset connected_connected_component connected_intermediate_closure frontier_closures insert_absorb subset_insertI)
qed blast
then show "connected (insert p (B \<union> A))"
by (simp add: sup.commute)
have "A \<subseteq> - U"
using A_def Usub \<open>frontier B \<subseteq> S\<close> connected_component_subset by fastforce
moreover have "B \<subseteq> - U"
using B_def Usub connected_component_subset frBA frontier_closures by fastforce
ultimately show "insert p (B \<union> A) \<subseteq> - U"
using p by auto
qed (auto simp: \<open>a \<in> A\<close> \<open>b \<in> B\<close>)
qed
qed
qed
lemma frontier_minimal_separating_closed_pointwise:
fixes S :: "'a::real_normed_vector set"
assumes S: "closed S" "a \<notin> S" and nconn: "\<not> connected_component (- S) a b"
and conn: "\<And>T. \<lbrakk>closed T; T \<subset> S\<rbrakk> \<Longrightarrow> connected_component (- T) a b"
shows "frontier(connected_component_set (- S) a) = S" (is "?F = S")
proof -
have "?F \<subseteq> S"
by (simp add: S componentsI frontier_of_components_closed_complement)
moreover have False if "?F \<subset> S"
proof -
have "connected_component (- ?F) a b"
by (simp add: conn that)
then obtain T where "connected T" "T \<subseteq> -?F" "a \<in> T" "b \<in> T"
by (auto simp: connected_component_def)
moreover have "T \<inter> ?F \<noteq> {}"
proof (rule connected_Int_frontier [OF \<open>connected T\<close>])
show "T \<inter> connected_component_set (- S) a \<noteq> {}"
using \<open>a \<notin> S\<close> \<open>a \<in> T\<close> by fastforce
show "T - connected_component_set (- S) a \<noteq> {}"
using \<open>b \<in> T\<close> nconn by blast
qed
ultimately show ?thesis
by blast
qed
ultimately show ?thesis
by blast
qed
subsection\<open>Unicoherence (closed)\<close>
definition\<^marker>\<open>tag important\<close> unicoherent where
"unicoherent U \<equiv>
\<forall>S T. connected S \<and> connected T \<and> S \<union> T = U \<and>
closedin (top_of_set U) S \<and> closedin (top_of_set U) T
\<longrightarrow> connected (S \<inter> T)"
lemma unicoherentI [intro?]:
assumes "\<And>S T. \<lbrakk>connected S; connected T; U = S \<union> T; closedin (top_of_set U) S; closedin (top_of_set U) T\<rbrakk>
\<Longrightarrow> connected (S \<inter> T)"
shows "unicoherent U"
using assms unfolding unicoherent_def by blast
lemma unicoherentD:
assumes "unicoherent U" "connected S" "connected T" "U = S \<union> T" "closedin (top_of_set U) S" "closedin (top_of_set U) T"
shows "connected (S \<inter> T)"
using assms unfolding unicoherent_def by blast
proposition homeomorphic_unicoherent:
assumes ST: "S homeomorphic T" and S: "unicoherent S"
shows "unicoherent T"
proof -
obtain f g where gf: "\<And>x. x \<in> S \<Longrightarrow> g (f x) = x" and fim: "T = f ` S" and gfim: "g ` f ` S = S"
and contf: "continuous_on S f" and contg: "continuous_on (f ` S) g"
using ST by (auto simp: homeomorphic_def homeomorphism_def)
show ?thesis
proof
fix U V
assume "connected U" "connected V" and T: "T = U \<union> V"
and cloU: "closedin (top_of_set T) U"
and cloV: "closedin (top_of_set T) V"
have "f ` (g ` U \<inter> g ` V) \<subseteq> U" "f ` (g ` U \<inter> g ` V) \<subseteq> V"
using gf fim T by auto (metis UnCI image_iff)+
moreover have "U \<inter> V \<subseteq> f ` (g ` U \<inter> g ` V)"
using gf fim by (force simp: image_iff T)
ultimately have "U \<inter> V = f ` (g ` U \<inter> g ` V)" by blast
moreover have "connected (f ` (g ` U \<inter> g ` V))"
proof (rule connected_continuous_image)
show "continuous_on (g ` U \<inter> g ` V) f"
apply (rule continuous_on_subset [OF contf])
using T fim gfim by blast
show "connected (g ` U \<inter> g ` V)"
proof (intro conjI unicoherentD [OF S])
show "connected (g ` U)" "connected (g ` V)"
using \<open>connected U\<close> cloU \<open>connected V\<close> cloV
by (metis Topological_Spaces.connected_continuous_image closedin_imp_subset contg continuous_on_subset fim)+
show "S = g ` U \<union> g ` V"
using T fim gfim by auto
have hom: "homeomorphism T S g f"
by (simp add: contf contg fim gf gfim homeomorphism_def)
have "closedin (top_of_set T) U" "closedin (top_of_set T) V"
by (simp_all add: cloU cloV)
then show "closedin (top_of_set S) (g ` U)"
"closedin (top_of_set S) (g ` V)"
by (blast intro: homeomorphism_imp_closed_map [OF hom])+
qed
qed
ultimately show "connected (U \<inter> V)" by metis
qed
qed
lemma homeomorphic_unicoherent_eq:
"S homeomorphic T \<Longrightarrow> (unicoherent S \<longleftrightarrow> unicoherent T)"
by (meson homeomorphic_sym homeomorphic_unicoherent)
lemma unicoherent_translation:
fixes S :: "'a::real_normed_vector set"
shows
"unicoherent (image (\<lambda>x. a + x) S) \<longleftrightarrow> unicoherent S"
using homeomorphic_translation homeomorphic_unicoherent_eq by blast
lemma unicoherent_injective_linear_image:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "linear f" "inj f"
shows "(unicoherent(f ` S) \<longleftrightarrow> unicoherent S)"
using assms homeomorphic_unicoherent_eq linear_homeomorphic_image by blast
lemma Borsukian_imp_unicoherent:
fixes U :: "'a::euclidean_space set"
assumes "Borsukian U" shows "unicoherent U"
unfolding unicoherent_def
proof clarify
fix S T
assume "connected S" "connected T" "U = S \<union> T"
and cloS: "closedin (top_of_set (S \<union> T)) S"
and cloT: "closedin (top_of_set (S \<union> T)) T"
show "connected (S \<inter> T)"
unfolding connected_closedin_eq
proof clarify
fix V W
assume "closedin (top_of_set (S \<inter> T)) V"
and "closedin (top_of_set (S \<inter> T)) W"
and VW: "V \<union> W = S \<inter> T" "V \<inter> W = {}" and "V \<noteq> {}" "W \<noteq> {}"
then have cloV: "closedin (top_of_set U) V" and cloW: "closedin (top_of_set U) W"
using \<open>U = S \<union> T\<close> cloS cloT closedin_trans by blast+
obtain q where contq: "continuous_on U q"
and q01: "\<And>x. x \<in> U \<Longrightarrow> q x \<in> {0..1::real}"
and qV: "\<And>x. x \<in> V \<Longrightarrow> q x = 0" and qW: "\<And>x. x \<in> W \<Longrightarrow> q x = 1"
by (rule Urysohn_local [OF cloV cloW \<open>V \<inter> W = {}\<close>, of 0 1])
(fastforce simp: closed_segment_eq_real_ivl)
let ?h = "\<lambda>x. if x \<in> S then exp(pi * \<i> * q x) else 1 / exp(pi * \<i> * q x)"
have eqST: "exp(pi * \<i> * q x) = 1 / exp(pi * \<i> * q x)" if "x \<in> S \<inter> T" for x
proof -
have "x \<in> V \<union> W"
using that \<open>V \<union> W = S \<inter> T\<close> by blast
with qV qW show ?thesis by force
qed
obtain g where contg: "continuous_on U g"
and circle: "g ` U \<subseteq> sphere 0 1"
and S: "\<And>x. x \<in> S \<Longrightarrow> g x = exp(pi * \<i> * q x)"
and T: "\<And>x. x \<in> T \<Longrightarrow> g x = 1 / exp(pi * \<i> * q x)"
proof
show "continuous_on U ?h"
unfolding \<open>U = S \<union> T\<close>
proof (rule continuous_on_cases_local [OF cloS cloT])
show "continuous_on S (\<lambda>x. exp (pi * \<i> * q x))"
apply (intro continuous_intros)
using \<open>U = S \<union> T\<close> continuous_on_subset contq by blast
show "continuous_on T (\<lambda>x. 1 / exp (pi * \<i> * q x))"
apply (intro continuous_intros)
using \<open>U = S \<union> T\<close> continuous_on_subset contq by auto
qed (use eqST in auto)
qed (use eqST in \<open>auto simp: norm_divide\<close>)
then obtain h where conth: "continuous_on U h" and heq: "\<And>x. x \<in> U \<Longrightarrow> g x = exp (h x)"
by (metis Borsukian_continuous_logarithm_circle assms)
obtain v w where "v \<in> V" "w \<in> W"
using \<open>V \<noteq> {}\<close> \<open>W \<noteq> {}\<close> by blast
then have vw: "v \<in> S \<inter> T" "w \<in> S \<inter> T"
using VW by auto
have iff: "2 * pi \<le> cmod (2 * of_int m * of_real pi * \<i> - 2 * of_int n * of_real pi * \<i>)
\<longleftrightarrow> 1 \<le> abs (m - n)" for m n
proof -
have "2 * pi \<le> cmod (2 * of_int m * of_real pi * \<i> - 2 * of_int n * of_real pi * \<i>)
\<longleftrightarrow> 2 * pi \<le> cmod ((2 * pi * \<i>) * (of_int m - of_int n))"
by (simp add: algebra_simps)
also have "... \<longleftrightarrow> 2 * pi \<le> 2 * pi * cmod (of_int m - of_int n)"
by (simp add: norm_mult)
also have "... \<longleftrightarrow> 1 \<le> abs (m - n)"
by simp (metis norm_of_int of_int_1_le_iff of_int_abs of_int_diff)
finally show ?thesis .
qed
have *: "\<exists>n::int. h x - (pi * \<i> * q x) = (of_int(2*n) * pi) * \<i>" if "x \<in> S" for x
using that S \<open>U = S \<union> T\<close> heq exp_eq [symmetric] by (simp add: algebra_simps)
moreover have "(\<lambda>x. h x - (pi * \<i> * q x)) constant_on S"
proof (rule continuous_discrete_range_constant [OF \<open>connected S\<close>])
have "continuous_on S h" "continuous_on S q"
using \<open>U = S \<union> T\<close> continuous_on_subset conth contq by blast+
then show "continuous_on S (\<lambda>x. h x - (pi * \<i> * q x))"
by (intro continuous_intros)
have "2*pi \<le> cmod (h y - (pi * \<i> * q y) - (h x - (pi * \<i> * q x)))"
if "x \<in> S" "y \<in> S" and ne: "h y - (pi * \<i> * q y) \<noteq> h x - (pi * \<i> * q x)" for x y
using * [OF \<open>x \<in> S\<close>] * [OF \<open>y \<in> S\<close>] ne by (auto simp: iff)
then show "\<And>x. x \<in> S \<Longrightarrow>
\<exists>e>0. \<forall>y. y \<in> S \<and> h y - (pi * \<i> * q y) \<noteq> h x - (pi * \<i> * q x) \<longrightarrow>
e \<le> cmod (h y - (pi * \<i> * q y) - (h x - (pi * \<i> * q x)))"
by (rule_tac x="2*pi" in exI) auto
qed
ultimately
obtain m where m: "\<And>x. x \<in> S \<Longrightarrow> h x - (pi * \<i> * q x) = (of_int(2*m) * pi) * \<i>"
using vw by (force simp: constant_on_def)
have *: "\<exists>n::int. h x = - (pi * \<i> * q x) + (of_int(2*n) * pi) * \<i>" if "x \<in> T" for x
unfolding exp_eq [symmetric]
using that T \<open>U = S \<union> T\<close> by (simp add: exp_minus field_simps heq [symmetric])
moreover have "(\<lambda>x. h x + (pi * \<i> * q x)) constant_on T"
proof (rule continuous_discrete_range_constant [OF \<open>connected T\<close>])
have "continuous_on T h" "continuous_on T q"
using \<open>U = S \<union> T\<close> continuous_on_subset conth contq by blast+
then show "continuous_on T (\<lambda>x. h x + (pi * \<i> * q x))"
by (intro continuous_intros)
have "2*pi \<le> cmod (h y + (pi * \<i> * q y) - (h x + (pi * \<i> * q x)))"
if "x \<in> T" "y \<in> T" and ne: "h y + (pi * \<i> * q y) \<noteq> h x + (pi * \<i> * q x)" for x y
using * [OF \<open>x \<in> T\<close>] * [OF \<open>y \<in> T\<close>] ne by (auto simp: iff)
then show "\<And>x. x \<in> T \<Longrightarrow>
\<exists>e>0. \<forall>y. y \<in> T \<and> h y + (pi * \<i> * q y) \<noteq> h x + (pi * \<i> * q x) \<longrightarrow>
e \<le> cmod (h y + (pi * \<i> * q y) - (h x + (pi * \<i> * q x)))"
by (rule_tac x="2*pi" in exI) auto
qed
ultimately
obtain n where n: "\<And>x. x \<in> T \<Longrightarrow> h x + (pi * \<i> * q x) = (of_int(2*n) * pi) * \<i>"
using vw by (force simp: constant_on_def)
show "False"
using m [of v] m [of w] n [of v] n [of w] vw
by (auto simp: algebra_simps \<open>v \<in> V\<close> \<open>w \<in> W\<close> qV qW)
qed
qed
corollary contractible_imp_unicoherent:
fixes U :: "'a::euclidean_space set"
assumes "contractible U" shows "unicoherent U"
by (simp add: Borsukian_imp_unicoherent assms contractible_imp_Borsukian)
corollary convex_imp_unicoherent:
fixes U :: "'a::euclidean_space set"
assumes "convex U" shows "unicoherent U"
by (simp add: Borsukian_imp_unicoherent assms convex_imp_Borsukian)
text\<open>If the type class constraint can be relaxed, I don't know how!\<close>
corollary unicoherent_UNIV: "unicoherent (UNIV :: 'a :: euclidean_space set)"
by (simp add: convex_imp_unicoherent)
lemma unicoherent_monotone_image_compact:
fixes T :: "'b :: t2_space set"
assumes S: "unicoherent S" "compact S" and contf: "continuous_on S f" and fim: "f ` S = T"
and conn: "\<And>y. y \<in> T \<Longrightarrow> connected (S \<inter> f -` {y})"
shows "unicoherent T"
proof
fix U V
assume UV: "connected U" "connected V" "T = U \<union> V"
and cloU: "closedin (top_of_set T) U"
and cloV: "closedin (top_of_set T) V"
moreover have "compact T"
using \<open>compact S\<close> compact_continuous_image contf fim by blast
ultimately have "closed U" "closed V"
by (auto simp: closedin_closed_eq compact_imp_closed)
let ?SUV = "(S \<inter> f -` U) \<inter> (S \<inter> f -` V)"
have UV_eq: "f ` ?SUV = U \<inter> V"
using \<open>T = U \<union> V\<close> fim by force+
have "connected (f ` ?SUV)"
proof (rule connected_continuous_image)
show "continuous_on ?SUV f"
by (meson contf continuous_on_subset inf_le1)
show "connected ?SUV"
proof (rule unicoherentD [OF \<open>unicoherent S\<close>, of "S \<inter> f -` U" "S \<inter> f -` V"])
have "\<And>C. closedin (top_of_set S) C \<Longrightarrow> closedin (top_of_set T) (f ` C)"
by (metis \<open>compact S\<close> closed_subset closedin_compact closedin_imp_subset compact_continuous_image compact_imp_closed contf continuous_on_subset fim image_mono)
then show "connected (S \<inter> f -` U)" "connected (S \<inter> f -` V)"
using UV by (auto simp: conn intro: connected_closed_monotone_preimage [OF contf fim])
show "S = (S \<inter> f -` U) \<union> (S \<inter> f -` V)"
using UV fim by blast
show "closedin (top_of_set S) (S \<inter> f -` U)"
"closedin (top_of_set S) (S \<inter> f -` V)"
by (auto simp: continuous_on_imp_closedin cloU cloV contf fim)
qed
qed
with UV_eq show "connected (U \<inter> V)"
by simp
qed
subsection\<open>Several common variants of unicoherence\<close>
lemma connected_frontier_simple:
fixes S :: "'a :: euclidean_space set"
assumes "connected S" "connected(- S)" shows "connected(frontier S)"
unfolding frontier_closures
apply (rule unicoherentD [OF unicoherent_UNIV])
apply (simp_all add: assms connected_imp_connected_closure)
by (simp add: closure_def)
lemma connected_frontier_component_complement:
fixes S :: "'a :: euclidean_space set"
assumes "connected S" and C: "C \<in> components(- S)" shows "connected(frontier C)"
apply (rule connected_frontier_simple)
using C in_components_connected apply blast
- by (metis Compl_eq_Diff_UNIV connected_UNIV assms top_greatest component_complement_connected)
+ by (metis assms component_complement_connected)
lemma connected_frontier_disjoint:
fixes S :: "'a :: euclidean_space set"
assumes "connected S" "connected T" "disjnt S T" and ST: "frontier S \<subseteq> frontier T"
shows "connected(frontier S)"
proof (cases "S = UNIV")
case True then show ?thesis
by simp
next
case False
then have "-S \<noteq> {}"
by blast
then obtain C where C: "C \<in> components(- S)" and "T \<subseteq> C"
by (metis ComplI disjnt_iff subsetI exists_component_superset \<open>disjnt S T\<close> \<open>connected T\<close>)
moreover have "frontier S = frontier C"
proof -
have "frontier C \<subseteq> frontier S"
using C frontier_complement frontier_of_components_subset by blast
moreover have "x \<in> frontier C" if "x \<in> frontier S" for x
proof -
have "x \<in> closure C"
using that unfolding frontier_def
by (metis (no_types) Diff_eq ST \<open>T \<subseteq> C\<close> closure_mono contra_subsetD frontier_def le_inf_iff that)
moreover have "x \<notin> interior C"
using that unfolding frontier_def
by (metis C Compl_eq_Diff_UNIV Diff_iff subsetD in_components_subset interior_diff interior_mono)
ultimately show ?thesis
by (auto simp: frontier_def)
qed
ultimately show ?thesis
by blast
qed
ultimately show ?thesis
using \<open>connected S\<close> connected_frontier_component_complement by auto
qed
subsection\<open>Some separation results\<close>
lemma separation_by_component_closed_pointwise:
fixes S :: "'a :: euclidean_space set"
assumes "closed S" "\<not> connected_component (- S) a b"
obtains C where "C \<in> components S" "\<not> connected_component(- C) a b"
proof (cases "a \<in> S \<or> b \<in> S")
case True
then show ?thesis
using connected_component_in componentsI that by fastforce
next
case False
obtain T where "T \<subseteq> S" "closed T" "T \<noteq> {}"
and nab: "\<not> connected_component (- T) a b"
and conn: "\<And>U. U \<subset> T \<Longrightarrow> connected_component (- U) a b"
using closed_irreducible_separator [OF assms] by metis
moreover have "connected T"
proof -
have ab: "frontier(connected_component_set (- T) a) = T" "frontier(connected_component_set (- T) b) = T"
using frontier_minimal_separating_closed_pointwise
by (metis False \<open>T \<subseteq> S\<close> \<open>closed T\<close> connected_component_sym conn connected_component_eq_empty connected_component_intermediate_subset empty_subsetI nab)+
have "connected (frontier (connected_component_set (- T) a))"
proof (rule connected_frontier_disjoint)
show "disjnt (connected_component_set (- T) a) (connected_component_set (- T) b)"
unfolding disjnt_iff
by (metis connected_component_eq connected_component_eq_empty connected_component_idemp mem_Collect_eq nab)
show "frontier (connected_component_set (- T) a) \<subseteq> frontier (connected_component_set (- T) b)"
by (simp add: ab)
qed auto
with ab \<open>closed T\<close> show ?thesis
by simp
qed
ultimately obtain C where "C \<in> components S" "T \<subseteq> C"
using exists_component_superset [of T S] by blast
then show ?thesis
by (meson Compl_anti_mono connected_component_of_subset nab that)
qed
lemma separation_by_component_closed:
fixes S :: "'a :: euclidean_space set"
assumes "closed S" "\<not> connected(- S)"
obtains C where "C \<in> components S" "\<not> connected(- C)"
proof -
obtain x y where "closed S" "x \<notin> S" "y \<notin> S" and "\<not> connected_component (- S) x y"
using assms by (auto simp: connected_iff_connected_component)
then obtain C where "C \<in> components S" "\<not> connected_component(- C) x y"
using separation_by_component_closed_pointwise by metis
then show "thesis"
apply (clarify elim!: componentsE)
by (metis Compl_iff \<open>C \<in> components S\<close> \<open>x \<notin> S\<close> \<open>y \<notin> S\<close> connected_component_eq connected_component_eq_eq connected_iff_connected_component that)
qed
lemma separation_by_Un_closed_pointwise:
fixes S :: "'a :: euclidean_space set"
assumes ST: "closed S" "closed T" "S \<inter> T = {}"
and conS: "connected_component (- S) a b" and conT: "connected_component (- T) a b"
shows "connected_component (- (S \<union> T)) a b"
proof (rule ccontr)
have "a \<notin> S" "b \<notin> S" "a \<notin> T" "b \<notin> T"
using conS conT connected_component_in by auto
assume "\<not> connected_component (- (S \<union> T)) a b"
then obtain C where "C \<in> components (S \<union> T)" and C: "\<not> connected_component(- C) a b"
using separation_by_component_closed_pointwise assms by blast
then have "C \<subseteq> S \<or> C \<subseteq> T"
proof -
have "connected C" "C \<subseteq> S \<union> T"
using \<open>C \<in> components (S \<union> T)\<close> in_components_subset by (blast elim: componentsE)+
moreover then have "C \<inter> T = {} \<or> C \<inter> S = {}"
by (metis Int_empty_right ST inf.commute connected_closed)
ultimately show ?thesis
by blast
qed
then show False
by (meson Compl_anti_mono C conS conT connected_component_of_subset)
qed
lemma separation_by_Un_closed:
fixes S :: "'a :: euclidean_space set"
assumes ST: "closed S" "closed T" "S \<inter> T = {}" and conS: "connected(- S)" and conT: "connected(- T)"
shows "connected(- (S \<union> T))"
using assms separation_by_Un_closed_pointwise
by (fastforce simp add: connected_iff_connected_component)
lemma open_unicoherent_UNIV:
fixes S :: "'a :: euclidean_space set"
assumes "open S" "open T" "connected S" "connected T" "S \<union> T = UNIV"
shows "connected(S \<inter> T)"
proof -
have "connected(- (-S \<union> -T))"
by (metis closed_Compl compl_sup compl_top_eq double_compl separation_by_Un_closed assms)
then show ?thesis
by simp
qed
lemma separation_by_component_open_aux:
fixes S :: "'a :: euclidean_space set"
assumes ST: "closed S" "closed T" "S \<inter> T = {}"
and "S \<noteq> {}" "T \<noteq> {}"
obtains C where "C \<in> components(-(S \<union> T))" "C \<noteq> {}" "frontier C \<inter> S \<noteq> {}" "frontier C \<inter> T \<noteq> {}"
proof (rule ccontr)
let ?S = "S \<union> \<Union>{C \<in> components(- (S \<union> T)). frontier C \<subseteq> S}"
let ?T = "T \<union> \<Union>{C \<in> components(- (S \<union> T)). frontier C \<subseteq> T}"
assume "\<not> thesis"
with that have *: "frontier C \<inter> S = {} \<or> frontier C \<inter> T = {}"
if C: "C \<in> components (- (S \<union> T))" "C \<noteq> {}" for C
using C by blast
have "\<exists>A B::'a set. closed A \<and> closed B \<and> UNIV \<subseteq> A \<union> B \<and> A \<inter> B = {} \<and> A \<noteq> {} \<and> B \<noteq> {}"
proof (intro exI conjI)
have "frontier (\<Union>{C \<in> components (- S \<inter> - T). frontier C \<subseteq> S}) \<subseteq> S"
apply (rule subset_trans [OF frontier_Union_subset_closure])
by (metis (no_types, lifting) SUP_least \<open>closed S\<close> closure_minimal mem_Collect_eq)
then have "frontier ?S \<subseteq> S"
by (simp add: frontier_subset_eq assms subset_trans [OF frontier_Un_subset])
then show "closed ?S"
using frontier_subset_eq by fastforce
have "frontier (\<Union>{C \<in> components (- S \<inter> - T). frontier C \<subseteq> T}) \<subseteq> T"
apply (rule subset_trans [OF frontier_Union_subset_closure])
by (metis (no_types, lifting) SUP_least \<open>closed T\<close> closure_minimal mem_Collect_eq)
then have "frontier ?T \<subseteq> T"
by (simp add: frontier_subset_eq assms subset_trans [OF frontier_Un_subset])
then show "closed ?T"
using frontier_subset_eq by fastforce
have "UNIV \<subseteq> (S \<union> T) \<union> \<Union>(components(- (S \<union> T)))"
using Union_components by blast
also have "... \<subseteq> ?S \<union> ?T"
proof -
have "C \<in> components (-(S \<union> T)) \<and> frontier C \<subseteq> S \<or>
C \<in> components (-(S \<union> T)) \<and> frontier C \<subseteq> T"
if "C \<in> components (- (S \<union> T))" "C \<noteq> {}" for C
using * [OF that] that
by clarify (metis (no_types, lifting) UnE \<open>closed S\<close> \<open>closed T\<close> closed_Un disjoint_iff_not_equal frontier_of_components_closed_complement subsetCE)
then show ?thesis
by blast
qed
finally show "UNIV \<subseteq> ?S \<union> ?T" .
have "\<Union>{C \<in> components (- (S \<union> T)). frontier C \<subseteq> S} \<union>
\<Union>{C \<in> components (- (S \<union> T)). frontier C \<subseteq> T} \<subseteq> - (S \<union> T)"
using in_components_subset by fastforce
moreover have "\<Union>{C \<in> components (- (S \<union> T)). frontier C \<subseteq> S} \<inter>
\<Union>{C \<in> components (- (S \<union> T)). frontier C \<subseteq> T} = {}"
proof -
have "C \<inter> C' = {}" if "C \<in> components (- (S \<union> T))" "frontier C \<subseteq> S"
"C' \<in> components (- (S \<union> T))" "frontier C' \<subseteq> T" for C C'
proof -
have NUN: "- S \<inter> - T \<noteq> UNIV"
using \<open>T \<noteq> {}\<close> by blast
have "C \<noteq> C'"
proof
assume "C = C'"
with that have "frontier C' \<subseteq> S \<inter> T"
by simp
also have "... = {}"
using \<open>S \<inter> T = {}\<close> by blast
finally have "C' = {} \<or> C' = UNIV"
using frontier_eq_empty by auto
then show False
using \<open>C = C'\<close> NUN that by (force simp: dest: in_components_nonempty in_components_subset)
qed
with that show ?thesis
by (simp add: components_nonoverlap [of _ "-(S \<union> T)"])
qed
then show ?thesis
by blast
qed
ultimately show "?S \<inter> ?T = {}"
using ST by blast
show "?S \<noteq> {}" "?T \<noteq> {}"
using \<open>S \<noteq> {}\<close> \<open>T \<noteq> {}\<close> by blast+
qed
then show False
by (metis Compl_disjoint connected_UNIV compl_bot_eq compl_unique connected_closedD inf_sup_absorb sup_compl_top_left1 top.extremum_uniqueI)
qed
proposition separation_by_component_open:
fixes S :: "'a :: euclidean_space set"
assumes "open S" and non: "\<not> connected(- S)"
obtains C where "C \<in> components S" "\<not> connected(- C)"
proof -
obtain T U
where "closed T" "closed U" and TU: "T \<union> U = - S" "T \<inter> U = {}" "T \<noteq> {}" "U \<noteq> {}"
using assms by (auto simp: connected_closed_set closed_def)
then obtain C where C: "C \<in> components(-(T \<union> U))" "C \<noteq> {}"
and "frontier C \<inter> T \<noteq> {}" "frontier C \<inter> U \<noteq> {}"
using separation_by_component_open_aux [OF \<open>closed T\<close> \<open>closed U\<close> \<open>T \<inter> U = {}\<close>] by force
show "thesis"
proof
show "C \<in> components S"
using C(1) TU(1) by auto
show "\<not> connected (- C)"
proof
assume "connected (- C)"
then have "connected (frontier C)"
using connected_frontier_simple [of C] \<open>C \<in> components S\<close> in_components_connected by blast
then show False
unfolding connected_closed
by (metis C(1) TU(2) \<open>closed T\<close> \<open>closed U\<close> \<open>frontier C \<inter> T \<noteq> {}\<close> \<open>frontier C \<inter> U \<noteq> {}\<close> closed_Un frontier_of_components_closed_complement inf_bot_right inf_commute)
qed
qed
qed
lemma separation_by_Un_open:
fixes S :: "'a :: euclidean_space set"
assumes "open S" "open T" "S \<inter> T = {}" and cS: "connected(-S)" and cT: "connected(-T)"
shows "connected(- (S \<union> T))"
using assms unicoherent_UNIV unfolding unicoherent_def by force
lemma nonseparation_by_component_eq:
fixes S :: "'a :: euclidean_space set"
assumes "open S \<or> closed S"
shows "((\<forall>C \<in> components S. connected(-C)) \<longleftrightarrow> connected(- S))" (is "?lhs = ?rhs")
proof
assume ?lhs with assms show ?rhs
by (meson separation_by_component_closed separation_by_component_open)
next
assume ?rhs with assms show ?lhs
using component_complement_connected by force
qed
text\<open>Another interesting equivalent of an inessential mapping into C-{0}\<close>
proposition inessential_eq_extensible:
fixes f :: "'a::euclidean_space \<Rightarrow> complex"
assumes "closed S"
shows "(\<exists>a. homotopic_with_canon (\<lambda>h. True) S (-{0}) f (\<lambda>t. a)) \<longleftrightarrow>
(\<exists>g. continuous_on UNIV g \<and> (\<forall>x \<in> S. g x = f x) \<and> (\<forall>x. g x \<noteq> 0))"
(is "?lhs = ?rhs")
proof
assume ?lhs
then obtain a where a: "homotopic_with_canon (\<lambda>h. True) S (-{0}) f (\<lambda>t. a)" ..
show ?rhs
proof (cases "S = {}")
case True
with a show ?thesis by force
next
case False
have anr: "ANR (-{0::complex})"
by (simp add: ANR_delete open_Compl open_imp_ANR)
obtain g where contg: "continuous_on UNIV g" and gim: "g ` UNIV \<subseteq> -{0}"
and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
proof (rule Borsuk_homotopy_extension_homotopic [OF _ _ continuous_on_const _ homotopic_with_symD [OF a]])
show "closedin (top_of_set UNIV) S"
using assms by auto
show "range (\<lambda>t. a) \<subseteq> - {0}"
using a homotopic_with_imp_subset2 False by blast
qed (use anr that in \<open>force+\<close>)
then show ?thesis
by force
qed
next
assume ?rhs
then obtain g where contg: "continuous_on UNIV g"
and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x" and non0: "\<And>x. g x \<noteq> 0"
by metis
obtain h k::"'a\<Rightarrow>'a" where hk: "homeomorphism (ball 0 1) UNIV h k"
using homeomorphic_ball01_UNIV homeomorphic_def by blast
then have "continuous_on (ball 0 1) (g \<circ> h)"
by (meson contg continuous_on_compose continuous_on_subset homeomorphism_cont1 top_greatest)
then obtain j where contj: "continuous_on (ball 0 1) j"
and j: "\<And>z. z \<in> ball 0 1 \<Longrightarrow> exp(j z) = (g \<circ> h) z"
by (metis (mono_tags, hide_lams) continuous_logarithm_on_ball comp_apply non0)
have [simp]: "\<And>x. x \<in> S \<Longrightarrow> h (k x) = x"
using hk homeomorphism_apply2 by blast
have "\<exists>\<zeta>. continuous_on S \<zeta>\<and> (\<forall>x\<in>S. f x = exp (\<zeta> x))"
proof (intro exI conjI ballI)
show "continuous_on S (j \<circ> k)"
proof (rule continuous_on_compose)
show "continuous_on S k"
by (meson continuous_on_subset hk homeomorphism_cont2 top_greatest)
show "continuous_on (k ` S) j"
apply (rule continuous_on_subset [OF contj])
using homeomorphism_image2 [OF hk] continuous_on_subset [OF contj] by blast
qed
show "f x = exp ((j \<circ> k) x)" if "x \<in> S" for x
proof -
have "f x = (g \<circ> h) (k x)"
by (simp add: gf that)
also have "... = exp (j (k x))"
by (metis rangeI homeomorphism_image2 [OF hk] j)
finally show ?thesis by simp
qed
qed
then show ?lhs
by (simp add: inessential_eq_continuous_logarithm)
qed
lemma inessential_on_clopen_Union:
fixes \<F> :: "'a::euclidean_space set set"
assumes T: "path_connected T"
and "\<And>S. S \<in> \<F> \<Longrightarrow> closedin (top_of_set (\<Union>\<F>)) S"
and "\<And>S. S \<in> \<F> \<Longrightarrow> openin (top_of_set (\<Union>\<F>)) S"
and hom: "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>a. homotopic_with_canon (\<lambda>x. True) S T f (\<lambda>x. a)"
obtains a where "homotopic_with_canon (\<lambda>x. True) (\<Union>\<F>) T f (\<lambda>x. a)"
proof (cases "\<Union>\<F> = {}")
case True
with that show ?thesis
by force
next
case False
then obtain C where "C \<in> \<F>" "C \<noteq> {}"
by blast
then obtain a where clo: "closedin (top_of_set (\<Union>\<F>)) C"
and ope: "openin (top_of_set (\<Union>\<F>)) C"
and "homotopic_with_canon (\<lambda>x. True) C T f (\<lambda>x. a)"
using assms by blast
with \<open>C \<noteq> {}\<close> have "f ` C \<subseteq> T" "a \<in> T"
using homotopic_with_imp_subset1 homotopic_with_imp_subset2 by blast+
have "homotopic_with_canon (\<lambda>x. True) (\<Union>\<F>) T f (\<lambda>x. a)"
proof (rule homotopic_on_clopen_Union)
show "\<And>S. S \<in> \<F> \<Longrightarrow> closedin (top_of_set (\<Union>\<F>)) S"
"\<And>S. S \<in> \<F> \<Longrightarrow> openin (top_of_set (\<Union>\<F>)) S"
by (simp_all add: assms)
show "homotopic_with_canon (\<lambda>x. True) S T f (\<lambda>x. a)" if "S \<in> \<F>" for S
proof (cases "S = {}")
case True
then show ?thesis
by auto
next
case False
then obtain b where "b \<in> S"
by blast
obtain c where c: "homotopic_with_canon (\<lambda>x. True) S T f (\<lambda>x. c)"
using \<open>S \<in> \<F>\<close> hom by blast
then have "c \<in> T"
using \<open>b \<in> S\<close> homotopic_with_imp_subset2 by blast
then have "homotopic_with_canon (\<lambda>x. True) S T (\<lambda>x. a) (\<lambda>x. c)"
using T \<open>a \<in> T\<close> homotopic_constant_maps path_connected_component
by (simp add: homotopic_constant_maps path_connected_component)
then show ?thesis
using c homotopic_with_symD homotopic_with_trans by blast
qed
qed
then show ?thesis ..
qed
proposition Janiszewski_dual:
fixes S :: "complex set"
assumes
"compact S" "compact T" "connected S" "connected T" "connected(- (S \<union> T))"
shows "connected(S \<inter> T)"
proof -
have ST: "compact (S \<union> T)"
by (simp add: assms compact_Un)
with Borsukian_imp_unicoherent [of "S \<union> T"] ST assms
show ?thesis
by (auto simp: closed_subset compact_imp_closed Borsukian_separation_compact unicoherent_def)
qed
end
diff --git a/src/HOL/Analysis/Gamma_Function.thy b/src/HOL/Analysis/Gamma_Function.thy
--- a/src/HOL/Analysis/Gamma_Function.thy
+++ b/src/HOL/Analysis/Gamma_Function.thy
@@ -1,3505 +1,3505 @@
(* Title: HOL/Analysis/Gamma_Function.thy
Author: Manuel Eberl, TU München
*)
section \<open>The Gamma Function\<close>
theory Gamma_Function
imports
Equivalence_Lebesgue_Henstock_Integration
Summation_Tests
Harmonic_Numbers
"HOL-Library.Nonpos_Ints"
"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>\<open>1 / (n::nat)^2\<close>.
\<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: field_split_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 strict_mono_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 strict_mono_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 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
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>The Euler form and the logarithmic Gamma function\<close>
text \<open>
We define the Gamma function by first defining its multiplicative inverse \<open>rGamma\<close>.
This is more convenient because \<open>rGamma\<close> is entire, which makes proofs of its
properties more convenient because one does not have to watch out for discontinuities.
(e.g. \<open>rGamma\<close> fulfils \<open>rGamma z = z * rGamma (z + 1)\<close> everywhere, whereas the \<open>\<Gamma>\<close> function
does not fulfil the analogous equation on the non-positive integers)
We define the \<open>\<Gamma>\<close> function (resp.\ its reciprocale) 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 \<open>Gamma (z + 1) = z * Gamma z\<close> follows
immediately from the definition.
\<close>
definition\<^marker>\<open>tag important\<close> 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: field_split_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
text \<open>
We now show that the series that defines the \<open>\<Gamma>\<close> 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\<^marker>\<open>tag important\<close> 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\<^marker>\<open>tag unimportant\<close> 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.
\<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\<in>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)
+ by (simp_all)
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: sum_nonneg simp: field_split_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: sum_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 sum_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 sum_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: field_split_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 sum.cong) simp_all
also note sum.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_sum sum_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: sum_distrib_left)
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)
theorem 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_sum)
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 prod.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_prod)
(subst prod_dividef [symmetric], simp_all add: field_simps)
also from m have "z * ... = (\<Prod>k=0..n. z + k) / fact n"
by (simp add: prod.atLeast0_atMost_Suc_shift prod.atLeast_Suc_atMost_Suc_shift del: prod.cl_ivl_Suc)
also have "(\<Prod>k=0..n. z + k) = pochhammer z (Suc n)"
unfolding pochhammer_prod
by (simp add: prod.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: field_split_simps powr_def)
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 sum.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 sum_subtractf sum_distrib_left divide_inverse)
also from n have "\<dots> - ?g n = 0"
- by (simp add: ln_Gamma_series_def sum_subtractf algebra_simps Ln_of_nat)
+ by (simp add: ln_Gamma_series_def sum_subtractf algebra_simps)
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: field_split_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 field_split_simps add_ac del: of_nat_Suc)
+ by (simp add: norm_divide norm_mult field_split_simps 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: field_split_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: field_split_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
subsection \<open>The Polygamma functions\<close>
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\<^marker>\<open>tag important\<close> 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\<^marker>\<open>tag important\<close> 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 sum.union_disjoint) auto
also have "(\<Sum>n=k..<m+k. f n) = (\<Sum>n=0..<m+k-k. f (n + k))"
using sum.shift_bounds_nat_ivl [of f 0 k m] by simp
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. sum f {..<a}) \<longlonglongrightarrow> lim (\<lambda>m. sum 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)
theorem Digamma_LIMSEQ:
fixes z :: "'a :: {banach,real_normed_field}"
assumes z: "z \<noteq> 0"
shows "(\<lambda>m. of_real (ln (real m)) - (\<Sum>n<m. inverse (z + of_nat n))) \<longlonglongrightarrow> Digamma z"
proof -
have "(\<lambda>n. of_real (ln (real n / (real (Suc n))))) \<longlonglongrightarrow> (of_real (ln 1) :: 'a)"
by (intro tendsto_intros LIMSEQ_n_over_Suc_n) simp_all
hence "(\<lambda>n. of_real (ln (real n / (real n + 1)))) \<longlonglongrightarrow> (0 :: 'a)" by (simp add: add_ac)
hence lim: "(\<lambda>n. of_real (ln (real n)) - of_real (ln (real n + 1))) \<longlonglongrightarrow> (0::'a)"
proof (rule Lim_transform_eventually)
show "eventually (\<lambda>n. of_real (ln (real n / (real n + 1))) =
of_real (ln (real n)) - (of_real (ln (real n + 1)) :: 'a)) at_top"
using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: ln_div)
qed
from summable_Digamma[OF z]
have "(\<lambda>n. inverse (of_nat (n+1)) - inverse (z + of_nat n))
sums (Digamma z + euler_mascheroni)"
by (simp add: Digamma_def summable_sums)
from sums_diff[OF this euler_mascheroni_sum]
have "(\<lambda>n. of_real (ln (real (Suc n) + 1)) - of_real (ln (real n + 1)) - inverse (z + of_nat n))
sums Digamma z" by (simp add: add_ac)
hence "(\<lambda>m. (\<Sum>n<m. of_real (ln (real (Suc n) + 1)) - of_real (ln (real n + 1))) -
(\<Sum>n<m. inverse (z + of_nat n))) \<longlonglongrightarrow> Digamma z"
by (simp add: sums_def sum_subtractf)
also have "(\<lambda>m. (\<Sum>n<m. of_real (ln (real (Suc n) + 1)) - of_real (ln (real n + 1)))) =
(\<lambda>m. of_real (ln (m + 1)) :: 'a)"
by (subst sum_lessThan_telescope) simp_all
finally show ?thesis by (rule Lim_transform) (insert lim, simp)
qed
theorem Polygamma_LIMSEQ:
fixes z :: "'a :: {banach,real_normed_field}"
assumes "z \<noteq> 0" and "n > 0"
shows "(\<lambda>k. inverse ((z + of_nat k)^Suc n)) sums ((-1) ^ Suc n * Polygamma n z / fact n)"
using Polygamma_converges'[OF assms(1), of "Suc n"] assms(2)
by (simp add: sums_iff Polygamma_def)
theorem 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
theorem 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)
theorem 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
theorem 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) sum_distrib_left) 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
lemma deriv_Polygamma:
assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
shows "deriv (Polygamma m) z =
Polygamma (Suc m) (z :: 'a :: {real_normed_field,euclidean_space})"
by (intro DERIV_imp_deriv has_field_derivative_Polygamma assms)
thm has_field_derivative_Polygamma
lemma higher_deriv_Polygamma:
assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
shows "(deriv ^^ n) (Polygamma m) z =
Polygamma (m + n) (z :: 'a :: {real_normed_field,euclidean_space})"
proof -
have "eventually (\<lambda>u. (deriv ^^ n) (Polygamma m) u = Polygamma (m + n) u) (nhds z)"
proof (induction n)
case (Suc n)
from Suc.IH have "eventually (\<lambda>z. eventually (\<lambda>u. (deriv ^^ n) (Polygamma m) u = Polygamma (m + n) u) (nhds z)) (nhds z)"
by (simp add: eventually_eventually)
hence "eventually (\<lambda>z. deriv ((deriv ^^ n) (Polygamma m)) z =
deriv (Polygamma (m + n)) z) (nhds z)"
by eventually_elim (intro deriv_cong_ev refl)
moreover have "eventually (\<lambda>z. z \<in> UNIV - \<int>\<^sub>\<le>\<^sub>0) (nhds z)" using assms
by (intro eventually_nhds_in_open open_Diff open_UNIV) auto
ultimately show ?case by eventually_elim (simp_all add: deriv_Polygamma)
qed simp_all
thus ?thesis by (rule eventually_nhds_x_imp_x)
qed
lemma deriv_ln_Gamma_complex:
assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0"
shows "deriv ln_Gamma z = Digamma (z :: complex)"
by (intro DERIV_imp_deriv has_field_derivative_ln_Gamma_complex assms)
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\<^marker>\<open>tag unimportant\<close> 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 * (prod 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. prod 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_prod pochhammer_Suc_prod
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)
theorem 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
field_split_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: field_split_simps rGamma_series_def add_ac)
+ by (subst pochhammer_rec') (simp_all add: field_split_simps rGamma_series_def)
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 (blast intro: 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
theorem 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)
theorem 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
theorem Gamma_fact: "Gamma (1 + of_nat n) = fact n"
by (simp add: pochhammer_fact pochhammer_Gamma of_nat_in_nonpos_Ints_iff flip: 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)
+ by (force elim!: eventually_mono simp: dist_real_def)
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)
+ by (force elim!: eventually_mono simp: dist_real_def)
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_prod) 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]
theorem 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
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]])
subsection\<^marker>\<open>tag unimportant\<close> \<open>The complex Gamma function\<close>
instantiation\<^marker>\<open>tag unimportant\<close> complex :: Gamma
begin
definition\<^marker>\<open>tag unimportant\<close> 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\<^marker>\<open>tag unimportant\<close>
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: field_split_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 (blast intro: 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
field_split_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)
+ 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 * prod 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_prod netlimit_at Let_def)
+ by (simp add: has_field_derivative_def has_derivative_def fact_prod 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. prod 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_prod pochhammer_Suc_prod 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_rGamma [holomorphic_intros]: "rGamma holomorphic_on A"
unfolding holomorphic_on_def by (auto intro!: field_differentiable_rGamma)
lemma holomorphic_rGamma' [holomorphic_intros]:
assumes "f holomorphic_on A"
shows "(\<lambda>x. rGamma (f x)) holomorphic_on A"
proof -
have "rGamma \<circ> f holomorphic_on A" using assms
by (intro holomorphic_on_compose assms holomorphic_rGamma)
thus ?thesis by (simp only: o_def)
qed
lemma analytic_rGamma: "rGamma analytic_on A"
unfolding analytic_on_def by (auto intro!: exI[of _ 1] holomorphic_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_Gamma [holomorphic_intros]: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> Gamma holomorphic_on A"
unfolding holomorphic_on_def by (auto intro!: field_differentiable_Gamma)
lemma holomorphic_Gamma' [holomorphic_intros]:
assumes "f holomorphic_on A" and "\<And>x. x \<in> A \<Longrightarrow> f x \<notin> \<int>\<^sub>\<le>\<^sub>0"
shows "(\<lambda>x. Gamma (f x)) holomorphic_on A"
proof -
have "Gamma \<circ> f holomorphic_on A" using assms
by (intro holomorphic_on_compose assms holomorphic_Gamma) auto
thus ?thesis by (simp only: o_def)
qed
lemma analytic_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_Gamma)
lemma field_differentiable_ln_Gamma_complex:
"z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> ln_Gamma field_differentiable (at (z::complex) within A)"
by (rule field_differentiable_within_subset[of _ _ UNIV])
(force simp: field_differentiable_def intro!: derivative_intros)+
lemma holomorphic_ln_Gamma [holomorphic_intros]: "A \<inter> \<real>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> ln_Gamma holomorphic_on A"
unfolding holomorphic_on_def by (auto intro!: field_differentiable_ln_Gamma_complex)
lemma analytic_ln_Gamma: "A \<inter> \<real>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> ln_Gamma analytic_on A"
by (rule analytic_on_subset[of _ "UNIV - \<real>\<^sub>\<le>\<^sub>0"], subst analytic_on_open)
(auto intro!: holomorphic_ln_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 [holomorphic_intros]: "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)
subsection\<^marker>\<open>tag unimportant\<close> \<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)
+ using n by (simp add: rGamma_series_def powr_def 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\<^marker>\<open>tag unimportant\<close> 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_field
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)
+ 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_field
simp: Polygamma_of_real rGamma_real_def [abs_def])
thus "let x = - of_nat n in (\<lambda>y. (rGamma y - rGamma x - (- 1) ^ n * prod 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_prod netlimit_at Let_def)
+ by (simp add: has_field_derivative_def has_derivative_def fact_prod 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. prod 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_prod pochhammer_Suc_prod 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)
+ with xn show ?thesis by (simp add: ln_Gamma_series_def 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 ln_Gamma_complex_conv_fact: "n > 0 \<Longrightarrow> ln_Gamma (of_nat n :: complex) = ln (fact (n - 1))"
using ln_Gamma_complex_of_real[of "real n"] Gamma_fact[of "n - 1", where 'a = real]
by (simp add: ln_Gamma_real_pos of_nat_diff Ln_of_real [symmetric])
lemma ln_Gamma_real_conv_fact: "n > 0 \<Longrightarrow> ln_Gamma (real n) = ln (fact (n - 1))"
using Gamma_fact[of "n - 1", where 'a = real]
by (simp add: ln_Gamma_real_pos of_nat_diff Ln_of_real [symmetric])
lemma Gamma_real_pos [simp, intro]: "x > (0::real) \<Longrightarrow> Gamma x > 0"
by (simp add: Gamma_real_pos_exp)
lemma Gamma_real_nonneg [simp, intro]: "x > (0::real) \<Longrightarrow> Gamma x \<ge> 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_field
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
lemma field_differentiable_ln_Gamma_real:
"x > 0 \<Longrightarrow> ln_Gamma field_differentiable (at (x::real) within A)"
by (rule field_differentiable_within_subset[of _ _ UNIV])
(auto simp: field_differentiable_def intro!: derivative_intros)+
declare has_field_derivative_ln_Gamma_real[THEN DERIV_chain2, derivative_intros]
lemma deriv_ln_Gamma_real:
assumes "z > 0"
shows "deriv ln_Gamma z = Digamma (z :: real)"
by (intro DERIV_imp_deriv has_field_derivative_ln_Gamma_real assms)
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_strict_mono: "(0::real) < x \<Longrightarrow> x < y \<Longrightarrow> Digamma x < Digamma y"
by (rule Polygamma_real_strict_mono) simp_all
lemma Digamma_real_mono: "(0::real) < x \<Longrightarrow> x \<le> y \<Longrightarrow> Digamma x \<le> Digamma y"
by (rule Polygamma_real_mono) 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
theorem 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>The uniqueness of the real Gamma function\<close>
text \<open>
The following is a proof of the Bohr--Mollerup theorem, which states that
any log-convex function \<open>G\<close> on the positive reals that fulfils \<open>G(1) = 1\<close> and
satisfies the functional equation \<open>G(x + 1) = x G(x)\<close> must be equal to the
Gamma function.
In principle, if \<open>G\<close> is a holomorphic complex function, one could then extend
this from the positive reals to the entire complex plane (minus the non-positive
integers, where the Gamma function is not defined).
\<close>
context\<^marker>\<open>tag unimportant\<close>
fixes G :: "real \<Rightarrow> real"
assumes G_1: "G 1 = 1"
assumes G_plus1: "x > 0 \<Longrightarrow> G (x + 1) = x * G x"
assumes G_pos: "x > 0 \<Longrightarrow> G x > 0"
assumes log_convex_G: "convex_on {0<..} (ln \<circ> G)"
begin
private lemma G_fact: "G (of_nat n + 1) = fact n"
using G_plus1[of "real n + 1" for n]
by (induction n) (simp_all add: G_1 G_plus1)
private definition S :: "real \<Rightarrow> real \<Rightarrow> real" where
"S x y = (ln (G y) - ln (G x)) / (y - x)"
private lemma S_eq:
"n \<ge> 2 \<Longrightarrow> S (of_nat n) (of_nat n + x) = (ln (G (real n + x)) - ln (fact (n - 1))) / x"
by (subst G_fact [symmetric]) (simp add: S_def add_ac of_nat_diff)
private lemma G_lower:
assumes x: "x > 0" and n: "n \<ge> 1"
shows "Gamma_series x n \<le> G x"
proof -
have "(ln \<circ> G) (real (Suc n)) \<le> ((ln \<circ> G) (real (Suc n) + x) -
(ln \<circ> G) (real (Suc n) - 1)) / (real (Suc n) + x - (real (Suc n) - 1)) *
(real (Suc n) - (real (Suc n) - 1)) + (ln \<circ> G) (real (Suc n) - 1)"
using x n by (intro convex_onD_Icc' convex_on_subset[OF log_convex_G]) auto
hence "S (of_nat n) (of_nat (Suc n)) \<le> S (of_nat (Suc n)) (of_nat (Suc n) + x)"
unfolding S_def using x by (simp add: field_simps)
also have "S (of_nat n) (of_nat (Suc n)) = ln (fact n) - ln (fact (n-1))"
unfolding S_def using n
by (subst (1 2) G_fact [symmetric]) (simp_all add: add_ac of_nat_diff)
also have "\<dots> = ln (fact n / fact (n-1))" by (subst ln_div) simp_all
also from n have "fact n / fact (n - 1) = n" by (cases n) simp_all
finally have "x * ln (real n) + ln (fact n) \<le> ln (G (real (Suc n) + x))"
using x n by (subst (asm) S_eq) (simp_all add: field_simps)
also have "x * ln (real n) + ln (fact n) = ln (exp (x * ln (real n)) * fact n)"
using x by (simp add: ln_mult)
finally have "exp (x * ln (real n)) * fact n \<le> G (real (Suc n) + x)" using x
by (subst (asm) ln_le_cancel_iff) (simp_all add: G_pos)
also have "G (real (Suc n) + x) = pochhammer x (Suc n) * G x"
using G_plus1[of "real (Suc n) + x" for n] G_plus1[of x] x
by (induction n) (simp_all add: pochhammer_Suc add_ac)
finally show "Gamma_series x n \<le> G x"
using x by (simp add: field_simps pochhammer_pos Gamma_series_def)
qed
private lemma G_upper:
assumes x: "x > 0" "x \<le> 1" and n: "n \<ge> 2"
shows "G x \<le> Gamma_series x n * (1 + x / real n)"
proof -
have "(ln \<circ> G) (real n + x) \<le> ((ln \<circ> G) (real n + 1) -
(ln \<circ> G) (real n)) / (real n + 1 - (real n)) *
((real n + x) - real n) + (ln \<circ> G) (real n)"
using x n by (intro convex_onD_Icc' convex_on_subset[OF log_convex_G]) auto
hence "S (of_nat n) (of_nat n + x) \<le> S (of_nat n) (of_nat n + 1)"
unfolding S_def using x by (simp add: field_simps)
also from n have "S (of_nat n) (of_nat n + 1) = ln (fact n) - ln (fact (n-1))"
by (subst (1 2) G_fact [symmetric]) (simp add: S_def add_ac of_nat_diff)
also have "\<dots> = ln (fact n / (fact (n-1)))" using n by (subst ln_div) simp_all
also from n have "fact n / fact (n - 1) = n" by (cases n) simp_all
finally have "ln (G (real n + x)) \<le> x * ln (real n) + ln (fact (n - 1))"
using x n by (subst (asm) S_eq) (simp_all add: field_simps)
also have "\<dots> = ln (exp (x * ln (real n)) * fact (n - 1))" using x
by (simp add: ln_mult)
finally have "G (real n + x) \<le> exp (x * ln (real n)) * fact (n - 1)" using x
by (subst (asm) ln_le_cancel_iff) (simp_all add: G_pos)
also have "G (real n + x) = pochhammer x n * G x"
using G_plus1[of "real n + x" for n] x
by (induction n) (simp_all add: pochhammer_Suc add_ac)
finally have "G x \<le> exp (x * ln (real n)) * fact (n- 1) / pochhammer x n"
using x by (simp add: field_simps pochhammer_pos)
also from n have "fact (n - 1) = fact n / n" by (cases n) simp_all
also have "exp (x * ln (real n)) * \<dots> / pochhammer x n =
Gamma_series x n * (1 + x / real n)" using n x
by (simp add: Gamma_series_def divide_simps pochhammer_Suc)
finally show ?thesis .
qed
private lemma G_eq_Gamma_aux:
assumes x: "x > 0" "x \<le> 1"
shows "G x = Gamma x"
proof (rule antisym)
show "G x \<ge> Gamma x"
proof (rule tendsto_upperbound)
from G_lower[of x] show "eventually (\<lambda>n. Gamma_series x n \<le> G x) sequentially"
using x by (auto intro: eventually_mono[OF eventually_ge_at_top[of "1::nat"]])
qed (simp_all add: Gamma_series_LIMSEQ)
next
show "G x \<le> Gamma x"
proof (rule tendsto_lowerbound)
have "(\<lambda>n. Gamma_series x n * (1 + x / real n)) \<longlonglongrightarrow> Gamma x * (1 + 0)"
by (rule tendsto_intros real_tendsto_divide_at_top
Gamma_series_LIMSEQ filterlim_real_sequentially)+
thus "(\<lambda>n. Gamma_series x n * (1 + x / real n)) \<longlonglongrightarrow> Gamma x" by simp
next
from G_upper[of x] show "eventually (\<lambda>n. Gamma_series x n * (1 + x / real n) \<ge> G x) sequentially"
using x by (auto intro: eventually_mono[OF eventually_ge_at_top[of "2::nat"]])
qed simp_all
qed
theorem Gamma_pos_real_unique:
assumes x: "x > 0"
shows "G x = Gamma x"
proof -
have G_eq: "G (real n + x) = Gamma (real n + x)" if "x \<in> {0<..1}" for n x using that
proof (induction n)
case (Suc n)
from Suc have "x + real n > 0" by simp
hence "x + real n \<notin> \<int>\<^sub>\<le>\<^sub>0" by auto
with Suc show ?case using G_plus1[of "real n + x"] Gamma_plus1[of "real n + x"]
by (auto simp: add_ac)
qed (simp_all add: G_eq_Gamma_aux)
show ?thesis
proof (cases "frac x = 0")
case True
hence "x = of_int (floor x)" by (simp add: frac_def)
with x have x_eq: "x = of_nat (nat (floor x) - 1) + 1" by simp
show ?thesis by (subst (1 2) x_eq, rule G_eq) simp_all
next
case False
from assms have x_eq: "x = of_nat (nat (floor x)) + frac x"
by (simp add: frac_def)
have frac_le_1: "frac x \<le> 1" unfolding frac_def by linarith
show ?thesis
by (subst (1 2) x_eq, rule G_eq, insert False frac_le_1) simp_all
qed
qed
end
subsection \<open>The 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)
theorem 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
theorem 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
theorem 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: field_split_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 "(*)2" "2*z"]
by (intro tendsto_intros Gamma_series'_LIMSEQ)
(simp_all add: o_def strict_mono_def Gamma_eq_zero_iff)
ultimately have "?h \<longlonglongrightarrow> ?powr 2 (2*z) * Gamma z * Gamma (z+1/2) / Gamma (2*z)"
by (blast intro: 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: field_split_simps Gamma_eq_zero_iff ring_distribs exp_diff exp_of_real ac_simps)
+ by (simp add: field_split_simps Gamma_eq_zero_iff ring_distribs exp_diff exp_of_real)
qed
text \<open>
The following lemma is somewhat annoying. With a little bit of complex analysis
(Cauchy's integral theorem, to be exact), this would be completely trivial. However,
we want to avoid depending on the complex analysis session at this point, so we prove it
the hard way.
\<close>
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)
+ from that have t: "t \<in> \<int> \<longleftrightarrow> t = 0" by (auto elim!: Ints_cases)
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)
+ using sums[of t] that by (simp add: sums_iff)
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)
+ by (auto elim!: eventually_mono)
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)
+ with z have z'': "z \<notin> \<int>\<^sub>\<le>\<^sub>0" "z \<notin> \<int>" by (auto elim!: Ints_cases)
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\<in>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)
+ from h'_eq[of t] t have "h' t = (h' (t/2) + h' ((t+1)/2)) / 4" by (simp)
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)
+ qed (insert bdd, auto)
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')
then obtain u where u: "u \<in> closed_segment 0 1" "Re (g 1) - Re (g 0) = Re (h u * g u)"
by auto
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)
+ then obtain c' where c: "\<And>z. g z = c'" by (force)
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 field_split_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 field_split_simps mult_ac)
+ using rGamma_reflection_complex'[of z] by (force simp add: Gamma_def field_split_simps)
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
theorem 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\<^marker>\<open>tag unimportant\<close> \<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: field_split_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: field_split_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 flip: 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_sum [symmetric]) (simp_all add: sum_distrib_left)
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_prod [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 prod.cong) (simp_all add: field_split_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: prod.nat_ivl_Suc' field_split_simps)
finally show ?thesis ..
qed
theorem 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 z by (auto simp: field_split_simps Gamma_series_def ring_distribs exp_diff ln_div add_ac
+ using z by (auto simp: field_split_simps Gamma_series_def ring_distribs exp_diff ln_div
intro: eventually_mono eventually_gt_at_top[of "0::nat"] 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 intro: 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 prod.distrib
prod_inversef[symmetric] divide_inverse)
also have "(\<Prod>k=1..n. (1 + z / of_nat k)) = pochhammer (z + 1) n / fact n"
proof (cases n)
case (Suc n')
then show ?thesis
unfolding pochhammer_prod fact_prod
by (simp add: atLeastLessThanSuc_atLeastAtMost field_simps prod_dividef
prod.atLeast_Suc_atMost_Suc_shift del: prod.cl_ivl_Suc)
qed auto
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\<^marker>\<open>tag unimportant\<close>
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)
theorem 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
sum.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_sum 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 field_split_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 prod_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 field_split_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 field_split_simps exp_of_real [symmetric] exp_minus)
+ by (simp add: binomial_gbinomial Gamma_def field_split_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)
+ by (simp_all add: Beta_pole1 one_minus_of_nat_in_nonpos_Ints_iff)
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
theorem 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: field_split_simps mult_ac add_ac)
+ using Gamma_plus1[of "z+1"] by (auto simp add: field_split_simps)
finally show ?thesis .
qed
subsubsection \<open>Integral form\<close>
lemma integrable_on_powr_from_0':
assumes a: "a > (-1::real)" and c: "c \<ge> 0"
shows "(\<lambda>x. x powr a) integrable_on {0<..c}"
proof -
from c have *: "{0<..c} - {0..c} = {}" "{0..c} - {0<..c} = {0}" by auto
show ?thesis
by (rule integrable_spike_set [OF integrable_on_powr_from_0[OF a c]]) (simp_all add: *)
qed
lemma absolutely_integrable_Gamma_integral:
assumes "Re z > 0" "a > 0"
shows "(\<lambda>t. complex_of_real t powr (z - 1) / of_real (exp (a * t)))
absolutely_integrable_on {0<..}" (is "?f absolutely_integrable_on _")
proof -
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 "a/2"] and \<open>a > 0\<close>
have "eventually (\<lambda>x. (Re z - 1) * ln x / x < a/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 < a/2 \<and> x > 0"
by (auto simp: eventually_at_top_linorder)
hence "x0 > 0" by simp
have "x powr (Re z - 1) / exp (a * x) < exp (-(a/2) * x)" if "x \<ge> x0" for x
proof -
from that and \<open>\<forall>x\<ge>x0. _\<close> have x: "(Re z - 1) * ln x / x < a / 2" "x > 0" by auto
have "x powr (Re z - 1) = exp ((Re z - 1) * ln x)"
using \<open>x > 0\<close> by (simp add: powr_def)
also from x have "(Re z - 1) * ln x < (a * x) / 2" by (simp add: field_simps)
finally show ?thesis by (simp add: field_simps exp_add [symmetric])
qed
note x0 = \<open>x0 > 0\<close> this
have "?f absolutely_integrable_on ({0<..x0} \<union> {x0..})"
proof (rule set_integrable_Un)
show "?f absolutely_integrable_on {0<..x0}"
unfolding set_integrable_def
proof (rule Bochner_Integration.integrable_bound [OF _ _ AE_I2])
show "integrable lebesgue (\<lambda>x. indicat_real {0<..x0} x *\<^sub>R x powr (Re z - 1))"
using x0(1) assms
by (intro nonnegative_absolutely_integrable_1 [unfolded set_integrable_def] integrable_on_powr_from_0') auto
show "(\<lambda>x. indicat_real {0<..x0} x *\<^sub>R (x powr (z - 1) / exp (a * x))) \<in> borel_measurable lebesgue"
by (intro measurable_completion)
(auto intro!: borel_measurable_continuous_on_indicator continuous_intros)
fix x :: real
have "x powr (Re z - 1) / exp (a * x) \<le> x powr (Re z - 1) / 1" if "x \<ge> 0"
using that assms by (intro divide_left_mono) auto
thus "norm (indicator {0<..x0} x *\<^sub>R ?f x) \<le>
norm (indicator {0<..x0} x *\<^sub>R x powr (Re z - 1))"
by (simp_all add: norm_divide norm_powr_real_powr indicator_def)
qed
next
show "?f absolutely_integrable_on {x0..}"
unfolding set_integrable_def
proof (rule Bochner_Integration.integrable_bound [OF _ _ AE_I2])
show "integrable lebesgue (\<lambda>x. indicat_real {x0..} x *\<^sub>R exp (- (a / 2) * x))" using assms
by (intro nonnegative_absolutely_integrable_1 [unfolded set_integrable_def] integrable_on_exp_minus_to_infinity) auto
show "(\<lambda>x. indicat_real {x0..} x *\<^sub>R (x powr (z - 1) / exp (a * x))) \<in> borel_measurable lebesgue" using x0(1)
by (intro measurable_completion)
(auto intro!: borel_measurable_continuous_on_indicator continuous_intros)
fix x :: real
show "norm (indicator {x0..} x *\<^sub>R ?f x) \<le>
norm (indicator {x0..} x *\<^sub>R exp (-(a/2) * x))" using x0
by (auto simp: norm_divide norm_powr_real_powr indicator_def less_imp_le)
qed
qed auto
also have "{0<..x0} \<union> {x0..} = {0<..}" using x0(1) by auto
finally show ?thesis .
qed
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_Un') (insert c, auto simp: max_def)
qed
theorem 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_field)
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_field)
show "(?g has_vector_derivative ?g' t) (at t)"
by (rule has_vector_derivative_real_field 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: field_split_simps pochhammer_rec
prod.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: 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 exp_of_real)
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)
+ by eventually_elim (simp add: powr_def algebra_simps 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 (blast intro: 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: field_split_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
lemma absolutely_integrable_Gamma_integral':
assumes "Re z > 0"
shows "(\<lambda>t. complex_of_real t powr (z - 1) / of_real (exp t)) absolutely_integrable_on {0<..}"
using absolutely_integrable_Gamma_integral [OF assms zero_less_one] by simp
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 "((\<lambda>t. of_real t powr (z - 1) / of_real (exp t)) has_integral Gamma z) {0..}"
by (rule Gamma_integral_complex) fact+
hence "((\<lambda>t. if t \<in> {0<..} then of_real t powr (z - 1) / of_real (exp t) else 0)
has_integral Gamma z) {0..}"
by (rule has_integral_spike [of "{0}", rotated 2]) auto
also have "?this = ?thesis"
by (subst has_integral_restrict) auto
finally show ?thesis .
qed
lemma Gamma_conv_nn_integral_real:
assumes "s > (0::real)"
shows "Gamma s = nn_integral lborel (\<lambda>t. ennreal (indicator {0..} t * t powr (s - 1) / exp t))"
using nn_integral_has_integral_lebesgue[OF _ Gamma_integral_real[OF assms]] by simp
lemma integrable_Beta:
assumes "a > 0" "b > (0::real)"
shows "set_integrable lborel {0..1} (\<lambda>t. t powr (a - 1) * (1 - t) powr (b - 1))"
proof -
define C where "C = max 1 ((1/2) powr (b - 1))"
define D where "D = max 1 ((1/2) powr (a - 1))"
have C: "(1 - x) powr (b - 1) \<le> C" if "x \<in> {0..1/2}" for x
proof (cases "b < 1")
case False
with that have "(1 - x) powr (b - 1) \<le> (1 powr (b - 1))" by (intro powr_mono2) auto
thus ?thesis by (auto simp: C_def)
qed (insert that, auto simp: max.coboundedI1 max.coboundedI2 powr_mono2' powr_mono2 C_def)
have D: "x powr (a - 1) \<le> D" if "x \<in> {1/2..1}" for x
proof (cases "a < 1")
case False
with that have "x powr (a - 1) \<le> (1 powr (a - 1))" by (intro powr_mono2) auto
thus ?thesis by (auto simp: D_def)
next
case True
qed (insert that, auto simp: max.coboundedI1 max.coboundedI2 powr_mono2' powr_mono2 D_def)
have [simp]: "C \<ge> 0" "D \<ge> 0" by (simp_all add: C_def D_def)
have I1: "set_integrable lborel {0..1/2} (\<lambda>t. t powr (a - 1) * (1 - t) powr (b - 1))"
unfolding set_integrable_def
proof (rule Bochner_Integration.integrable_bound[OF _ _ AE_I2])
have "(\<lambda>t. t powr (a - 1)) integrable_on {0..1/2}"
by (rule integrable_on_powr_from_0) (use assms in auto)
hence "(\<lambda>t. t powr (a - 1)) absolutely_integrable_on {0..1/2}"
by (subst absolutely_integrable_on_iff_nonneg) auto
from integrable_mult_right[OF this [unfolded set_integrable_def], of C]
show "integrable lborel (\<lambda>x. indicat_real {0..1/2} x *\<^sub>R (C * x powr (a - 1)))"
by (subst (asm) integrable_completion) (auto simp: mult_ac)
next
fix x :: real
have "x powr (a - 1) * (1 - x) powr (b - 1) \<le> x powr (a - 1) * C" if "x \<in> {0..1/2}"
using that by (intro mult_left_mono powr_mono2 C) auto
thus "norm (indicator {0..1/2} x *\<^sub>R (x powr (a - 1) * (1 - x) powr (b - 1))) \<le>
norm (indicator {0..1/2} x *\<^sub>R (C * x powr (a - 1)))"
by (auto simp: indicator_def abs_mult mult_ac)
qed (auto intro!: AE_I2 simp: indicator_def)
have I2: "set_integrable lborel {1/2..1} (\<lambda>t. t powr (a - 1) * (1 - t) powr (b - 1))"
unfolding set_integrable_def
proof (rule Bochner_Integration.integrable_bound[OF _ _ AE_I2])
have "(\<lambda>t. t powr (b - 1)) integrable_on {0..1/2}"
by (rule integrable_on_powr_from_0) (use assms in auto)
hence "(\<lambda>t. t powr (b - 1)) integrable_on (cbox 0 (1/2))" by simp
from integrable_affinity[OF this, of "-1" 1]
have "(\<lambda>t. (1 - t) powr (b - 1)) integrable_on {1/2..1}" by simp
hence "(\<lambda>t. (1 - t) powr (b - 1)) absolutely_integrable_on {1/2..1}"
by (subst absolutely_integrable_on_iff_nonneg) auto
from integrable_mult_right[OF this [unfolded set_integrable_def], of D]
show "integrable lborel (\<lambda>x. indicat_real {1/2..1} x *\<^sub>R (D * (1 - x) powr (b - 1)))"
by (subst (asm) integrable_completion) (auto simp: mult_ac)
next
fix x :: real
have "x powr (a - 1) * (1 - x) powr (b - 1) \<le> D * (1 - x) powr (b - 1)" if "x \<in> {1/2..1}"
using that by (intro mult_right_mono powr_mono2 D) auto
thus "norm (indicator {1/2..1} x *\<^sub>R (x powr (a - 1) * (1 - x) powr (b - 1))) \<le>
norm (indicator {1/2..1} x *\<^sub>R (D * (1 - x) powr (b - 1)))"
by (auto simp: indicator_def abs_mult mult_ac)
qed (auto intro!: AE_I2 simp: indicator_def)
have "set_integrable lborel ({0..1/2} \<union> {1/2..1}) (\<lambda>t. t powr (a - 1) * (1 - t) powr (b - 1))"
by (intro set_integrable_Un I1 I2) auto
also have "{0..1/2} \<union> {1/2..1} = {0..(1::real)}" by auto
finally show ?thesis .
qed
lemma integrable_Beta':
assumes "a > 0" "b > (0::real)"
shows "(\<lambda>t. t powr (a - 1) * (1 - t) powr (b - 1)) integrable_on {0..1}"
using integrable_Beta[OF assms] by (rule set_borel_integral_eq_integral)
theorem has_integral_Beta_real:
assumes a: "a > 0" and b: "b > (0 :: real)"
shows "((\<lambda>t. t powr (a - 1) * (1 - t) powr (b - 1)) has_integral Beta a b) {0..1}"
proof -
define B where "B = integral {0..1} (\<lambda>x. x powr (a - 1) * (1 - x) powr (b - 1))"
have [simp]: "B \<ge> 0" unfolding B_def using a b
by (intro integral_nonneg integrable_Beta') auto
from a b have "ennreal (Gamma a * Gamma b) =
(\<integral>\<^sup>+ t. ennreal (indicator {0..} t * t powr (a - 1) / exp t) \<partial>lborel) *
(\<integral>\<^sup>+ t. ennreal (indicator {0..} t * t powr (b - 1) / exp t) \<partial>lborel)"
by (subst ennreal_mult') (simp_all add: Gamma_conv_nn_integral_real)
also have "\<dots> = (\<integral>\<^sup>+t. \<integral>\<^sup>+u. ennreal (indicator {0..} t * t powr (a - 1) / exp t) *
ennreal (indicator {0..} u * u powr (b - 1) / exp u) \<partial>lborel \<partial>lborel)"
by (simp add: nn_integral_cmult nn_integral_multc)
also have "\<dots> = (\<integral>\<^sup>+t. \<integral>\<^sup>+u. ennreal (indicator ({0..}\<times>{0..}) (t,u) * t powr (a - 1) * u powr (b - 1)
/ exp (t + u)) \<partial>lborel \<partial>lborel)"
by (intro nn_integral_cong)
(auto simp: indicator_def divide_ennreal ennreal_mult' [symmetric] exp_add)
also have "\<dots> = (\<integral>\<^sup>+t. \<integral>\<^sup>+u. ennreal (indicator ({0..}\<times>{t..}) (t,u) * t powr (a - 1) *
(u - t) powr (b - 1) / exp u) \<partial>lborel \<partial>lborel)"
proof (rule nn_integral_cong, goal_cases)
case (1 t)
have "(\<integral>\<^sup>+u. ennreal (indicator ({0..}\<times>{0..}) (t,u) * t powr (a - 1) *
u powr (b - 1) / exp (t + u)) \<partial>distr lborel borel ((+) (-t))) =
(\<integral>\<^sup>+u. ennreal (indicator ({0..}\<times>{t..}) (t,u) * t powr (a - 1) *
(u - t) powr (b - 1) / exp u) \<partial>lborel)"
by (subst nn_integral_distr) (auto intro!: nn_integral_cong simp: indicator_def)
thus ?case by (subst (asm) lborel_distr_plus)
qed
also have "\<dots> = (\<integral>\<^sup>+u. \<integral>\<^sup>+t. ennreal (indicator ({0..}\<times>{t..}) (t,u) * t powr (a - 1) *
(u - t) powr (b - 1) / exp u) \<partial>lborel \<partial>lborel)"
by (subst lborel_pair.Fubini')
(auto simp: case_prod_unfold indicator_def cong: measurable_cong_sets)
also have "\<dots> = (\<integral>\<^sup>+u. \<integral>\<^sup>+t. ennreal (indicator {0..u} t * t powr (a - 1) * (u - t) powr (b - 1)) *
ennreal (indicator {0..} u / exp u) \<partial>lborel \<partial>lborel)"
by (intro nn_integral_cong) (auto simp: indicator_def ennreal_mult' [symmetric])
also have "\<dots> = (\<integral>\<^sup>+u. (\<integral>\<^sup>+t. ennreal (indicator {0..u} t * t powr (a - 1) * (u - t) powr (b - 1))
\<partial>lborel) * ennreal (indicator {0..} u / exp u) \<partial>lborel)"
by (subst nn_integral_multc [symmetric]) auto
also have "\<dots> = (\<integral>\<^sup>+u. (\<integral>\<^sup>+t. ennreal (indicator {0..u} t * t powr (a - 1) * (u - t) powr (b - 1))
\<partial>lborel) * ennreal (indicator {0<..} u / exp u) \<partial>lborel)"
by (intro nn_integral_cong_AE eventually_mono[OF AE_lborel_singleton[of 0]])
(auto simp: indicator_def)
also have "\<dots> = (\<integral>\<^sup>+u. ennreal B * ennreal (indicator {0..} u / exp u * u powr (a + b - 1)) \<partial>lborel)"
proof (intro nn_integral_cong, goal_cases)
case (1 u)
show ?case
proof (cases "u > 0")
case True
have "(\<integral>\<^sup>+t. ennreal (indicator {0..u} t * t powr (a - 1) * (u - t) powr (b - 1)) \<partial>lborel) =
(\<integral>\<^sup>+t. ennreal (indicator {0..1} t * (u * t) powr (a - 1) * (u - u * t) powr (b - 1))
\<partial>distr lborel borel ((*) (1 / u)))" (is "_ = nn_integral _ ?f")
using True
by (subst nn_integral_distr) (auto simp: indicator_def field_simps intro!: nn_integral_cong)
also have "distr lborel borel ((*) (1 / u)) = density lborel (\<lambda>_. u)"
using \<open>u > 0\<close> by (subst lborel_distr_mult) auto
also have "nn_integral \<dots> ?f = (\<integral>\<^sup>+x. ennreal (indicator {0..1} x * (u * (u * x) powr (a - 1) *
(u * (1 - x)) powr (b - 1))) \<partial>lborel)" using \<open>u > 0\<close>
by (subst nn_integral_density) (auto simp: ennreal_mult' [symmetric] algebra_simps)
also have "\<dots> = (\<integral>\<^sup>+x. ennreal (u powr (a + b - 1)) *
ennreal (indicator {0..1} x * x powr (a - 1) *
(1 - x) powr (b - 1)) \<partial>lborel)" using \<open>u > 0\<close> a b
by (intro nn_integral_cong)
(auto simp: indicator_def powr_mult powr_add powr_diff mult_ac ennreal_mult' [symmetric])
also have "\<dots> = ennreal (u powr (a + b - 1)) *
(\<integral>\<^sup>+x. ennreal (indicator {0..1} x * x powr (a - 1) *
(1 - x) powr (b - 1)) \<partial>lborel)"
by (subst nn_integral_cmult) auto
also have "((\<lambda>x. x powr (a - 1) * (1 - x) powr (b - 1)) has_integral
integral {0..1} (\<lambda>x. x powr (a - 1) * (1 - x) powr (b - 1))) {0..1}"
using a b by (intro integrable_integral integrable_Beta')
from nn_integral_has_integral_lebesgue[OF _ this] a b
have "(\<integral>\<^sup>+x. ennreal (indicator {0..1} x * x powr (a - 1) *
(1 - x) powr (b - 1)) \<partial>lborel) = B" by (simp add: mult_ac B_def)
finally show ?thesis using \<open>u > 0\<close> by (simp add: ennreal_mult' [symmetric] mult_ac)
qed auto
qed
also have "\<dots> = ennreal B * ennreal (Gamma (a + b))"
using a b by (subst nn_integral_cmult) (auto simp: Gamma_conv_nn_integral_real)
also have "\<dots> = ennreal (B * Gamma (a + b))"
by (subst (1 2) mult.commute, intro ennreal_mult' [symmetric]) (use a b in auto)
finally have "B = Beta a b" using a b Gamma_real_pos[of "a + b"]
by (subst (asm) ennreal_inj) (auto simp: field_simps Beta_def Gamma_eq_zero_iff)
moreover have "(\<lambda>t. t powr (a - 1) * (1 - t) powr (b - 1)) integrable_on {0..1}"
by (intro integrable_Beta' a b)
ultimately show ?thesis by (simp add: has_integral_iff B_def)
qed
subsection \<open>The Weierstraß product formula for the sine\<close>
theorem 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 prod.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 prod.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: field_split_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: field_split_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
theorem wallis: "(\<lambda>n. \<Prod>k=1..n. (4*real k^2) / (4*real k^2 - 1)) \<longlonglongrightarrow> pi / 2"
proof -
from tendsto_inverse[OF tendsto_mult[OF
sin_product_formula_real[of "1/2"] tendsto_const[of "2/pi"]]]
have "(\<lambda>n. (\<Prod>k=1..n. inverse (1 - (1/2)\<^sup>2 / (real k)\<^sup>2))) \<longlonglongrightarrow> pi/2"
by (simp add: prod_inversef [symmetric])
also have "(\<lambda>n. (\<Prod>k=1..n. inverse (1 - (1/2)\<^sup>2 / (real k)\<^sup>2))) =
(\<lambda>n. (\<Prod>k=1..n. (4*real k^2)/(4*real k^2 - 1)))"
by (intro ext prod.cong refl) (simp add: field_split_simps)
finally show ?thesis .
qed
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: prod.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)
+ from x have "x^2 < 1" by (auto simp: 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 prod_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_prod 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: field_split_simps eval_nat_numeral power_mult_distrib mult_ac)
+ by (simp add: field_split_simps eval_nat_numeral)
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 field_split_simps eval_nat_numeral)
+ show ?thesis by (simp add: field_split_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
diff --git a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy
--- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy
+++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy
@@ -1,7665 +1,7665 @@
(* Author: John Harrison
Author: Robert Himmelmann, TU Muenchen (Translation from HOL light)
Huge cleanup by LCP
*)
section \<open>Henstock-Kurzweil Gauge Integration in Many Dimensions\<close>
theory Henstock_Kurzweil_Integration
imports
Lebesgue_Measure Tagged_Division
begin
lemma norm_diff2: "\<lbrakk>y = y1 + y2; x = x1 + x2; e = e1 + e2; norm(y1 - x1) \<le> e1; norm(y2 - x2) \<le> e2\<rbrakk>
\<Longrightarrow> norm(y-x) \<le> e"
using norm_triangle_mono [of "y1 - x1" "e1" "y2 - x2" "e2"]
by (simp add: add_diff_add)
lemma setcomp_dot1: "{z. P (z \<bullet> (i,0))} = {(x,y). P(x \<bullet> i)}"
by auto
lemma setcomp_dot2: "{z. P (z \<bullet> (0,i))} = {(x,y). P(y \<bullet> i)}"
by auto
lemma Sigma_Int_Paircomp1: "(Sigma A B) \<inter> {(x, y). P x} = Sigma (A \<inter> {x. P x}) B"
by blast
lemma Sigma_Int_Paircomp2: "(Sigma A B) \<inter> {(x, y). P y} = Sigma A (\<lambda>z. B z \<inter> {y. P y})"
by blast
(* END MOVE *)
subsection \<open>Content (length, area, volume...) of an interval\<close>
abbreviation content :: "'a::euclidean_space set \<Rightarrow> real"
where "content s \<equiv> measure lborel s"
lemma content_cbox_cases:
"content (cbox a b) = (if \<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i then prod (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis else 0)"
by (simp add: measure_lborel_cbox_eq inner_diff)
lemma content_cbox: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow> content (cbox a b) = (\<Prod>i\<in>Basis. b\<bullet>i - a\<bullet>i)"
unfolding content_cbox_cases by simp
lemma content_cbox': "cbox a b \<noteq> {} \<Longrightarrow> content (cbox a b) = (\<Prod>i\<in>Basis. b\<bullet>i - a\<bullet>i)"
by (simp add: box_ne_empty inner_diff)
lemma content_cbox_if: "content (cbox a b) = (if cbox a b = {} then 0 else \<Prod>i\<in>Basis. b\<bullet>i - a\<bullet>i)"
by (simp add: content_cbox')
lemma content_cbox_cart:
"cbox a b \<noteq> {} \<Longrightarrow> content(cbox a b) = prod (\<lambda>i. b$i - a$i) UNIV"
by (simp add: content_cbox_if Basis_vec_def cart_eq_inner_axis axis_eq_axis prod.UNION_disjoint)
lemma content_cbox_if_cart:
"content(cbox a b) = (if cbox a b = {} then 0 else prod (\<lambda>i. b$i - a$i) UNIV)"
by (simp add: content_cbox_cart)
lemma content_division_of:
assumes "K \<in> \<D>" "\<D> division_of S"
shows "content K = (\<Prod>i \<in> Basis. interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)"
proof -
obtain a b where "K = cbox a b"
using cbox_division_memE assms by metis
then show ?thesis
using assms by (force simp: division_of_def content_cbox')
qed
lemma content_real: "a \<le> b \<Longrightarrow> content {a..b} = b - a"
by simp
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: "content {a} = 0"
by simp
lemma content_unit[iff]: "content (cbox 0 (One::'a::euclidean_space)) = 1"
by simp
lemma content_pos_le [iff]: "0 \<le> content X"
by simp
corollary\<^marker>\<open>tag unimportant\<close> content_nonneg [simp]: "\<not> content (cbox a b) < 0"
using not_le by blast
lemma content_pos_lt: "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i \<Longrightarrow> 0 < content (cbox a b)"
by (auto simp: less_imp_le inner_diff box_eq_empty intro!: prod_pos)
lemma content_eq_0: "content (cbox a b) = 0 \<longleftrightarrow> (\<exists>i\<in>Basis. b\<bullet>i \<le> a\<bullet>i)"
by (auto simp: content_cbox_cases not_le intro: less_imp_le antisym eq_refl)
lemma content_eq_0_interior: "content (cbox a b) = 0 \<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)"
by (auto simp add: content_cbox_cases less_le prod_nonneg)
lemma content_empty [simp]: "content {} = 0"
by simp
lemma content_real_if [simp]: "content {a..b} = (if a \<le> b then b - a else 0)"
by (simp add: content_real)
lemma content_subset: "cbox a b \<subseteq> cbox c d \<Longrightarrow> content (cbox a b) \<le> content (cbox c d)"
unfolding measure_def
by (intro enn2real_mono emeasure_mono) (auto simp: emeasure_lborel_cbox_eq)
lemma content_lt_nz: "0 < content (cbox a b) \<longleftrightarrow> content (cbox a b) \<noteq> 0"
unfolding content_pos_lt_eq content_eq_0 unfolding not_ex not_le by fastforce
lemma content_Pair: "content (cbox (a,c) (b,d)) = content (cbox a b) * content (cbox c d)"
unfolding measure_lborel_cbox_eq Basis_prod_def
apply (subst prod.union_disjoint)
apply (auto simp: bex_Un ball_Un)
apply (subst (1 2) prod.reindex_nontrivial)
apply auto
done
lemma content_cbox_pair_eq0_D:
"content (cbox (a,c) (b,d)) = 0 \<Longrightarrow> content (cbox a b) = 0 \<or> content (cbox c d) = 0"
by (simp add: content_Pair)
lemma content_cbox_plus:
fixes x :: "'a::euclidean_space"
shows "content(cbox x (x + h *\<^sub>R One)) = (if h \<ge> 0 then h ^ DIM('a) else 0)"
by (simp add: algebra_simps content_cbox_if box_eq_empty)
lemma content_0_subset: "content(cbox a b) = 0 \<Longrightarrow> s \<subseteq> cbox a b \<Longrightarrow> content s = 0"
using emeasure_mono[of s "cbox a b" lborel]
by (auto simp: measure_def enn2real_eq_0_iff emeasure_lborel_cbox_eq)
lemma content_ball_pos:
assumes "r > 0"
shows "content (ball c r) > 0"
proof -
from rational_boxes[OF assms, of c] obtain a b where ab: "c \<in> box a b" "box a b \<subseteq> ball c r"
by auto
from ab have "0 < content (box a b)"
by (subst measure_lborel_box_eq) (auto intro!: prod_pos simp: algebra_simps box_def)
have "emeasure lborel (box a b) \<le> emeasure lborel (ball c r)"
using ab by (intro emeasure_mono) auto
also have "emeasure lborel (box a b) = ennreal (content (box a b))"
using emeasure_lborel_box_finite[of a b] by (intro emeasure_eq_ennreal_measure) auto
also have "emeasure lborel (ball c r) = ennreal (content (ball c r))"
using emeasure_lborel_ball_finite[of c r] by (intro emeasure_eq_ennreal_measure) auto
finally show ?thesis
using \<open>content (box a b) > 0\<close> by simp
qed
lemma content_cball_pos:
assumes "r > 0"
shows "content (cball c r) > 0"
proof -
from rational_boxes[OF assms, of c] obtain a b where ab: "c \<in> box a b" "box a b \<subseteq> ball c r"
by auto
from ab have "0 < content (box a b)"
by (subst measure_lborel_box_eq) (auto intro!: prod_pos simp: algebra_simps box_def)
have "emeasure lborel (box a b) \<le> emeasure lborel (ball c r)"
using ab by (intro emeasure_mono) auto
also have "\<dots> \<le> emeasure lborel (cball c r)"
by (intro emeasure_mono) auto
also have "emeasure lborel (box a b) = ennreal (content (box a b))"
using emeasure_lborel_box_finite[of a b] by (intro emeasure_eq_ennreal_measure) auto
also have "emeasure lborel (cball c r) = ennreal (content (cball c r))"
using emeasure_lborel_cball_finite[of c r] by (intro emeasure_eq_ennreal_measure) auto
finally show ?thesis
using \<open>content (box a b) > 0\<close> by simp
qed
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})"
\<comment> \<open>Prove using measure theory\<close>
proof (cases "\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i")
case True
have 1: "\<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))"
by (simp add: if_distrib prod.delta_remove assms)
note simps = interval_split[OF assms] content_cbox_cases
have 2: "(\<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)"
by (metis (no_types, lifting) assms finite_Basis mult.commute prod.remove)
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!: prod.cong)
have "\<not> a \<bullet> k \<le> c \<Longrightarrow> \<not> c \<le> b \<bullet> k \<Longrightarrow> False"
unfolding not_le using True assms by auto
ultimately show ?thesis
using assms unfolding simps ** 1[of "\<lambda>i x. b\<bullet>i - x"] 1[of "\<lambda>i x. x - a\<bullet>i"] 2
by auto
next
case False
then have "cbox a b = {}"
unfolding box_eq_empty by (auto simp: not_le)
then show ?thesis
by (auto simp: not_le)
qed
lemma division_of_content_0:
assumes "content (cbox a b) = 0" "d division_of (cbox a b)" "K \<in> d"
shows "content K = 0"
unfolding forall_in_division[OF assms(2)]
by (meson assms content_0_subset division_of_def)
lemma sum_content_null:
assumes "content (cbox a b) = 0"
and "p tagged_division_of (cbox a b)"
shows "(\<Sum>(x,K)\<in>p. content K *\<^sub>R f x) = (0::'a::real_normed_vector)"
proof (rule sum.neutral, rule)
fix y
assume y: "y \<in> p"
obtain x K where xk: "y = (x, K)"
using surj_pair[of y] by blast
then obtain c d where k: "K = cbox c d" "K \<subseteq> cbox a b"
by (metis assms(2) tagged_division_ofD(3) tagged_division_ofD(4) y)
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 assms(1) content_0_subset k(2) by auto
finally show "(\<lambda>(x, k). content k *\<^sub>R f x) y = 0" .
qed
global_interpretation sum_content: operative plus 0 content
rewrites "comm_monoid_set.F plus 0 = sum"
proof -
interpret operative plus 0 content
by standard (auto simp add: content_split [symmetric] content_eq_0_interior)
show "operative plus 0 content"
by standard
show "comm_monoid_set.F plus 0 = sum"
by (simp add: sum_def)
qed
lemma additive_content_division: "d division_of (cbox a b) \<Longrightarrow> sum content d = content (cbox a b)"
by (fact sum_content.division)
lemma additive_content_tagged_division:
"d tagged_division_of (cbox a b) \<Longrightarrow> sum (\<lambda>(x,l). content l) d = content (cbox a b)"
by (fact sum_content.tagged_division)
lemma subadditive_content_division:
assumes "\<D> division_of S" "S \<subseteq> cbox a b"
shows "sum content \<D> \<le> content(cbox a b)"
proof -
have "\<D> division_of \<Union>\<D>" "\<Union>\<D> \<subseteq> cbox a b"
using assms by auto
then obtain \<D>' where "\<D> \<subseteq> \<D>'" "\<D>' division_of cbox a b"
using partial_division_extend_interval by metis
then have "sum content \<D> \<le> sum content \<D>'"
using sum_mono2 by blast
also have "... \<le> content(cbox a b)"
by (simp add: \<open>\<D>' division_of cbox a b\<close> additive_content_division less_eq_real_def)
finally show ?thesis .
qed
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 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
lemma interval_bounds_nz_content [simp]:
assumes "content (cbox a b) \<noteq> 0"
shows "interval_upperbound (cbox a b) = b"
and "interval_lowerbound (cbox a b) = a"
by (metis assms content_empty interval_bounds')+
subsection \<open>Gauge integral\<close>
text \<open>Case distinction to define it first on compact intervals first, then use a limit. This is only
much later unified. In Fremlin: Measure Theory, Volume 4I this is generalized using residual sets.\<close>
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 I) s \<longleftrightarrow>
(if \<exists>a b. s = cbox a b
then ((\<lambda>p. \<Sum>(x,k)\<in>p. content k *\<^sub>R f x) \<longlongrightarrow> I) (division_filter s)
else (\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
(\<exists>z. ((\<lambda>p. \<Sum>(x,k)\<in>p. content k *\<^sub>R (if x \<in> s then f x else 0)) \<longlongrightarrow> z) (division_filter (cbox a b)) \<and>
norm (z - I) < e)))"
lemma has_integral_cbox:
"(f has_integral I) (cbox a b) \<longleftrightarrow> ((\<lambda>p. \<Sum>(x,k)\<in>p. content k *\<^sub>R f x) \<longlongrightarrow> I) (division_filter (cbox a b))"
by (auto simp add: has_integral_def)
lemma has_integral:
"(f has_integral y) (cbox a b) \<longleftrightarrow>
(\<forall>e>0. \<exists>\<gamma>. gauge \<gamma> \<and>
(\<forall>\<D>. \<D> tagged_division_of (cbox a b) \<and> \<gamma> fine \<D> \<longrightarrow>
norm (sum (\<lambda>(x,k). content(k) *\<^sub>R f x) \<D> - y) < e))"
by (auto simp: dist_norm eventually_division_filter has_integral_def tendsto_iff)
lemma has_integral_real:
"(f has_integral y) {a..b::real} \<longleftrightarrow>
(\<forall>e>0. \<exists>\<gamma>. gauge \<gamma> \<and>
(\<forall>\<D>. \<D> tagged_division_of {a..b} \<and> \<gamma> fine \<D> \<longrightarrow>
norm (sum (\<lambda>(x,k). content(k) *\<^sub>R f x) \<D> - 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 \<gamma>
where "gauge \<gamma>"
and "\<And>\<D>. \<D> tagged_division_of (cbox a b) \<Longrightarrow> \<gamma> fine \<D> \<Longrightarrow>
norm ((\<Sum>(x,k)\<in>\<D>. content k *\<^sub>R f x) - 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)))"
by (subst has_integral_def) (auto simp add: has_integral_cbox)
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 has_integral_alt[of f y i] 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> \<not> f integrable_on i \<and> y=0)"
lemma integrable_integral[intro]: "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: "\<not> f integrable_on i \<Longrightarrow> integral i f = 0"
unfolding integrable_on_def integral_def by blast
lemma has_integral_integrable[dest]: "(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
subsection \<open>Basic theorems about integrals\<close>
lemma has_integral_eq_rhs: "(f has_integral j) S \<Longrightarrow> i = j \<Longrightarrow> (f has_integral i) S"
by (rule forw_subst)
lemma has_integral_unique_cbox:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector"
shows "(f has_integral k1) (cbox a b) \<Longrightarrow> (f has_integral k2) (cbox a b) \<Longrightarrow> k1 = k2"
by (auto simp: has_integral_cbox intro: tendsto_unique[OF division_filter_not_empty])
lemma has_integral_unique:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector"
assumes "(f has_integral k1) i" "(f has_integral k2) i"
shows "k1 = k2"
proof (rule ccontr)
let ?e = "norm (k1 - k2)/2"
let ?F = "(\<lambda>x. if x \<in> i then f x else 0)"
assume "k1 \<noteq> k2"
then have e: "?e > 0"
by auto
have nonbox: "\<not> (\<exists>a b. i = cbox a b)"
using \<open>k1 \<noteq> k2\<close> assms has_integral_unique_cbox by blast
obtain B1 where B1:
"0 < B1"
"\<And>a b. ball 0 B1 \<subseteq> cbox a b \<Longrightarrow>
\<exists>z. (?F has_integral z) (cbox a b) \<and> norm (z - k1) < norm (k1 - k2)/2"
by (rule has_integral_altD[OF assms(1) nonbox,OF e]) blast
obtain B2 where B2:
"0 < B2"
"\<And>a b. ball 0 B2 \<subseteq> cbox a b \<Longrightarrow>
\<exists>z. (?F has_integral z) (cbox a b) \<and> norm (z - k2) < norm (k1 - k2)/2"
by (rule has_integral_altD[OF assms(2) nonbox,OF e]) blast
obtain a b :: 'n where ab: "ball 0 B1 \<subseteq> cbox a b" "ball 0 B2 \<subseteq> cbox a b"
by (metis Un_subset_iff bounded_Un bounded_ball bounded_subset_cbox_symmetric)
obtain w where w: "(?F has_integral w) (cbox a b)" "norm (w - k1) < norm (k1 - k2)/2"
using B1(2)[OF ab(1)] by blast
obtain z where z: "(?F has_integral z) (cbox a b)" "norm (z - k2) < norm (k1 - k2)/2"
using B2(2)[OF ab(2)] by blast
have "z = w"
using has_integral_unique_cbox[OF w(1) z(1)] by auto
then have "norm (k1 - k2) \<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"
by (metis add_strict_mono z(2) w(2))
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 has_integral_iff: "(f has_integral i) S \<longleftrightarrow> (f integrable_on S \<and> integral S f = i)"
by blast
lemma eq_integralD: "integral k f = y \<Longrightarrow> (f has_integral y) k \<or> \<not> 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_const [intro]:
fixes a b :: "'a::euclidean_space"
shows "((\<lambda>x. c) has_integral (content (cbox a b) *\<^sub>R c)) (cbox a b)"
using eventually_division_filter_tagged_division[of "cbox a b"]
additive_content_tagged_division[of _ a b]
by (auto simp: has_integral_cbox split_beta' scaleR_sum_left[symmetric]
elim!: eventually_mono intro!: tendsto_cong[THEN iffD1, OF _ tendsto_const])
lemma has_integral_const_real [intro]:
fixes a b :: real
shows "((\<lambda>x. c) has_integral (content {a..b} *\<^sub>R c)) {a..b}"
by (metis box_real(2) has_integral_const)
lemma has_integral_integrable_integral: "(f has_integral i) s \<longleftrightarrow> f integrable_on s \<and> integral s f = i"
by blast
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)
lemma has_integral_is_0_cbox:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector"
assumes "\<And>x. x \<in> cbox a b \<Longrightarrow> f x = 0"
shows "(f has_integral 0) (cbox a b)"
unfolding has_integral_cbox
using eventually_division_filter_tagged_division[of "cbox a b"] assms
by (subst tendsto_cong[where g="\<lambda>_. 0"])
(auto elim!: eventually_mono intro!: sum.neutral simp: tag_in_interval)
lemma has_integral_is_0:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector"
assumes "\<And>x. x \<in> S \<Longrightarrow> f x = 0"
shows "(f has_integral 0) S"
proof (cases "(\<exists>a b. S = cbox a b)")
case True with assms has_integral_is_0_cbox show ?thesis
by blast
next
case False
have *: "(\<lambda>x. if x \<in> S then f x else 0) = (\<lambda>x. 0)"
by (auto simp add: assms)
show ?thesis
using has_integral_is_0_cbox False
by (subst has_integral_alt) (force simp add: *)
qed
lemma has_integral_0[simp]: "((\<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_cbox:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector"
assumes f: "(f has_integral y) (cbox a b)"
and h: "bounded_linear h"
shows "((h \<circ> f) has_integral (h y)) (cbox a b)"
proof -
interpret bounded_linear h using h .
show ?thesis
unfolding has_integral_cbox using tendsto [OF f [unfolded has_integral_cbox]]
by (simp add: sum scaleR split_beta')
qed
lemma has_integral_linear:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector"
assumes f: "(f has_integral y) S"
and h: "bounded_linear h"
shows "((h \<circ> f) has_integral (h y)) S"
proof (cases "(\<exists>a b. S = cbox a b)")
case True with f h has_integral_linear_cbox show ?thesis
by blast
next
case False
interpret bounded_linear h using h .
from pos_bounded obtain B where B: "0 < B" "\<And>x. norm (h x) \<le> norm x * B"
by blast
let ?S = "\<lambda>f x. if x \<in> S then f x else 0"
show ?thesis
proof (subst has_integral_alt, clarsimp simp: False)
fix e :: real
assume e: "e > 0"
have *: "0 < e/B" using e B(1) by simp
obtain M where M:
"M > 0"
"\<And>a b. ball 0 M \<subseteq> cbox a b \<Longrightarrow>
\<exists>z. (?S f has_integral z) (cbox a b) \<and> norm (z - y) < e/B"
using has_integral_altD[OF f False *] by blast
show "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
(\<exists>z. (?S(h \<circ> f) has_integral z) (cbox a b) \<and> norm (z - h y) < e)"
proof (rule exI, intro allI conjI impI)
show "M > 0" using M by metis
next
fix a b::'n
assume sb: "ball 0 M \<subseteq> cbox a b"
obtain z where z: "(?S f has_integral z) (cbox a b)" "norm (z - y) < e/B"
using M(2)[OF sb] by blast
have *: "?S(h \<circ> f) = h \<circ> ?S f"
using zero by auto
show "\<exists>z. (?S(h \<circ> f) has_integral z) (cbox a b) \<and> norm (z - h y) < e"
apply (rule_tac x="h z" in exI)
apply (simp add: * has_integral_linear_cbox[OF z(1) h])
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 integrable_on_scaleR_left:
assumes "f integrable_on A"
shows "(\<lambda>x. f x *\<^sub>R y) integrable_on A"
using assms has_integral_scaleR_left unfolding integrable_on_def by blast
lemma has_integral_mult_left:
fixes c :: "_ :: real_normed_algebra"
shows "(f has_integral y) S \<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)
lemma has_integral_divide:
fixes c :: "_ :: real_normed_div_algebra"
shows "(f has_integral y) S \<Longrightarrow> ((\<lambda>x. f x / c) has_integral (y / c)) S"
unfolding divide_inverse by (simp add: has_integral_mult_left)
text\<open>The case analysis eliminates the condition \<^term>\<open>f integrable_on S\<close> 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 "\<not> (\<lambda>x. f x * c) integrable_on S"
using has_integral_mult_left [of "(\<lambda>x. f x * c)" _ S "inverse c"]
by (auto simp add: mult.assoc)
with False show ?thesis by (simp add: not_integrable_integral)
qed
corollary integral_mult_right [simp]:
fixes c:: "'a::{real_normed_field}"
shows "integral S (\<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_neg_iff: "((\<lambda>x. - f x) has_integral k) S \<longleftrightarrow> (f has_integral - k) S"
using has_integral_neg[of f "- k"] has_integral_neg[of "\<lambda>x. - f x" k] by auto
lemma has_integral_add_cbox:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector"
assumes "(f has_integral k) (cbox a b)" "(g has_integral l) (cbox a b)"
shows "((\<lambda>x. f x + g x) has_integral (k + l)) (cbox a b)"
using assms
unfolding has_integral_cbox
by (simp add: split_beta' scaleR_add_right sum.distrib[abs_def] tendsto_add)
lemma has_integral_add:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector"
assumes f: "(f has_integral k) S" and g: "(g has_integral l) S"
shows "((\<lambda>x. f x + g x) has_integral (k + l)) S"
proof (cases "\<exists>a b. S = cbox a b")
case True with has_integral_add_cbox assms show ?thesis
by blast
next
let ?S = "\<lambda>f x. if x \<in> S then f x else 0"
case False
then show ?thesis
proof (subst has_integral_alt, clarsimp, goal_cases)
case (1 e)
then have e2: "e/2 > 0"
by auto
obtain Bf where "0 < Bf"
and Bf: "\<And>a b. ball 0 Bf \<subseteq> cbox a b \<Longrightarrow>
\<exists>z. (?S f has_integral z) (cbox a b) \<and> norm (z - k) < e/2"
using has_integral_altD[OF f False e2] by blast
obtain Bg where "0 < Bg"
and Bg: "\<And>a b. ball 0 Bg \<subseteq> (cbox a b) \<Longrightarrow>
\<exists>z. (?S g has_integral z) (cbox a b) \<and> norm (z - l) < e/2"
using has_integral_altD[OF g False e2] by blast
show ?case
proof (rule_tac x="max Bf Bg" in exI, clarsimp simp add: max.strict_coboundedI1 \<open>0 < Bf\<close>)
fix a b
assume "ball 0 (max Bf Bg) \<subseteq> cbox a (b::'n)"
then have fs: "ball 0 Bf \<subseteq> cbox a (b::'n)" and gs: "ball 0 Bg \<subseteq> cbox a (b::'n)"
by auto
obtain w where w: "(?S f has_integral w) (cbox a b)" "norm (w - k) < e/2"
using Bf[OF fs] by blast
obtain z where z: "(?S g has_integral z) (cbox a b)" "norm (z - l) < e/2"
using Bg[OF gs] by blast
have *: "\<And>x. (if x \<in> S then f x + g x else 0) = (?S f x) + (?S g x)"
by auto
show "\<exists>z. (?S(\<lambda>x. f x + g x) has_integral z) (cbox a b) \<and> norm (z - (k + l)) < e"
apply (rule_tac x="w + z" in exI)
apply (simp add: has_integral_add_cbox[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_diff:
"(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 integrable_integral show ?thesis
by fastforce
next
case False then have "\<not> (\<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 auto
with False show ?thesis by (simp add: not_integrable_integral)
qed
lemma integral_mult:
fixes K::real
shows "f integrable_on X \<Longrightarrow> K * integral X f = integral X (\<lambda>x. K * f x)"
unfolding real_scaleR_def[symmetric] integral_cmul ..
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 "\<not> (\<lambda>x. - f x) integrable_on S"
using has_integral_neg [of "(\<lambda>x. - f x)" _ S ] by auto
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_diff)
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_scaleR_iff [simp]:
fixes c :: real
assumes "c \<noteq> 0"
shows "(\<lambda>x. c *\<^sub>R f x) integrable_on S \<longleftrightarrow> f integrable_on S"
using integrable_cmul[of "\<lambda>x. c *\<^sub>R f x" S "1 / c"] integrable_cmul[of f S c] \<open>c \<noteq> 0\<close>
by auto
lemma integrable_on_cmult_iff [simp]:
fixes c :: real
assumes "c \<noteq> 0"
shows "(\<lambda>x. c * f x) integrable_on S \<longleftrightarrow> f integrable_on S"
using integrable_on_scaleR_iff [of c f] assms by simp
lemma integrable_on_cmult_left:
assumes "f integrable_on S"
shows "(\<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_neg_iff: "(\<lambda>x. -f(x)) integrable_on S \<longleftrightarrow> f integrable_on S"
using integrable_neg by fastforce
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_diff)
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 integrable_on_cnj_iff:
"(\<lambda>x. cnj (f x)) integrable_on A \<longleftrightarrow> f integrable_on A"
using integrable_linear[OF _ bounded_linear_cnj, of f A]
integrable_linear[OF _ bounded_linear_cnj, of "cnj \<circ> f" A]
by (auto simp: o_def)
lemma integral_cnj: "cnj (integral A f) = integral A (\<lambda>x. cnj (f x))"
by (cases "f integrable_on A")
(simp_all add: integral_linear[OF _ bounded_linear_cnj, symmetric]
o_def integrable_on_cnj_iff not_integrable_integral)
lemma integral_component_eq[simp]:
fixes f :: "'n::euclidean_space \<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_inner_left,unfolded o_def] ..
lemma has_integral_sum:
assumes "finite T"
and "\<And>a. a \<in> T \<Longrightarrow> ((f a) has_integral (i a)) S"
shows "((\<lambda>x. sum (\<lambda>a. f a x) T) has_integral (sum i T)) S"
using assms(1) subset_refl[of T]
proof (induct rule: finite_subset_induct)
case empty
then show ?case by auto
next
case (insert x F)
with assms show ?case
by (simp add: has_integral_add)
qed
lemma integral_sum:
"\<lbrakk>finite I; \<And>a. a \<in> I \<Longrightarrow> f a integrable_on S\<rbrakk> \<Longrightarrow>
integral S (\<lambda>x. \<Sum>a\<in>I. f a x) = (\<Sum>a\<in>I. integral S (f a))"
by (simp add: has_integral_sum integrable_integral integral_unique)
lemma integrable_sum:
"\<lbrakk>finite I; \<And>a. a \<in> I \<Longrightarrow> f a integrable_on S\<rbrakk> \<Longrightarrow> (\<lambda>x. \<Sum>a\<in>I. f a x) integrable_on S"
unfolding integrable_on_def using has_integral_sum[of I] by metis
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_diff[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: "\<lbrakk>f integrable_on s; \<And>x. x \<in> s \<Longrightarrow> f x = g x\<rbrakk> \<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 flip: 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 flip: of_real_inverse)
lemma has_integral_null [intro]: "content(cbox a b) = 0 \<Longrightarrow> (f has_integral 0) (cbox a b)"
unfolding has_integral_cbox
using eventually_division_filter_tagged_division[of "cbox a b"]
by (subst tendsto_cong[where g="\<lambda>_. 0"]) (auto elim: eventually_mono intro: sum_content_null)
lemma has_integral_null_real [intro]: "content {a..b::real} = 0 \<Longrightarrow> (f has_integral 0) {a..b}"
by (metis box_real(2) has_integral_null)
lemma has_integral_null_eq[simp]: "content (cbox a b) = 0 \<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 (meson ex_in_conv 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 -
show "(f has_integral 0) (cbox a a)"
by (rule has_integral_null) simp
then show "(f has_integral 0) {a}"
by simp
qed
lemma integrable_on_refl[intro]: "f integrable_on cbox a a"
unfolding integrable_on_def by auto
lemma integral_refl [simp]: "integral (cbox a a) f = 0"
by (rule integral_unique) auto
lemma integral_singleton [simp]: "integral {a} f = 0"
by auto
lemma integral_blinfun_apply:
assumes "f integrable_on s"
shows "integral s (\<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_inner_left, 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_sum) (simp_all add: o_def)
thus "(f has_integral y) A" by (simp add: euclidean_representation)
qed
lemma has_integral_componentwise:
fixes f :: "'a :: euclidean_space \<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_sum) (simp_all add: o_def)
thus "f integrable_on A" by (auto simp: integrable_on_def o_def euclidean_representation)
qed
lemma integrable_componentwise:
fixes f :: "'a :: euclidean_space \<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_sum) (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_inner_left[of y]]) (simp add: o_def)
subsection \<open>Cauchy-type criterion for integrability\<close>
proposition 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>\<gamma>. gauge \<gamma> \<and>
(\<forall>\<D>1 \<D>2. \<D>1 tagged_division_of (cbox a b) \<and> \<gamma> fine \<D>1 \<and>
\<D>2 tagged_division_of (cbox a b) \<and> \<gamma> fine \<D>2 \<longrightarrow>
norm ((\<Sum>(x,K)\<in>\<D>1. content K *\<^sub>R f x) - (\<Sum>(x,K)\<in>\<D>2. content K *\<^sub>R f x)) < e))"
(is "?l = (\<forall>e>0. \<exists>\<gamma>. ?P e \<gamma>)")
proof (intro iffI allI impI)
assume ?l
then obtain y
where y: "\<And>e. e > 0 \<Longrightarrow>
\<exists>\<gamma>. gauge \<gamma> \<and>
(\<forall>\<D>. \<D> tagged_division_of cbox a b \<and> \<gamma> fine \<D> \<longrightarrow>
norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - y) < e)"
by (auto simp: integrable_on_def has_integral)
show "\<exists>\<gamma>. ?P e \<gamma>" if "e > 0" for e
proof -
have "e/2 > 0" using that by auto
with y obtain \<gamma> where "gauge \<gamma>"
and \<gamma>: "\<And>\<D>. \<D> tagged_division_of cbox a b \<and> \<gamma> fine \<D> \<Longrightarrow>
norm ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f x) - y) < e/2"
by meson
show ?thesis
apply (rule_tac x=\<gamma> in exI, clarsimp simp: \<open>gauge \<gamma>\<close>)
by (blast intro!: \<gamma> dist_triangle_half_l[where y=y,unfolded dist_norm])
qed
next
assume "\<forall>e>0. \<exists>\<gamma>. ?P e \<gamma>"
then have "\<forall>n::nat. \<exists>\<gamma>. ?P (1 / (n + 1)) \<gamma>"
by auto
then obtain \<gamma> :: "nat \<Rightarrow> 'n \<Rightarrow> 'n set" where \<gamma>:
"\<And>m. gauge (\<gamma> m)"
"\<And>m \<D>1 \<D>2. \<lbrakk>\<D>1 tagged_division_of cbox a b;
\<gamma> m fine \<D>1; \<D>2 tagged_division_of cbox a b; \<gamma> m fine \<D>2\<rbrakk>
\<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>1. content K *\<^sub>R f x) - (\<Sum>(x,K) \<in> \<D>2. content K *\<^sub>R f x))
< 1 / (m + 1)"
by metis
have "\<And>n. gauge (\<lambda>x. \<Inter>{\<gamma> i x |i. i \<in> {0..n}})"
apply (rule gauge_Inter)
using \<gamma> by auto
then have "\<forall>n. \<exists>p. p tagged_division_of (cbox a b) \<and> (\<lambda>x. \<Inter>{\<gamma> i x |i. i \<in> {0..n}}) fine p"
by (meson fine_division_exists)
then obtain p where p: "\<And>z. p z tagged_division_of cbox a b"
"\<And>z. (\<lambda>x. \<Inter>{\<gamma> i x |i. i \<in> {0..z}}) fine p z"
by meson
have dp: "\<And>i n. i\<le>n \<Longrightarrow> \<gamma> i fine p n"
using p unfolding fine_Inter
using atLeastAtMost_iff by blast
have "Cauchy (\<lambda>n. sum (\<lambda>(x,K). content K *\<^sub>R (f x)) (p n))"
proof (rule CauchyI)
fix e::real
assume "0 < e"
then obtain N where "N \<noteq> 0" and N: "inverse (real N) < e"
using real_arch_inverse[of e] by blast
show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. 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"
proof (intro exI allI impI)
fix m n
assume mn: "N \<le> m" "N \<le> n"
have "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)) < 1 / (real N + 1)"
by (simp add: p(1) dp mn \<gamma>)
also have "... < e"
using N \<open>N \<noteq> 0\<close> \<open>0 < e\<close> by (auto simp: field_simps)
finally 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" .
qed
qed
then obtain y where y: "\<exists>no. \<forall>n\<ge>no. norm ((\<Sum>(x,K) \<in> p n. content K *\<^sub>R f x) - y) < r" if "r > 0" for r
by (auto simp: convergent_eq_Cauchy[symmetric] dest: LIMSEQ_D)
show ?l
unfolding integrable_on_def has_integral
proof (rule_tac x=y in exI, clarify)
fix e :: real
assume "e>0"
then have e2: "e/2 > 0" by auto
then obtain N1::nat where N1: "N1 \<noteq> 0" "inverse (real N1) < e/2"
using real_arch_inverse by blast
obtain N2::nat where N2: "\<And>n. n \<ge> N2 \<Longrightarrow> norm ((\<Sum>(x,K) \<in> p n. content K *\<^sub>R f x) - y) < e/2"
using y[OF e2] by metis
show "\<exists>\<gamma>. gauge \<gamma> \<and>
(\<forall>\<D>. \<D> tagged_division_of (cbox a b) \<and> \<gamma> fine \<D> \<longrightarrow>
norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - y) < e)"
proof (intro exI conjI allI impI)
show "gauge (\<gamma> (N1+N2))"
using \<gamma> by auto
show "norm ((\<Sum>(x,K) \<in> q. content K *\<^sub>R f x) - y) < e"
if "q tagged_division_of cbox a b \<and> \<gamma> (N1+N2) fine q" for q
proof (rule norm_triangle_half_r)
have "norm ((\<Sum>(x,K) \<in> p (N1+N2). content K *\<^sub>R f x) - (\<Sum>(x,K) \<in> q. content K *\<^sub>R f x))
< 1 / (real (N1+N2) + 1)"
by (rule \<gamma>; simp add: dp p that)
also have "... < e/2"
using N1 \<open>0 < e\<close> by (auto simp: field_simps intro: less_le_trans)
finally show "norm ((\<Sum>(x,K) \<in> p (N1+N2). content K *\<^sub>R f x) - (\<Sum>(x,K) \<in> q. content K *\<^sub>R f x)) < e/2" .
show "norm ((\<Sum>(x,K) \<in> p (N1+N2). content K *\<^sub>R f x) - y) < e/2"
using N2 le_add_same_cancel2 by blast
qed
qed
qed
qed
subsection \<open>Additivity of integral on abutting intervals\<close>
lemma tagged_division_split_left_inj_content:
assumes \<D>: "\<D> tagged_division_of S"
and "(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 -
from tagged_division_ofD(4)[OF \<D> \<open>(x1, K1) \<in> \<D>\<close>] obtain a b where K1: "K1 = cbox a b"
by auto
then have "interior (K1 \<inter> {x. x \<bullet> k \<le> c}) = {}"
by (metis tagged_division_split_left_inj assms)
then show ?thesis
unfolding K1 interval_split[OF \<open>k \<in> Basis\<close>] by (auto simp: content_eq_0_interior)
qed
lemma tagged_division_split_right_inj_content:
assumes \<D>: "\<D> tagged_division_of S"
and "(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 -
from tagged_division_ofD(4)[OF \<D> \<open>(x1, K1) \<in> \<D>\<close>] obtain a b where K1: "K1 = cbox a b"
by auto
then have "interior (K1 \<inter> {x. c \<le> x \<bullet> k}) = {}"
by (metis tagged_division_split_right_inj assms)
then show ?thesis
unfolding K1 interval_split[OF \<open>k \<in> Basis\<close>]
by (auto simp: content_eq_0_interior)
qed
proposition 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)"
unfolding has_integral
proof clarify
fix e::real
assume "0 < e"
then have e: "e/2 > 0"
by auto
obtain \<gamma>1 where \<gamma>1: "gauge \<gamma>1"
and \<gamma>1norm:
"\<And>\<D>. \<lbrakk>\<D> tagged_division_of cbox a b \<inter> {x. x \<bullet> k \<le> c}; \<gamma>1 fine \<D>\<rbrakk>
\<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. 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 \<gamma>2 where \<gamma>2: "gauge \<gamma>2"
and \<gamma>2norm:
"\<And>\<D>. \<lbrakk>\<D> tagged_division_of cbox a b \<inter> {x. c \<le> x \<bullet> k}; \<gamma>2 fine \<D>\<rbrakk>
\<Longrightarrow> norm ((\<Sum>(x, k) \<in> \<D>. 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 ?\<gamma> = "\<lambda>x. if x\<bullet>k = c then (\<gamma>1 x \<inter> \<gamma>2 x) else ball x \<bar>x\<bullet>k - c\<bar> \<inter> \<gamma>1 x \<inter> \<gamma>2 x"
have "gauge ?\<gamma>"
using \<gamma>1 \<gamma>2 unfolding gauge_def by auto
then show "\<exists>\<gamma>. gauge \<gamma> \<and>
(\<forall>\<D>. \<D> tagged_division_of cbox a b \<and> \<gamma> fine \<D> \<longrightarrow>
norm ((\<Sum>(x, k)\<in>\<D>. content k *\<^sub>R f x) - (i + j)) < e)"
proof (rule_tac x="?\<gamma>" in exI, safe)
fix p
assume p: "p tagged_division_of (cbox a b)" and "?\<gamma> fine p"
have ab_eqp: "cbox a b = \<Union>{K. \<exists>x. (x, K) \<in> p}"
using p by blast
have xk_le_c: "x\<bullet>k \<le> c" if as: "(x,K) \<in> p" and K: "K \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}" for x K
proof (rule ccontr)
assume **: "\<not> x \<bullet> k \<le> c"
then have "K \<subseteq> ball x \<bar>x \<bullet> k - c\<bar>"
using \<open>?\<gamma> fine p\<close> as by (fastforce simp: not_le algebra_simps)
with K 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
have xk_ge_c: "x\<bullet>k \<ge> c" if as: "(x,K) \<in> p" and K: "K \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}" for x K
proof (rule ccontr)
assume **: "\<not> x \<bullet> k \<ge> c"
then have "K \<subseteq> ball x \<bar>x \<bullet> k - c\<bar>"
using \<open>?\<gamma> fine p\<close> as by (fastforce simp: not_le algebra_simps)
with K 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
have fin_finite: "finite {(x,f K) | x K. (x,K) \<in> s \<and> P x K}"
if "finite s" for s and f :: "'a set \<Rightarrow> 'a set" and P :: "'a \<Rightarrow> 'a set \<Rightarrow> bool"
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 "finite p"
using p by blast
let ?M1 = "{(x, K \<inter> {x. x\<bullet>k \<le> c}) |x K. (x,K) \<in> p \<and> K \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}}"
have \<gamma>1_fine: "\<gamma>1 fine ?M1"
using \<open>?\<gamma> fine p\<close> by (fastforce simp: fine_def split: if_split_asm)
have "norm ((\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) - i) < e/2"
proof (rule \<gamma>1norm [OF tagged_division_ofI \<gamma>1_fine])
show "finite ?M1"
by (rule fin_finite) (use p in blast)
show "\<Union>{k. \<exists>x. (x, k) \<in> ?M1} = cbox a b \<inter> {x. x\<bullet>k \<le> c}"
by (auto simp: ab_eqp)
fix x L
assume xL: "(x, L) \<in> ?M1"
then obtain x' L' where xL': "x = x'" "L = L' \<inter> {x. x \<bullet> k \<le> c}"
"(x', L') \<in> p" "L' \<inter> {x. x \<bullet> k \<le> c} \<noteq> {}"
by blast
then obtain a' b' where ab': "L' = cbox a' b'"
using p by blast
show "x \<in> L" "L \<subseteq> cbox a b \<inter> {x. x \<bullet> k \<le> c}"
using p xk_le_c xL' by auto
show "\<exists>a b. L = cbox a b"
using p xL' ab' by (auto simp add: interval_split[OF k,where c=c])
fix y R
assume yR: "(y, R) \<in> ?M1"
then obtain y' R' where yR': "y = y'" "R = R' \<inter> {x. x \<bullet> k \<le> c}"
"(y', R') \<in> p" "R' \<inter> {x. x \<bullet> k \<le> c} \<noteq> {}"
by blast
assume as: "(x, L) \<noteq> (y, R)"
show "interior L \<inter> interior R = {}"
proof (cases "L' = R' \<longrightarrow> x' = y'")
case False
have "interior R' = {}"
by (metis (no_types) False Pair_inject inf.idem tagged_division_ofD(5) [OF p] xL'(3) yR'(3))
then show ?thesis
using yR' by simp
next
case True
then have "L' \<noteq> R'"
using as unfolding xL' yR' by auto
have "interior L' \<inter> interior R' = {}"
by (metis (no_types) Pair_inject \<open>L' \<noteq> R'\<close> p tagged_division_ofD(5) xL'(3) yR'(3))
then show ?thesis
using xL'(2) yR'(2) by auto
qed
qed
moreover
let ?M2 = "{(x,K \<inter> {x. x\<bullet>k \<ge> c}) |x K. (x,K) \<in> p \<and> K \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}}"
have \<gamma>2_fine: "\<gamma>2 fine ?M2"
using \<open>?\<gamma> fine p\<close> by (fastforce simp: fine_def split: if_split_asm)
have "norm ((\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - j) < e/2"
proof (rule \<gamma>2norm [OF tagged_division_ofI \<gamma>2_fine])
show "finite ?M2"
by (rule fin_finite) (use p in blast)
show "\<Union>{k. \<exists>x. (x, k) \<in> ?M2} = cbox a b \<inter> {x. x\<bullet>k \<ge> c}"
by (auto simp: ab_eqp)
fix x L
assume xL: "(x, L) \<in> ?M2"
then obtain x' L' where xL': "x = x'" "L = L' \<inter> {x. x \<bullet> k \<ge> c}"
"(x', L') \<in> p" "L' \<inter> {x. x \<bullet> k \<ge> c} \<noteq> {}"
by blast
then obtain a' b' where ab': "L' = cbox a' b'"
using p by blast
show "x \<in> L" "L \<subseteq> cbox a b \<inter> {x. x \<bullet> k \<ge> c}"
using p xk_ge_c xL' by auto
show "\<exists>a b. L = cbox a b"
using p xL' ab' by (auto simp add: interval_split[OF k,where c=c])
fix y R
assume yR: "(y, R) \<in> ?M2"
then obtain y' R' where yR': "y = y'" "R = R' \<inter> {x. x \<bullet> k \<ge> c}"
"(y', R') \<in> p" "R' \<inter> {x. x \<bullet> k \<ge> c} \<noteq> {}"
by blast
assume as: "(x, L) \<noteq> (y, R)"
show "interior L \<inter> interior R = {}"
proof (cases "L' = R' \<longrightarrow> x' = y'")
case False
have "interior R' = {}"
by (metis (no_types) False Pair_inject inf.idem tagged_division_ofD(5) [OF p] xL'(3) yR'(3))
then show ?thesis
using yR' by simp
next
case True
then have "L' \<noteq> R'"
using as unfolding xL' yR' by auto
have "interior L' \<inter> interior R' = {}"
by (metis (no_types) Pair_inject \<open>L' \<noteq> R'\<close> p tagged_division_ofD(5) xL'(3) yR'(3))
then show ?thesis
using xL'(2) yR'(2) 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
moreover 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, ka)\<in>p. content ka *\<^sub>R f x) - (i + j)"
proof -
have eq0: "\<And>x y. x = (0::real) \<Longrightarrow> x *\<^sub>R (y::'b) = 0"
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 *: "\<And>\<G> :: 'a set \<Rightarrow> 'a set.
(\<Sum>(x,K)\<in>{(x, \<G> K) |x K. (x,K) \<in> p \<and> \<G> K \<noteq> {}}. content K *\<^sub>R f x) =
(\<Sum>(x,K)\<in>(\<lambda>(x,K). (x, \<G> K)) ` p. content K *\<^sub>R f x)"
by (rule sum.mono_neutral_left) (auto simp: \<open>finite p\<close>)
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
moreover have "\<dots> = (\<Sum>(x,K) \<in> p. content (K \<inter> {x. x \<bullet> k \<le> c}) *\<^sub>R f x) +
(\<Sum>(x,K) \<in> p. content (K \<inter> {x. c \<le> x \<bullet> k}) *\<^sub>R f x) - (i + j)"
unfolding *
apply (subst (1 2) sum.reindex_nontrivial)
apply (auto intro!: k p eq0 tagged_division_split_left_inj_content tagged_division_split_right_inj_content
simp: cont_eq \<open>finite p\<close>)
done
moreover have "\<And>x. x \<in> p \<Longrightarrow> (\<lambda>(a,B). content (B \<inter> {a. a \<bullet> k \<le> c}) *\<^sub>R f a) x +
(\<lambda>(a,B). content (B \<inter> {a. c \<le> a \<bullet> k}) *\<^sub>R f a) x =
(\<lambda>(a,B). content B *\<^sub>R f a) x"
proof clarify
fix a B
assume "(a, B) \<in> p"
with p obtain u v where uv: "B = cbox u v" by blast
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"
by (auto simp: scaleR_left_distrib uv content_split[OF k,of u v c])
qed
ultimately show ?thesis
by (auto simp: sum.distrib[symmetric])
qed
ultimately show "norm ((\<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 has_integral_separate_sides:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
assumes f: "(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 ((sum (\<lambda>(x,k). content k *\<^sub>R f x) p1 + sum (\<lambda>(x,k). content k *\<^sub>R f x) p2) - i) < e"
proof -
obtain \<gamma> where d: "gauge \<gamma>"
"\<And>p. \<lbrakk>p tagged_division_of cbox a b; \<gamma> fine p\<rbrakk>
\<Longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - i) < e"
using has_integralD[OF f \<open>e > 0\<close>] by metis
{ fix p1 p2
assume tdiv1: "p1 tagged_division_of (cbox a b) \<inter> {x. x \<bullet> k \<le> c}" and "\<gamma> fine p1"
note p1=tagged_division_ofD[OF this(1)]
assume tdiv2: "p2 tagged_division_of (cbox a b) \<inter> {x. c \<le> x \<bullet> k}" and "\<gamma> fine p2"
note p2=tagged_division_ofD[OF this(1)]
note tagged_division_Un_interval[OF tdiv1 tdiv2]
note p12 = tagged_division_ofD[OF this] this
{ fix a b
assume ab: "(a, b) \<in> p1 \<inter> p2"
have "(a, b) \<in> p1"
using ab by auto
obtain u v where uv: "b = cbox u v"
using \<open>(a, b) \<in> p1\<close> p1(4) by moura
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 obtain \<epsilon> where "0 < \<epsilon>" and \<epsilon>: "ball x \<epsilon> \<subseteq> {x. x \<bullet> k = c}"
using mem_interior by metis
have x: "x\<bullet>k = c"
using x interior_subset by fastforce
have *: "\<And>i. i \<in> Basis \<Longrightarrow> \<bar>(x - (x + (\<epsilon>/2) *\<^sub>R k)) \<bullet> i\<bar> = (if i = k then \<epsilon>/2 else 0)"
using \<open>0 < \<epsilon>\<close> k by (auto simp: inner_simps inner_not_same_Basis)
have "(\<Sum>i\<in>Basis. \<bar>(x - (x + (\<epsilon>/2 ) *\<^sub>R k)) \<bullet> i\<bar>) =
(\<Sum>i\<in>Basis. (if i = k then \<epsilon>/2 else 0))"
using "*" by (blast intro: sum.cong)
also have "\<dots> < \<epsilon>"
by (subst sum.delta) (use \<open>0 < \<epsilon>\<close> in auto)
finally have "x + (\<epsilon>/2) *\<^sub>R k \<in> ball x \<epsilon>"
unfolding mem_ball dist_norm by(rule le_less_trans[OF norm_le_l1])
then have "x + (\<epsilon>/2) *\<^sub>R k \<in> {x. x\<bullet>k = c}"
using \<epsilon> by auto
then show False
using \<open>0 < \<epsilon>\<close> 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 sum.union_inter_neutral) (auto simp: p1 p2)
also have "\<dots> < e"
using d(2) p12 by (simp add: fine_Un k \<open>\<gamma> fine p1\<close> \<open>\<gamma> fine p2\<close>)
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
using d(1) that by auto
qed
lemma integrable_split [intro]:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::{real_normed_vector,complete_space}"
assumes f: "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 ?thesis1)
and "f integrable_on (cbox a b \<inter> {x. x\<bullet>k \<ge> c})" (is ?thesis2)
proof -
obtain y where y: "(f has_integral y) (cbox a b)"
using f by blast
define a' where "a' = (\<Sum>i\<in>Basis. (if i = k then max (a\<bullet>k) c else a\<bullet>i)*\<^sub>R i)"
define b' where "b' = (\<Sum>i\<in>Basis. (if i = k then min (b\<bullet>k) c else b\<bullet>i)*\<^sub>R i)"
have "\<exists>d. gauge d \<and>
(\<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 \<le> c} \<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)"
if "e > 0" for e
proof -
have "e/2 > 0" using that by auto
with has_integral_separate_sides[OF y this k, of c]
obtain d
where "gauge d"
and d: "\<And>p1 p2. \<lbrakk>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. c \<le> x \<bullet> k}; d fine p2\<rbrakk>
\<Longrightarrow> norm ((\<Sum>(x,K)\<in>p1. content K *\<^sub>R f x) + (\<Sum>(x,K)\<in>p2. content K *\<^sub>R f x) - y) < e/2"
by metis
show ?thesis
proof (rule_tac x=d in exI, clarsimp simp add: \<open>gauge d\<close>)
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 \<open>gauge d\<close>, of a' b])
fix p
assume "p tagged_division_of cbox a' b" "d fine p"
then show ?thesis
using as norm_triangle_half_l[OF d[of p1 p] d[of p2 p]]
unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric]
by (auto simp add: algebra_simps)
qed
qed
qed
with f show ?thesis1
by (simp add: interval_split[OF k] integrable_Cauchy)
have "\<exists>d. gauge d \<and>
(\<forall>p1 p2. p1 tagged_division_of cbox a b \<inter> {x. x \<bullet> k \<ge> 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 ((\<Sum>(x,K) \<in> p1. content K *\<^sub>R f x) - (\<Sum>(x,K) \<in> p2. content K *\<^sub>R f x)) < e)"
if "e > 0" for e
proof -
have "e/2 > 0" using that by auto
with has_integral_separate_sides[OF y this k, of c]
obtain d
where "gauge d"
and d: "\<And>p1 p2. \<lbrakk>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. c \<le> x \<bullet> k}; d fine p2\<rbrakk>
\<Longrightarrow> norm ((\<Sum>(x,K)\<in>p1. content K *\<^sub>R f x) + (\<Sum>(x,K)\<in>p2. content K *\<^sub>R f x) - y) < e/2"
by metis
show ?thesis
proof (rule_tac x=d in exI, clarsimp simp add: \<open>gauge d\<close>)
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 \<open>gauge d\<close>, of a b'])
fix p
assume "p tagged_division_of cbox a b'" "d fine p"
then show ?thesis
using as norm_triangle_half_l[OF d[of p p1] d[of p p2]]
unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric]
by (auto simp add: algebra_simps)
qed
qed
qed
with f show ?thesis2
by (simp add: interval_split[OF k] integrable_Cauchy)
qed
lemma operative_integralI:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
shows "operative (lift_option (+)) (Some 0)
(\<lambda>i. if f integrable_on i then Some (integral i f) else None)"
proof -
interpret comm_monoid "lift_option plus" "Some (0::'b)"
by (rule comm_monoid_lift_option)
(rule add.comm_monoid_axioms)
show ?thesis
proof
fix a b c
fix k :: 'a
assume k: "k \<in> Basis"
show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) =
lift_option (+) (if f integrable_on cbox a b \<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 "box a b = {}"
then show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) = Some 0"
using has_integral_null_eq
by (auto simp: integrable_on_null content_eq_0_interior)
qed
qed
subsection \<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 (sum (\<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>) = sum content p"
apply (rule sum.cong)
using assms
apply simp
- apply (metis abs_of_nonneg assms(1) content_pos_le division_ofD(4))
+ apply (metis abs_of_nonneg content_pos_le)
done
have e: "0 \<le> e"
using assms(2) norm_ge_zero order_trans by blast
have "norm (sum (\<lambda>l. content l *\<^sub>R c) p) \<le> (\<Sum>i\<in>p. norm (content i *\<^sub>R c))"
using norm_sum by blast
also have "... \<le> e * (\<Sum>i\<in>p. \<bar>content i\<bar>)"
by (simp add: sum_distrib_left[symmetric] mult.commute assms(2) mult_right_mono sum_nonneg)
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 (sum (\<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 sum_le: "sum (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 (sum (\<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_sum)
also have "... \<le> e * content (cbox a b)"
unfolding split_def norm_scaleR
apply (rule order_trans[OF sum_mono])
apply (rule mult_left_mono[OF _ abs_ge_zero, of _ e])
apply (metis norm)
unfolding sum_distrib_right[symmetric]
using con sum_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 (sum (\<lambda>(x,k). content k *\<^sub>R f x) p - sum (\<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 sum_subtractf eq_refl)
done
lemma has_integral_bound:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
assumes "0 \<le> B"
and f: "(f has_integral i) (cbox a b)"
and "\<And>x. x\<in>cbox a b \<Longrightarrow> 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
with f[unfolded has_integral]
obtain \<gamma> where "gauge \<gamma>" and \<gamma>:
"\<And>p. \<lbrakk>p tagged_division_of cbox a b; \<gamma> fine p\<rbrakk>
\<Longrightarrow> norm ((\<Sum>(x, K)\<in>p. content K *\<^sub>R f x) - i) < norm i - B * content (cbox a b)"
by metis
then obtain p where p: "p tagged_division_of cbox a b" and "\<gamma> fine p"
using fine_division_exists by blast
have "\<And>s B. norm s \<le> B \<Longrightarrow> \<not> norm (s - i) < norm i - B"
unfolding not_less
by (metis diff_left_mono dist_commute dist_norm norm_triangle_ineq2 order_trans)
then show False
using \<gamma> [OF p \<open>\<gamma> fine p\<close>] rsum_bound[OF p] assms by metis
qed
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: "p tagged_division_of (cbox a b)"
and "\<And>x. x \<in> cbox a b \<Longrightarrow> (f x)\<bullet>i \<le> (g x)\<bullet>i"
shows "(\<Sum>(x, K)\<in>p. content K *\<^sub>R f x) \<bullet> i \<le> (\<Sum>(x, K)\<in>p. content K *\<^sub>R g x) \<bullet> i"
unfolding inner_sum_left
proof (rule sum_mono, clarify)
fix x K
assume ab: "(x, K) \<in> p"
with p obtain u v where K: "K = cbox u v"
by blast
then show "(content K *\<^sub>R f x) \<bullet> i \<le> (content K *\<^sub>R g x) \<bullet> i"
by (metis ab assms inner_scaleR_left measure_nonneg mult_left_mono tag_in_interval)
qed
lemma has_integral_component_le:
fixes f g :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes k: "k \<in> Basis"
assumes "(f has_integral i) S" "(g has_integral j) S"
and f_le_g: "\<And>x. x \<in> S \<Longrightarrow> (f x)\<bullet>k \<le> (g x)\<bullet>k"
shows "i\<bullet>k \<le> j\<bullet>k"
proof -
have ik_le_jk: "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
obtain \<gamma>1 where "gauge \<gamma>1"
and \<gamma>1: "\<And>p. \<lbrakk>p tagged_division_of cbox a b; \<gamma>1 fine p\<rbrakk>
\<Longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - i) < (i \<bullet> k - j \<bullet> k) / 3"
using f_i[unfolded has_integral,rule_format,OF *] by fastforce
obtain \<gamma>2 where "gauge \<gamma>2"
and \<gamma>2: "\<And>p. \<lbrakk>p tagged_division_of cbox a b; \<gamma>2 fine p\<rbrakk>
\<Longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g x) - j) < (i \<bullet> k - j \<bullet> k) / 3"
using g_j[unfolded has_integral,rule_format,OF *] by fastforce
obtain p where p: "p tagged_division_of cbox a b" and "\<gamma>1 fine p" "\<gamma>2 fine p"
using fine_division_exists[OF gauge_Int[OF \<open>gauge \<gamma>1\<close> \<open>gauge \<gamma>2\<close>], of a b] unfolding fine_Int
by metis
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 le_less_trans[OF Basis_le_norm[OF k]] k \<gamma>1 \<gamma>2 by metis+
then show False
unfolding inner_simps
using rsum_component_le[OF p] le
by (fastforce simp add: abs_real_def split: if_split_asm)
qed
show ?thesis
proof (cases "\<exists>a b. S = cbox a b")
case True
with ik_le_jk assms show ?thesis
by auto
next
case False
show ?thesis
proof (rule ccontr)
assume "\<not> i\<bullet>k \<le> j\<bullet>k"
then have ij: "(i\<bullet>k - j\<bullet>k) / 3 > 0"
by auto
obtain B1 where "0 < B1"
and 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 - i) < (i \<bullet> k - j \<bullet> k) / 3"
using has_integral_altD[OF _ False ij] assms by blast
obtain B2 where "0 < B2"
and 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 - j) < (i \<bullet> k - j \<bullet> k) / 3"
using has_integral_altD[OF _ False ij] assms by blast
have "bounded (ball 0 B1 \<union> ball (0::'a) B2)"
unfolding bounded_Un by(rule conjI bounded_ball)+
from bounded_subset_cbox_symmetric[OF this]
obtain a b::'a where ab: "ball 0 B1 \<subseteq> cbox a b" "ball 0 B2 \<subseteq> cbox a b"
by (meson Un_subset_iff)
then obtain w1 w2 where int_w1: "((\<lambda>x. if x \<in> S then f x else 0) has_integral w1) (cbox a b)"
and norm_w1: "norm (w1 - i) < (i \<bullet> k - j \<bullet> k) / 3"
and int_w2: "((\<lambda>x. if x \<in> S then g x else 0) has_integral w2) (cbox a b)"
and norm_w2: "norm (w2 - j) < (i \<bullet> k - j \<bullet> k) / 3"
using B1 B2 by blast
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)
have "\<bar>(w1 - i) \<bullet> k\<bar> < (i \<bullet> k - j \<bullet> k) / 3"
"\<bar>(w2 - j) \<bullet> k\<bar> < (i \<bullet> k - j \<bullet> k) / 3"
using Basis_le_norm k le_less_trans norm_w1 norm_w2 by blast+
moreover
have "w1\<bullet>k \<le> w2\<bullet>k"
using ik_le_jk int_w1 int_w2 f_le_g by auto
ultimately show False
unfolding inner_simps by(rule *)
qed
qed
qed
lemma integral_component_le:
fixes g f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "k \<in> Basis"
and "f integrable_on S" "g integrable_on S"
and "\<And>x. x \<in> S \<Longrightarrow> (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 "\<And>x. x \<in> S \<Longrightarrow> 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 "\<And>x. x \<in> S \<Longrightarrow> 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 "\<And>x. x \<in> S \<Longrightarrow> (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 "\<And>e. e > 0 \<Longrightarrow> \<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 "1 / (real n + 1) > 0" for n
by auto
then have "\<exists>g. (\<forall>x\<in>cbox a b. norm (f x - g x) \<le> 1 / (real n + 1)) \<and> g integrable_on cbox a b" for n
using assms by blast
then obtain g where g_near_f: "\<And>n x. x \<in> cbox a b \<Longrightarrow> norm (f x - g n x) \<le> 1 / (real n + 1)"
and int_g: "\<And>n. g n integrable_on cbox a b"
by metis
then obtain h where h: "\<And>n. (g n has_integral h n) (cbox a b)"
unfolding integrable_on_def by metis
have "Cauchy h"
unfolding Cauchy_def
proof clarify
fix e :: real
assume "e>0"
then have "e/4 / content (cbox a b) > 0"
using True by (auto simp: field_simps)
then obtain M where "M \<noteq> 0" and M: "1 / (real M) < e/4 / content (cbox a b)"
by (metis inverse_eq_divide real_arch_inverse)
show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (h m) (h 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
then obtain gm gn where "gauge gm" "gauge gn"
and gm: "\<And>\<D>. \<D> tagged_division_of cbox a b \<and> gm fine \<D>
\<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g m x) - h m) < e/4"
and gn: "\<And>\<D>. \<D> tagged_division_of cbox a b \<and> gn fine \<D> \<Longrightarrow>
norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g n x) - h n) < e/4"
using h[unfolded has_integral] by meson
then obtain \<D> where \<D>: "\<D> tagged_division_of cbox a b" "(\<lambda>x. gm x \<inter> gn x) fine \<D>"
by (metis (full_types) fine_division_exists gauge_Int)
have triangle3: "norm (i1 - i2) < e"
if no: "norm(s2 - s1) \<le> e/2" "norm (s1 - i1) < e/4" "norm (s2 - i2) < e/4" for s1 s2 i1 and i2::'b
proof -
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: algebra_simps)
also have "\<dots> < e"
using no by (auto simp: algebra_simps norm_minus_commute)
finally show ?thesis .
qed
have finep: "gm fine \<D>" "gn fine \<D>"
using fine_Int \<D> by auto
have norm_le: "norm (g n x - g m x) \<le> 2 / real M" if x: "x \<in> cbox a b" for x
proof -
have "norm (f x - g n x) + norm (f x - g m x) \<le> 1 / (real n + 1) + 1 / (real m + 1)"
using g_near_f[OF x, of n] g_near_f[OF x, of m] by simp
also have "\<dots> \<le> 1 / (real M) + 1 / (real M)"
apply (rule add_mono)
using \<open>M \<noteq> 0\<close> m n by (auto simp: field_split_simps)
also have "\<dots> = 2 / real M"
by auto
finally show "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: algebra_simps simp add: norm_minus_commute)
qed
have "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g n x) - (\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g m x)) \<le> 2 / real M * content (cbox a b)"
by (blast intro: norm_le rsum_diff_bound[OF \<D>(1), where e="2 / real M"])
also have "... \<le> e/2"
using M True
by (auto simp: field_simps)
finally have le_e2: "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g n x) - (\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g m x)) \<le> e/2" .
then show "dist (h m) (h n) < e"
unfolding dist_norm using gm gn \<D> finep by (auto intro!: triangle3)
qed
qed
then obtain s where s: "h \<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 (h n - s) < e/3"
using LIMSEQ_D [OF s] by metis
from e True have "e/3 / content (cbox a b) > 0"
by (auto simp: field_simps)
then obtain N2 :: nat
where "N2 \<noteq> 0" and N2: "1 / (real N2) < e/3 / content (cbox a b)"
by (metis inverse_eq_divide real_arch_inverse)
obtain g' where "gauge g'"
and g': "\<And>\<D>. \<D> tagged_division_of cbox a b \<and> g' fine \<D> \<Longrightarrow>
norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g (N1 + N2) x) - h (N1 + N2)) < e/3"
by (metis h has_integral \<open>e/3 > 0\<close>)
have *: "norm (sf - s) < e"
if no: "norm (sf - sg) \<le> e/3" "norm(h - s) < e/3" "norm (sg - h) < e/3" for sf sg h
proof -
have "norm (sf - s) \<le> norm (sf - sg) + norm (sg - h) + norm (h - s)"
using norm_triangle_ineq[of "sf - sg" "sg - s"]
using norm_triangle_ineq[of "sg - h" " h - s"]
by (auto simp: algebra_simps)
also have "\<dots> < e"
using no by (auto simp: algebra_simps norm_minus_commute)
finally show ?thesis .
qed
{ fix \<D>
assume ptag: "\<D> tagged_division_of (cbox a b)" and "g' fine \<D>"
then have norm_less: "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g (N1 + N2) x) - h (N1 + N2)) < e/3"
using g' by blast
have "content (cbox a b) < e/3 * (of_nat N2)"
using \<open>N2 \<noteq> 0\<close> N2 using True by (auto simp: field_split_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: "1 / (real (N1 + N2) + 1) * content (cbox a b) \<le> e/3"
unfolding inverse_eq_divide
by (auto simp: field_simps)
have ne3: "norm (h (N1 + N2) - s) < e/3"
using N1 by auto
have "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - (\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g (N1 + N2) x))
\<le> 1 / (real (N1 + N2) + 1) * content (cbox a b)"
by (blast intro: g_near_f rsum_diff_bound[OF ptag])
then have "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - s) < e"
by (rule *[OF order_trans [OF _ le_e3] ne3 norm_less])
}
then show "\<exists>d. gauge d \<and>
(\<forall>\<D>. \<D> tagged_division_of cbox a b \<and> d fine \<D> \<longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - s) < e)"
by (blast intro: g' \<open>gauge g'\<close>)
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))"
subsubsection \<open>Negligibility of hyperplane\<close>
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
assume *: "a \<bullet> k \<le> c \<and> c \<le> b \<bullet> k \<and> (\<forall>j\<in>Basis. a \<bullet> j \<le> b \<bullet> j)"
define a' where "a' d = (\<Sum>j\<in>Basis. (if j = k then max (a\<bullet>j) (c - d) else a\<bullet>j) *\<^sub>R j)" for d
define b' where "b' d = (\<Sum>j\<in>Basis. (if j = k then min (b\<bullet>j) (c + d) else b\<bullet>j) *\<^sub>R j)" for d
have "((\<lambda>d. \<Prod>j\<in>Basis. (b' d - a' d) \<bullet> j) \<longlongrightarrow> (\<Prod>j\<in>Basis. (b' 0 - a' 0) \<bullet> j)) (at_right 0)"
by (auto simp: b'_def a'_def intro!: tendsto_min tendsto_max tendsto_eq_intros)
also have "(\<Prod>j\<in>Basis. (b' 0 - a' 0) \<bullet> j) = 0"
using k *
by (intro prod_zero bexI[OF _ k])
(auto simp: b'_def a'_def inner_diff inner_sum_left inner_not_same_Basis intro!: sum.cong)
also have "((\<lambda>d. \<Prod>j\<in>Basis. (b' d - a' d) \<bullet> j) \<longlongrightarrow> 0) (at_right 0) =
((\<lambda>d. content (cbox a b \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> d})) \<longlongrightarrow> 0) (at_right 0)"
proof (intro tendsto_cong eventually_at_rightI)
fix d :: real assume d: "d \<in> {0<..<1}"
have "cbox a b \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> d} = cbox (a' d) (b' d)" for d
using * d k by (auto simp add: cbox_def set_eq_iff Int_def ball_conj_distrib abs_diff_le_iff a'_def b'_def)
moreover have "j \<in> Basis \<Longrightarrow> a' d \<bullet> j \<le> b' d \<bullet> j" for j
using * d k by (auto simp: a'_def b'_def)
ultimately show "(\<Prod>j\<in>Basis. (b' d - a' d) \<bullet> j) = content (cbox a b \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> d})"
by simp
qed simp
finally have "((\<lambda>d. content (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})) \<longlongrightarrow> 0) (at_right 0)" .
from order_tendstoD(2)[OF this \<open>0<e\<close>]
obtain d' where "0 < d'" and d': "\<And>y. y > 0 \<Longrightarrow> y < d' \<Longrightarrow> content (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> y}) < e"
by (subst (asm) eventually_at_right[of _ 1]) auto
show ?thesis
by (rule that[of "d'/2"], insert \<open>0<d'\<close> d'[of "d'/2"], auto)
next
assume *: "\<not> (a \<bullet> k \<le> c \<and> c \<le> b \<bullet> k \<and> (\<forall>j\<in>Basis. a \<bullet> j \<le> b \<bullet> j))"
then have "(\<exists>j\<in>Basis. b \<bullet> j < a \<bullet> j) \<or> (c < a \<bullet> k \<or> b \<bullet> k < c)"
by (auto simp: not_le)
show thesis
proof cases
assume "\<exists>j\<in>Basis. b \<bullet> j < a \<bullet> j"
then have [simp]: "cbox a b = {}"
using box_ne_empty(1)[of a b] by auto
show ?thesis
by (rule that[of 1]) (simp_all add: \<open>0<e\<close>)
next
assume "\<not> (\<exists>j\<in>Basis. b \<bullet> j < a \<bullet> j)"
with * have "c < a \<bullet> k \<or> b \<bullet> k < c"
by auto
then show thesis
proof
assume c: "c < a \<bullet> k"
moreover have "x \<in> cbox a b \<Longrightarrow> c \<le> x \<bullet> k" for x
using k c by (auto simp: cbox_def)
ultimately have "cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> (a \<bullet> k - c)/2} = {}"
using k by (auto simp: cbox_def)
with \<open>0<e\<close> c that[of "(a \<bullet> k - c)/2"] show ?thesis
by auto
next
assume c: "b \<bullet> k < c"
moreover have "x \<in> cbox a b \<Longrightarrow> x \<bullet> k \<le> c" for x
using k c by (auto simp: cbox_def)
ultimately have "cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> (c - b \<bullet> k)/2} = {}"
using k by (auto simp: cbox_def)
with \<open>0<e\<close> c that[of "(c - b \<bullet> k)/2"] show ?thesis
by auto
qed
qed
qed
proposition 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 clarsimp
fix a b and e::real assume "e > 0"
with k obtain d where "0 < d" and d: "content (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) < e"
by (metis content_doublesplit)
let ?i = "indicator {x::'a. x\<bullet>k = c} :: 'a\<Rightarrow>real"
show "\<exists>\<gamma>. gauge \<gamma> \<and>
(\<forall>\<D>. \<D> tagged_division_of cbox a b \<and> \<gamma> fine \<D> \<longrightarrow>
\<bar>\<Sum>(x,K) \<in> \<D>. content K * ?i x\<bar> < e)"
proof (intro exI, safe)
show "gauge (\<lambda>x. ball x d)"
using \<open>0 < d\<close> by blast
next
fix \<D>
assume p: "\<D> tagged_division_of (cbox a b)" "(\<lambda>x. ball x d) fine \<D>"
have "content L = content (L \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})"
if "(x, L) \<in> \<D>" "?i x \<noteq> 0" for x L
proof -
have xk: "x\<bullet>k = c"
using that by (simp add: indicator_def split: if_split_asm)
have "L \<subseteq> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}"
proof
fix y
assume y: "y \<in> L"
have "L \<subseteq> ball x d"
using p(2) that(1) by auto
then have "norm (x - y) < d"
by (simp add: dist_norm subset_iff y)
then have "\<bar>(x - y) \<bullet> k\<bar> < d"
using k norm_bound_Basis_lt by blast
then show "y \<in> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}"
unfolding inner_simps xk by auto
qed
then show "content L = content (L \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})"
by (metis inf.orderE)
qed
then have *: "(\<Sum>(x,K)\<in>\<D>. content K * ?i x) = (\<Sum>(x,K)\<in>\<D>. content (K \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> d}) *\<^sub>R ?i x)"
by (force simp add: split_paired_all intro!: sum.cong [OF refl])
note p'= tagged_division_ofD[OF p(1)] and p''=division_of_tagged_division[OF p(1)]
have "(\<Sum>(x,K)\<in>\<D>. content (K \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) * indicator {x. x \<bullet> k = c} x) < e"
proof -
have "(\<Sum>(x,K)\<in>\<D>. content (K \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) * ?i x) \<le> (\<Sum>(x,K)\<in>\<D>. content (K \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}))"
by (force simp add: indicator_def intro!: sum_mono)
also have "\<dots> < e"
proof (subst sum.over_tagged_division_lemma[OF p(1)])
fix u v::'a
assume "box u v = {}"
then have *: "content (cbox u v) = 0"
unfolding content_eq_0_interior by simp
have "cbox u v \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<subseteq> cbox u v"
by auto
then 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] by (rule content_subset)
then show "content (cbox u v \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) = 0"
unfolding * interval_doublesplit[OF k]
by (blast intro: antisym)
next
have "(\<Sum>l\<in>snd ` \<D>. content (l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})) =
sum content ((\<lambda>l. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})`{l\<in>snd ` \<D>. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}})"
proof (subst (2) sum.reindex_nontrivial)
fix x y assume "x \<in> {l \<in> snd ` \<D>. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}}" "y \<in> {l \<in> snd ` \<D>. 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> \<D>" "x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}" "(y', y) \<in> \<D>" "y \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}"
by (auto)
from p'(5)[OF \<open>(x', x) \<in> \<D>\<close> \<open>(y', y) \<in> \<D>\<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> \<D>\<close>] by (auto simp: interval_doublesplit[OF k] content_eq_0_interior simp del: interior_Int)
qed (insert p'(1), auto intro!: sum.mono_neutral_right)
also have "\<dots> \<le> norm (\<Sum>l\<in>(\<lambda>l. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})`{l\<in>snd ` \<D>. 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 by simp
finally show "(\<Sum>K\<in>snd ` \<D>. content (K \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})) < e" .
qed
finally show "(\<Sum>(x, K)\<in>\<D>. content (K \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) * ?i x) < e" .
qed
then show "\<bar>\<Sum>(x, K)\<in>\<D>. content K * ?i x\<bar> < e"
unfolding *
apply (subst abs_of_nonneg)
using measure_nonneg
by (force simp add: indicator_def intro: sum_nonneg)+
qed
qed
corollary negligible_standard_hyperplane_cart:
fixes k :: "'a::finite"
shows "negligible {x. x$k = (0::real)}"
by (simp add: cart_eq_inner_axis negligible_standard_hyperplane)
subsubsection \<open>Hence the main theorem about negligible sets\<close>
lemma has_integral_negligible_cbox:
fixes f :: "'b::euclidean_space \<Rightarrow> 'a::real_normed_vector"
assumes negs: "negligible S"
and 0: "\<And>x. x \<notin> S \<Longrightarrow> f x = 0"
shows "(f has_integral 0) (cbox a b)"
unfolding has_integral
proof clarify
fix e::real
assume "e > 0"
then have nn_gt0: "e/2 / ((real n+1) * (2 ^ n)) > 0" for n
by simp
then have "\<exists>\<gamma>. gauge \<gamma> \<and>
(\<forall>\<D>. \<D> tagged_division_of cbox a b \<and> \<gamma> fine \<D> \<longrightarrow>
\<bar>\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R indicator S x\<bar>
< e/2 / ((real n + 1) * 2 ^ n))" for n
using negs [unfolded negligible_def has_integral] by auto
then obtain \<gamma> where
gd: "\<And>n. gauge (\<gamma> n)"
and \<gamma>: "\<And>n \<D>. \<lbrakk>\<D> tagged_division_of cbox a b; \<gamma> n fine \<D>\<rbrakk>
\<Longrightarrow> \<bar>\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R indicator S x\<bar> < e/2 / ((real n + 1) * 2 ^ n)"
by metis
show "\<exists>\<gamma>. gauge \<gamma> \<and>
(\<forall>\<D>. \<D> tagged_division_of cbox a b \<and> \<gamma> fine \<D> \<longrightarrow>
norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - 0) < e)"
proof (intro exI, safe)
show "gauge (\<lambda>x. \<gamma> (nat \<lfloor>norm (f x)\<rfloor>) x)"
using gd by (auto simp: gauge_def)
show "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - 0) < e"
if "\<D> tagged_division_of (cbox a b)" "(\<lambda>x. \<gamma> (nat \<lfloor>norm (f x)\<rfloor>) x) fine \<D>" for \<D>
proof (cases "\<D> = {}")
case True with \<open>0 < e\<close> show ?thesis by simp
next
case False
obtain N where "Max ((\<lambda>(x, K). norm (f x)) ` \<D>) \<le> real N"
using real_arch_simple by blast
then have N: "\<And>x. x \<in> (\<lambda>(x, K). norm (f x)) ` \<D> \<Longrightarrow> x \<le> real N"
by (meson Max_ge that(1) dual_order.trans finite_imageI tagged_division_of_finite)
have "\<forall>i. \<exists>q. q tagged_division_of (cbox a b) \<and> (\<gamma> i) fine q \<and> (\<forall>(x,K) \<in> \<D>. K \<subseteq> (\<gamma> i) x \<longrightarrow> (x, K) \<in> q)"
by (auto intro: tagged_division_finer[OF that(1) gd])
from choice[OF this]
obtain q where q: "\<And>n. q n tagged_division_of cbox a b"
"\<And>n. \<gamma> n fine q n"
"\<And>n x K. \<lbrakk>(x, K) \<in> \<D>; K \<subseteq> \<gamma> n x\<rbrakk> \<Longrightarrow> (x, K) \<in> q n"
by fastforce
have "finite \<D>"
using that(1) by blast
then have sum_le_inc: "\<lbrakk>finite T; \<And>x y. (x,y) \<in> T \<Longrightarrow> (0::real) \<le> g(x,y);
\<And>y. y\<in>\<D> \<Longrightarrow> \<exists>x. (x,y) \<in> T \<and> f(y) \<le> g(x,y)\<rbrakk> \<Longrightarrow> sum f \<D> \<le> sum g T" for f g T
by (rule sum_le_included[of \<D> T g snd f]; force)
have "norm (\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) \<le> (\<Sum>(x,K) \<in> \<D>. norm (content K *\<^sub>R f x))"
unfolding split_def by (rule norm_sum)
also have "... \<le> (\<Sum>(i, j) \<in> Sigma {..N + 1} q.
(real i + 1) * (case j of (x, K) \<Rightarrow> content K *\<^sub>R indicator S x))"
proof (rule sum_le_inc, safe)
show "finite (Sigma {..N+1} q)"
by (meson finite_SigmaI finite_atMost tagged_division_of_finite q(1))
next
fix x K
assume xk: "(x, K) \<in> \<D>"
define n where "n = nat \<lfloor>norm (f x)\<rfloor>"
have *: "norm (f x) \<in> (\<lambda>(x, K). norm (f x)) ` \<D>"
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[OF *] by auto
moreover have "K \<subseteq> \<gamma> (nat \<lfloor>norm (f x)\<rfloor>) x"
using that(2) xk by auto
moreover then have "(x, K) \<in> q (nat \<lfloor>norm (f x)\<rfloor>)"
by (simp add: q(3) xk)
moreover then have "(x, K) \<in> q n"
using n_def by blast
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 by (simp add: 0)
next
case True
have *: "content K \<ge> 0"
using tagged_division_ofD(4)[OF that(1) xk] by auto
moreover have "content K * norm (f x) \<le> content K * (real n + 1)"
by (simp add: mult_left_mono nfx(2))
ultimately show ?thesis
using nfx True by (auto simp: field_simps)
qed
ultimately show "\<exists>y. (y, x, K) \<in> (Sigma {..N + 1} q) \<and> norm (content K *\<^sub>R f x) \<le>
(real y + 1) * (content K *\<^sub>R indicator S x)"
by force
qed auto
also have "... = (\<Sum>i\<le>N + 1. \<Sum>j\<in>q i. (real i + 1) * (case j of (x, K) \<Rightarrow> content K *\<^sub>R indicator S x))"
apply (rule sum_Sigma_product [symmetric])
using q(1) apply auto
done
also have "... \<le> (\<Sum>i\<le>N + 1. (real i + 1) * \<bar>\<Sum>(x,K) \<in> q i. content K *\<^sub>R indicator S x\<bar>)"
by (rule sum_mono) (simp add: sum_distrib_left [symmetric])
also have "... \<le> (\<Sum>i\<le>N + 1. e/2/2 ^ i)"
proof (rule sum_mono)
show "(real i + 1) * \<bar>\<Sum>(x,K) \<in> q i. content K *\<^sub>R indicator S x\<bar> \<le> e/2/2 ^ i"
if "i \<in> {..N + 1}" for i
using \<gamma>[of "q i" i] q by (simp add: divide_simps mult.left_commute)
qed
also have "... = e/2 * (\<Sum>i\<le>N + 1. (1/2) ^ i)"
unfolding sum_distrib_left by (metis divide_inverse inverse_eq_divide power_one_over)
also have "\<dots> < e/2 * 2"
proof (rule mult_strict_left_mono)
have "sum (power (1/2)) {..N + 1} = sum (power (1/2::real)) {..<N + 2}"
using lessThan_Suc_atMost by auto
also have "... < 2"
by (auto simp: geometric_sum)
finally show "sum (power (1/2::real)) {..N + 1} < 2" .
qed (use \<open>0 < e\<close> in auto)
finally show ?thesis by auto
qed
qed
qed
proposition has_integral_negligible:
fixes f :: "'b::euclidean_space \<Rightarrow> 'a::real_normed_vector"
assumes negs: "negligible S"
and "\<And>x. x \<in> (T - S) \<Longrightarrow> f x = 0"
shows "(f has_integral 0) T"
proof (cases "\<exists>a b. T = cbox a b")
case True
then have "((\<lambda>x. if x \<in> T then f x else 0) has_integral 0) T"
using assms by (auto intro!: has_integral_negligible_cbox)
then show ?thesis
by (rule has_integral_eq [rotated]) auto
next
case False
let ?f = "(\<lambda>x. if x \<in> T then f x else 0)"
have "((\<lambda>x. if x \<in> T then f x else 0) has_integral 0) T"
apply (auto simp: False has_integral_alt [of ?f])
apply (rule_tac x=1 in exI, auto)
apply (rule_tac x=0 in exI, simp add: has_integral_negligible_cbox [OF negs] assms)
done
then show ?thesis
by (rule_tac f="?f" in has_integral_eq) auto
qed
lemma
assumes "negligible S"
shows integrable_negligible: "f integrable_on S" and integral_negligible: "integral S f = 0"
using has_integral_negligible [OF assms]
by (auto simp: has_integral_iff)
lemma has_integral_spike:
fixes f :: "'b::euclidean_space \<Rightarrow> 'a::real_normed_vector"
assumes "negligible S"
and gf: "\<And>x. x \<in> T - S \<Longrightarrow> g x = f x"
and fint: "(f has_integral y) T"
shows "(g has_integral y) T"
proof -
have *: "(g has_integral y) (cbox a b)"
if "(f has_integral y) (cbox a b)" "\<forall>x \<in> cbox a b - S. g x = f x" for a b f and g:: "'b \<Rightarrow> 'a" and y
proof -
have "((\<lambda>x. f x + (g x - f x)) has_integral (y + 0)) (cbox a b)"
using that by (intro has_integral_add has_integral_negligible) (auto intro!: \<open>negligible S\<close>)
then show ?thesis
by auto
qed
show ?thesis
using fint gf
apply (subst has_integral_alt)
apply (subst (asm) has_integral_alt)
apply (simp split: if_split_asm)
apply (blast dest: *)
apply (erule_tac V = "\<forall>a b. T \<noteq> cbox a b" in thin_rl)
apply (elim all_forward imp_forward ex_forward all_forward conj_forward asm_rl)
apply (auto dest!: *[where f="\<lambda>x. if x\<in>T then f x else 0" and g="\<lambda>x. if x \<in> T then g x else 0"])
done
qed
lemma has_integral_spike_eq:
assumes "negligible S"
and gf: "\<And>x. x \<in> T - S \<Longrightarrow> g x = f x"
shows "(f has_integral y) T \<longleftrightarrow> (g has_integral y) T"
using has_integral_spike [OF \<open>negligible S\<close>] gf
by metis
lemma integrable_spike:
assumes "f integrable_on T" "negligible S" "\<And>x. x \<in> T - S \<Longrightarrow> g x = f x"
shows "g integrable_on T"
using assms unfolding integrable_on_def by (blast intro: has_integral_spike)
lemma integral_spike:
assumes "negligible S"
and "\<And>x. x \<in> T - S \<Longrightarrow> g x = f x"
shows "integral T f = integral T g"
using has_integral_spike_eq[OF assms]
by (auto simp: integral_def integrable_on_def)
subsection \<open>Some other trivialities about negligible sets\<close>
lemma negligible_subset:
assumes "negligible s" "t \<subseteq> s"
shows "negligible t"
unfolding negligible_def
by (metis (no_types) Diff_iff assms contra_subsetD has_integral_negligible indicator_simps(2))
lemma negligible_diff[intro?]:
assumes "negligible s"
shows "negligible (s - t)"
using assms by (meson Diff_subset negligible_subset)
lemma negligible_Int:
assumes "negligible s \<or> negligible t"
shows "negligible (s \<inter> t)"
using assms negligible_subset by force
lemma negligible_Un:
assumes "negligible S" and T: "negligible T"
shows "negligible (S \<union> T)"
proof -
have "(indicat_real (S \<union> T) has_integral 0) (cbox a b)"
if S0: "(indicat_real S has_integral 0) (cbox a b)"
and "(indicat_real T has_integral 0) (cbox a b)" for a b
proof (subst has_integral_spike_eq[OF T])
show "indicat_real S x = indicat_real (S \<union> T) x" if "x \<in> cbox a b - T" for x
by (metis Diff_iff Un_iff indicator_def that)
show "(indicat_real S has_integral 0) (cbox a b)"
by (simp add: S0)
qed
with assms show ?thesis
unfolding negligible_def by blast
qed
lemma negligible_Un_eq[simp]: "negligible (s \<union> t) \<longleftrightarrow> negligible s \<and> negligible t"
using negligible_Un negligible_subset by blast
lemma negligible_sing[intro]: "negligible {a::'a::euclidean_space}"
using negligible_standard_hyperplane[OF SOME_Basis, of "a \<bullet> (SOME i. i \<in> Basis)"] negligible_subset by blast
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 {}"
using negligible_insert by blast
text\<open>Useful in this form for backchaining\<close>
lemma empty_imp_negligible: "S = {} \<Longrightarrow> negligible S"
by simp
lemma negligible_finite[intro]:
assumes "finite s"
shows "negligible s"
using assms by (induct s) auto
lemma negligible_Union[intro]:
assumes "finite \<T>"
and "\<And>t. t \<in> \<T> \<Longrightarrow> negligible t"
shows "negligible(\<Union>\<T>)"
using assms by induct auto
lemma negligible: "negligible S \<longleftrightarrow> (\<forall>T. (indicat_real S has_integral 0) T)"
proof (intro iffI allI)
fix T
assume "negligible S"
then show "(indicator S has_integral 0) T"
by (meson Diff_iff has_integral_negligible indicator_simps(2))
qed (simp add: negligible_def)
subsection \<open>Finite case of the spike theorem is quite commonly needed\<close>
lemma has_integral_spike_finite:
assumes "finite S"
and "\<And>x. x \<in> T - S \<Longrightarrow> g x = f x"
and "(f has_integral y) T"
shows "(g has_integral y) T"
using assms has_integral_spike negligible_finite by blast
lemma has_integral_spike_finite_eq:
assumes "finite S"
and "\<And>x. x \<in> T - S \<Longrightarrow> g x = f x"
shows "((f has_integral y) T \<longleftrightarrow> (g has_integral y) T)"
by (metis assms has_integral_spike_finite)
lemma integrable_spike_finite:
assumes "finite S"
and "\<And>x. x \<in> T - S \<Longrightarrow> g x = f x"
and "f integrable_on T"
shows "g integrable_on T"
using assms has_integral_spike_finite by blast
lemma has_integral_bound_spike_finite:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
assumes "0 \<le> B" "finite S"
and f: "(f has_integral i) (cbox a b)"
and leB: "\<And>x. x \<in> cbox a b - S \<Longrightarrow> norm (f x) \<le> B"
shows "norm i \<le> B * content (cbox a b)"
proof -
define g where "g \<equiv> (\<lambda>x. if x \<in> S then 0 else f x)"
then have "\<And>x. x \<in> cbox a b - S \<Longrightarrow> norm (g x) \<le> B"
using leB by simp
moreover have "(g has_integral i) (cbox a b)"
using has_integral_spike_finite [OF \<open>finite S\<close> _ f]
by (simp add: g_def)
ultimately show ?thesis
by (simp add: \<open>0 \<le> B\<close> g_def has_integral_bound)
qed
corollary has_integral_bound_real:
fixes f :: "real \<Rightarrow> 'b::real_normed_vector"
assumes "0 \<le> B" "finite S"
and "(f has_integral i) {a..b}"
and "\<And>x. x \<in> {a..b} - S \<Longrightarrow> norm (f x) \<le> B"
shows "norm i \<le> B * content {a..b}"
by (metis assms box_real(2) has_integral_bound_spike_finite)
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 "negligible ?A"
by (force simp add: negligible_Union[OF finite_imageI])
moreover have "cbox a b - box a b \<subseteq> ?A"
by (force simp add: mem_box)
ultimately show ?thesis
by (rule negligible_subset)
qed
lemma has_integral_spike_interior:
assumes f: "(f has_integral y) (cbox a b)" and gf: "\<And>x. x \<in> box a b \<Longrightarrow> g x = f x"
shows "(g has_integral y) (cbox a b)"
apply (rule has_integral_spike[OF negligible_frontier_interval _ f])
using gf by auto
lemma has_integral_spike_interior_eq:
assumes "\<And>x. x \<in> box a b \<Longrightarrow> g x = f x"
shows "(f has_integral y) (cbox a b) \<longleftrightarrow> (g has_integral y) (cbox a b)"
by (metis assms has_integral_spike_interior)
lemma integrable_spike_interior:
assumes "\<And>x. x \<in> box a b \<Longrightarrow> g x = f x"
and "f integrable_on cbox a b"
shows "g integrable_on cbox a b"
using assms has_integral_spike_interior_eq by blast
subsection \<open>Integrability of continuous functions\<close>
lemma operative_approximableI:
fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach"
assumes "0 \<le> e"
shows "operative conj True (\<lambda>i. \<exists>g. (\<forall>x\<in>i. norm (f x - g (x::'b)) \<le> e) \<and> g integrable_on i)"
proof -
interpret comm_monoid conj True
by standard auto
show ?thesis
proof (standard, 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 "box a b = {}" for a b
apply (rule_tac x=f in exI)
using assms that by (auto simp: content_eq_0_interior)
{
fix c g and k :: 'b
assume fg: "\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e" and g: "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 fg integrable_split[OF g k] by auto
}
show "\<exists>g. (\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b"
if fg1: "\<forall>x\<in>cbox a b \<inter> {x. x \<bullet> k \<le> c}. norm (f x - g1 x) \<le> e"
and g1: "g1 integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c}"
and fg2: "\<forall>x\<in>cbox a b \<inter> {x. c \<le> x \<bullet> k}. norm (f x - g2 x) \<le> e"
and g2: "g2 integrable_on cbox a b \<inter> {x. c \<le> x \<bullet> k}"
and k: "k \<in> Basis"
for c k g1 g2
proof -
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"
proof (intro exI conjI ballI)
show "norm (f x - ?g x) \<le> e" if "x \<in> cbox a b" for x
by (auto simp: that assms fg1 fg2)
show "?g integrable_on cbox a b"
proof -
have "?g integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c}" "?g integrable_on cbox a b \<inter> {x. x \<bullet> k \<ge> c}"
by(rule integrable_spike[OF _ negligible_standard_hyperplane[of k c]], use k g1 g2 in auto)+
with has_integral_split[OF _ _ k] show ?thesis
unfolding integrable_on_def by blast
qed
qed
qed
qed
qed
lemma comm_monoid_set_F_and: "comm_monoid_set.F (\<and>) True f s \<longleftrightarrow> (finite s \<longrightarrow> (\<forall>x\<in>s. f x))"
proof -
interpret bool: comm_monoid_set "(\<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: "d division_of (cbox a b)"
and f: "\<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 -
interpret operative conj True "\<lambda>i. \<exists>g. (\<forall>x\<in>i. norm (f x - g (x::'b)) \<le> e) \<and> g integrable_on i"
using \<open>0 \<le> e\<close> by (rule operative_approximableI)
from f local.division [OF d] that show thesis
by auto
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)
fix e :: real
assume e: "e > 0"
then obtain d where "0 < d" and d: "\<And>x x'. \<lbrakk>x \<in> cbox a b; x' \<in> cbox a b; dist x' x < d\<rbrakk> \<Longrightarrow> dist (f x') (f x) < e"
using compact_uniformly_continuous[OF assms compact_cbox] unfolding uniformly_continuous_on_def by metis
obtain p where ptag: "p tagged_division_of cbox a b" and finep: "(\<lambda>x. ball x d) fine p"
using fine_division_exists[OF gauge_ball[OF \<open>0 < d\<close>], of a b] .
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"
obtain a b where l: "l = cbox a b"
using as ptag by blast
then have x: "x \<in> cbox a b"
using as ptag by auto
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 by blast
next
fix y
assume y: "y \<in> l"
then have "y \<in> ball x d"
using as finep by fastforce
then show "norm (f y - f x) \<le> e"
using d x y as l
by (metis dist_commute dist_norm less_imp_le mem_ball ptag subsetCE tagged_division_ofD(3))
qed
qed
from e have "e \<ge> 0"
by auto
from approximable_on_division[OF this division_of_tagged_division[OF ptag] *]
show "\<exists>g. (\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b"
by metis
qed
lemma integrable_continuous_interval:
fixes f :: "'b::ordered_euclidean_space \<Rightarrow> 'a::banach"
assumes "continuous_on {a..b} f"
shows "f integrable_on {a..b}"
by (metis assms integrable_continuous interval_cbox)
lemmas integrable_continuous_real = integrable_continuous_interval[where 'b=real]
lemma integrable_continuous_closed_segment:
fixes f :: "real \<Rightarrow> 'a::banach"
assumes "continuous_on (closed_segment a b) f"
shows "f integrable_on (closed_segment a b)"
using assms
by (auto intro!: integrable_continuous_interval simp: closed_segment_eq_real_ivl)
subsection \<open>Specialization of additivity to one dimension\<close>
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 (sum (\<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
have "\<And>e p. p tagged_division_of cbox a b \<Longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)) \<le> e * content (cbox a b)"
unfolding sum_content_null[OF True] True by force
moreover have "i = 0"
if "\<And>e. e > 0 \<Longrightarrow> \<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) - i) \<le> e * content (cbox a b))"
using that [of 1]
by (force simp add: True sum_content_null[OF True] intro: fine_division_exists[of _ a b])
ultimately show ?thesis
unfolding has_integral_null_eq[OF True]
by (force simp add: )
next
case False
then have F: "0 < content (cbox a b)"
using zero_less_measure_iff by blast
let ?P = "\<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 (<)"
then show "?P (e * content (cbox a b)) (\<le>)"
apply (rule allE [where x="e * content (cbox a b)"])
apply (elim impE ex_forward conj_forward)
using F e apply (auto simp add: algebra_simps)
done }
{ assume "\<forall>e>0. ?P (e * content (cbox a b)) (\<le>)"
then show "?P e (<)"
apply (rule allE [where x="e/2 / content (cbox a b)"])
apply (elim impE ex_forward conj_forward)
using F e apply (auto simp add: algebra_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 (sum (\<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
theorem fundamental_theorem_of_calculus:
fixes f :: "real \<Rightarrow> 'a::banach"
assumes "a \<le> b"
and vecd: "\<And>x. x \<in> {a..b} \<Longrightarrow> (f has_vector_derivative f' x) (at x within {a..b})"
shows "(f' has_integral (f b - f a)) {a..b}"
unfolding has_integral_factor_content box_real[symmetric]
proof safe
fix e :: real
assume "e > 0"
then have "\<forall>x. \<exists>d>0. x \<in> {a..b} \<longrightarrow>
(\<forall>y\<in>{a..b}. norm (y-x) < d \<longrightarrow> norm (f y - f x - (y-x) *\<^sub>R f' x) \<le> e * norm (y-x))"
using vecd[unfolded has_vector_derivative_def has_derivative_within_alt] by blast
then obtain d where d: "\<And>x. 0 < d x"
"\<And>x y. \<lbrakk>x \<in> {a..b}; y \<in> {a..b}; norm (y-x) < d x\<rbrakk>
\<Longrightarrow> norm (f y - f x - (y-x) *\<^sub>R f' x) \<le> e * norm (y-x)"
by metis
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))"
proof (rule exI, safe)
show "gauge (\<lambda>x. ball x (d x))"
using d(1) gauge_ball_dependent by blast
next
fix p
assume ptag: "p tagged_division_of cbox a b" and finep: "(\<lambda>x. ball x (d x)) fine p"
have ba: "b - a = (\<Sum>(x,K)\<in>p. Sup K - Inf K)" "f b - f a = (\<Sum>(x,K)\<in>p. f(Sup K) - f(Inf K))"
using additive_tagged_division_1[where f= "\<lambda>x. x"] additive_tagged_division_1[where f= f]
\<open>a \<le> b\<close> ptag by auto
have "norm (\<Sum>(x, K) \<in> p. (content K *\<^sub>R f' x) - (f (Sup K) - f (Inf K)))
\<le> (\<Sum>n\<in>p. e * (case n of (x, k) \<Rightarrow> Sup k - Inf k))"
proof (rule sum_norm_le,safe)
fix x K
assume "(x, K) \<in> p"
then have "x \<in> K" and kab: "K \<subseteq> cbox a b"
using ptag by blast+
then obtain u v where k: "K = cbox u v" and "x \<in> K" and kab: "K \<subseteq> cbox a b"
using ptag \<open>(x, K) \<in> p\<close> by auto
have "u \<le> v"
using \<open>x \<in> K\<close> unfolding k by auto
have ball: "\<forall>y\<in>K. y \<in> ball x (d x)"
using finep \<open>(x, K) \<in> p\<close> by blast
have "u \<in> K" "v \<in> K"
by (simp_all add: \<open>u \<le> v\<close> k)
have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) = norm (f u - f x - (u - x) *\<^sub>R f' x - (f v - f x - (v - x) *\<^sub>R f' x))"
by (auto simp add: algebra_simps)
also have "... \<le> norm (f u - f x - (u - x) *\<^sub>R f' x) + norm (f v - f x - (v - x) *\<^sub>R f' x)"
by (rule norm_triangle_ineq4)
finally have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) \<le>
norm (f u - f x - (u - x) *\<^sub>R f' x) + norm (f v - f x - (v - x) *\<^sub>R f' x)" .
also have "\<dots> \<le> e * norm (u - x) + e * norm (v - x)"
proof (rule add_mono)
show "norm (f u - f x - (u - x) *\<^sub>R f' x) \<le> e * norm (u - x)"
apply (rule d(2)[of x u])
using \<open>x \<in> K\<close> kab \<open>u \<in> K\<close> ball dist_real_def by (auto simp add:dist_real_def)
show "norm (f v - f x - (v - x) *\<^sub>R f' x) \<le> e * norm (v - x)"
apply (rule d(2)[of x v])
using \<open>x \<in> K\<close> kab \<open>v \<in> K\<close> ball dist_real_def by (auto simp add:dist_real_def)
qed
also have "\<dots> \<le> e * (Sup K - Inf K)"
using \<open>x \<in> K\<close> by (auto simp: k interval_bounds_real[OF \<open>u \<le> v\<close>] field_simps)
finally show "norm (content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \<le> e * (Sup K - Inf K)"
using \<open>u \<le> v\<close> by (simp add: k)
qed
with \<open>a \<le> b\<close> show "norm ((\<Sum>(x, K)\<in>p. content K *\<^sub>R f' x) - (f b - f a)) \<le> e * content (cbox a b)"
by (auto simp: ba split_def sum_subtractf [symmetric] sum_distrib_left)
qed
qed
lemma ident_has_integral:
fixes a::real
assumes "a \<le> b"
shows "((\<lambda>x. x) has_integral (b\<^sup>2 - a\<^sup>2)/2) {a..b}"
proof -
have "((\<lambda>x. x) has_integral inverse 2 * b\<^sup>2 - inverse 2 * a\<^sup>2) {a..b}"
apply (rule fundamental_theorem_of_calculus [OF assms])
unfolding power2_eq_square
by (rule derivative_eq_intros | simp)+
then show ?thesis
by (simp add: field_simps)
qed
lemma integral_ident [simp]:
fixes a::real
assumes "a \<le> b"
shows "integral {a..b} (\<lambda>x. x) = (if a \<le> b then (b\<^sup>2 - a\<^sup>2)/2 else 0)"
by (metis assms ident_has_integral integral_unique)
lemma ident_integrable_on:
fixes a::real
shows "(\<lambda>x. x) integrable_on {a..b}"
by (metis atLeastatMost_empty_iff integrable_on_def has_integral_empty ident_has_integral)
lemma integral_sin [simp]:
fixes a::real
assumes "a \<le> b" shows "integral {a..b} sin = cos a - cos b"
proof -
have "(sin has_integral (- cos b - - cos a)) {a..b}"
proof (rule fundamental_theorem_of_calculus)
show "((\<lambda>a. - cos a) has_vector_derivative sin x) (at x within {a..b})" for x
unfolding has_field_derivative_iff_has_vector_derivative [symmetric]
by (rule derivative_eq_intros | force)+
qed (use assms in auto)
then show ?thesis
by (simp add: integral_unique)
qed
lemma integral_cos [simp]:
fixes a::real
assumes "a \<le> b" shows "integral {a..b} cos = sin b - sin a"
proof -
have "(cos has_integral (sin b - sin a)) {a..b}"
proof (rule fundamental_theorem_of_calculus)
show "(sin has_vector_derivative cos x) (at x within {a..b})" for x
unfolding has_field_derivative_iff_has_vector_derivative [symmetric]
by (rule derivative_eq_intros | force)+
qed (use assms in auto)
then show ?thesis
by (simp add: integral_unique)
qed
lemma has_integral_sin_nx: "((\<lambda>x. sin(real_of_int n * x)) has_integral 0) {-pi..pi}"
proof (cases "n = 0")
case False
have "((\<lambda>x. sin (n * x)) has_integral (- cos (n * pi)/n - - cos (n * - pi)/n)) {-pi..pi}"
proof (rule fundamental_theorem_of_calculus)
show "((\<lambda>x. - cos (n * x) / n) has_vector_derivative sin (n * a)) (at a within {-pi..pi})"
if "a \<in> {-pi..pi}" for a :: real
using that False
apply (simp only: has_vector_derivative_def)
apply (rule derivative_eq_intros | force)+
done
qed auto
then show ?thesis
by simp
qed auto
lemma integral_sin_nx:
"integral {-pi..pi} (\<lambda>x. sin(x * real_of_int n)) = 0"
using has_integral_sin_nx [of n] by (force simp: mult.commute)
lemma has_integral_cos_nx:
"((\<lambda>x. cos(real_of_int n * x)) has_integral (if n = 0 then 2 * pi else 0)) {-pi..pi}"
proof (cases "n = 0")
case True
then show ?thesis
using has_integral_const_real [of "1::real" "-pi" pi] by auto
next
case False
have "((\<lambda>x. cos (n * x)) has_integral (sin (n * pi)/n - sin (n * - pi)/n)) {-pi..pi}"
proof (rule fundamental_theorem_of_calculus)
show "((\<lambda>x. sin (n * x) / n) has_vector_derivative cos (n * x)) (at x within {-pi..pi})"
if "x \<in> {-pi..pi}"
for x :: real
using that False
apply (simp only: has_vector_derivative_def)
apply (rule derivative_eq_intros | force)+
done
qed auto
with False show ?thesis
by (simp add: mult.commute)
qed
lemma integral_cos_nx:
"integral {-pi..pi} (\<lambda>x. cos(x * real_of_int n)) = (if n = 0 then 2 * pi else 0)"
using has_integral_cos_nx [of n] by (force simp: mult.commute)
subsection \<open>Taylor series expansion\<close>
lemma mvt_integral:
fixes f::"'a::real_normed_vector\<Rightarrow>'b::banach"
assumes f'[derivative_intros]:
"\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
assumes line_in: "\<And>t. t \<in> {0..1} \<Longrightarrow> x + t *\<^sub>R y \<in> S"
shows "f (x + y) - f x = integral {0..1} (\<lambda>t. f' (x + t *\<^sub>R y) y)" (is ?th1)
proof -
from assms have subset: "(\<lambda>xa. x + xa *\<^sub>R y) ` {0..1} \<subseteq> S" by auto
note [derivative_intros] =
has_derivative_subset[OF _ subset]
has_derivative_in_compose[where f="(\<lambda>xa. x + xa *\<^sub>R y)" and g = f]
note [continuous_intros] =
continuous_on_compose2[where f="(\<lambda>xa. x + xa *\<^sub>R y)"]
continuous_on_subset[OF _ subset]
have "\<And>t. t \<in> {0..1} \<Longrightarrow>
((\<lambda>t. f (x + t *\<^sub>R y)) has_vector_derivative f' (x + t *\<^sub>R y) y)
(at t within {0..1})"
using assms
by (auto simp: has_vector_derivative_def
linear_cmul[OF has_derivative_linear[OF f'], symmetric]
intro!: derivative_eq_intros)
from fundamental_theorem_of_calculus[rule_format, OF _ this]
show ?th1
by (auto intro!: integral_unique[symmetric])
qed
lemma (in bounded_bilinear) sum_prod_derivatives_has_vector_derivative:
assumes "p>0"
and f0: "Df 0 = f"
and Df: "\<And>m t. m < p \<Longrightarrow> a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow>
(Df m has_vector_derivative Df (Suc m) t) (at t within {a..b})"
and g0: "Dg 0 = g"
and Dg: "\<And>m t. m < p \<Longrightarrow> a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow>
(Dg m has_vector_derivative Dg (Suc m) t) (at t within {a..b})"
and ivl: "a \<le> t" "t \<le> b"
shows "((\<lambda>t. \<Sum>i<p. (-1)^i *\<^sub>R prod (Df i t) (Dg (p - Suc i) t))
has_vector_derivative
prod (f t) (Dg p t) - (-1)^p *\<^sub>R prod (Df p t) (g t))
(at t within {a..b})"
using assms
proof cases
assume p: "p \<noteq> 1"
define p' where "p' = p - 2"
from assms p have p': "{..<p} = {..Suc p'}" "p = Suc (Suc p')"
by (auto simp: p'_def)
have *: "\<And>i. i \<le> p' \<Longrightarrow> Suc (Suc p' - i) = (Suc (Suc p') - i)"
by auto
let ?f = "\<lambda>i. (-1) ^ i *\<^sub>R (prod (Df i t) (Dg ((p - i)) t))"
have "(\<Sum>i<p. (-1) ^ i *\<^sub>R (prod (Df i t) (Dg (Suc (p - Suc i)) t) +
prod (Df (Suc i) t) (Dg (p - Suc i) t))) =
(\<Sum>i\<le>(Suc p'). ?f i - ?f (Suc i))"
by (auto simp: algebra_simps p'(2) numeral_2_eq_2 * lessThan_Suc_atMost)
also note sum_telescope
finally
have "(\<Sum>i<p. (-1) ^ i *\<^sub>R (prod (Df i t) (Dg (Suc (p - Suc i)) t) +
prod (Df (Suc i) t) (Dg (p - Suc i) t)))
= prod (f t) (Dg p t) - (- 1) ^ p *\<^sub>R prod (Df p t) (g t)"
unfolding p'[symmetric]
by (simp add: assms)
thus ?thesis
using assms
by (auto intro!: derivative_eq_intros has_vector_derivative)
qed (auto intro!: derivative_eq_intros has_vector_derivative)
lemma
fixes f::"real\<Rightarrow>'a::banach"
assumes "p>0"
and f0: "Df 0 = f"
and Df: "\<And>m t. m < p \<Longrightarrow> a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow>
(Df m has_vector_derivative Df (Suc m) t) (at t within {a..b})"
and ivl: "a \<le> b"
defines "i \<equiv> \<lambda>x. ((b - x) ^ (p - 1) / fact (p - 1)) *\<^sub>R Df p x"
shows Taylor_has_integral:
"(i has_integral f b - (\<Sum>i<p. ((b-a) ^ i / fact i) *\<^sub>R Df i a)) {a..b}"
and Taylor_integral:
"f b = (\<Sum>i<p. ((b-a) ^ i / fact i) *\<^sub>R Df i a) + integral {a..b} i"
and Taylor_integrable:
"i integrable_on {a..b}"
proof goal_cases
case 1
interpret bounded_bilinear "scaleR::real\<Rightarrow>'a\<Rightarrow>'a"
by (rule bounded_bilinear_scaleR)
define g where "g s = (b - s)^(p - 1)/fact (p - 1)" for s
define Dg where [abs_def]:
"Dg n s = (if n < p then (-1)^n * (b - s)^(p - 1 - n) / fact (p - 1 - n) else 0)" for n s
have g0: "Dg 0 = g"
using \<open>p > 0\<close>
by (auto simp add: Dg_def field_split_simps g_def split: if_split_asm)
{
fix m
assume "p > Suc m"
hence "p - Suc m = Suc (p - Suc (Suc m))"
by auto
hence "real (p - Suc m) * fact (p - Suc (Suc m)) = fact (p - Suc m)"
by auto
} note fact_eq = this
have Dg: "\<And>m t. m < p \<Longrightarrow> a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow>
(Dg m has_vector_derivative Dg (Suc m) t) (at t within {a..b})"
unfolding Dg_def
by (auto intro!: derivative_eq_intros simp: has_vector_derivative_def fact_eq field_split_simps)
let ?sum = "\<lambda>t. \<Sum>i<p. (- 1) ^ i *\<^sub>R Dg i t *\<^sub>R Df (p - Suc i) t"
from sum_prod_derivatives_has_vector_derivative[of _ Dg _ _ _ Df,
OF \<open>p > 0\<close> g0 Dg f0 Df]
have deriv: "\<And>t. a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow>
(?sum has_vector_derivative
g t *\<^sub>R Df p t - (- 1) ^ p *\<^sub>R Dg p t *\<^sub>R f t) (at t within {a..b})"
by auto
from fundamental_theorem_of_calculus[rule_format, OF \<open>a \<le> b\<close> deriv]
have "(i has_integral ?sum b - ?sum a) {a..b}"
using atLeastatMost_empty'[simp del]
by (simp add: i_def g_def Dg_def)
also
have one: "(- 1) ^ p' * (- 1) ^ p' = (1::real)"
and "{..<p} \<inter> {i. p = Suc i} = {p - 1}"
for p'
using \<open>p > 0\<close>
by (auto simp: power_mult_distrib[symmetric])
then have "?sum b = f b"
using Suc_pred'[OF \<open>p > 0\<close>]
by (simp add: diff_eq_eq Dg_def power_0_left le_Suc_eq if_distrib
if_distribR sum.If_cases f0)
also
have "{..<p} = (\<lambda>x. p - x - 1) ` {..<p}"
proof safe
fix x
assume "x < p"
thus "x \<in> (\<lambda>x. p - x - 1) ` {..<p}"
by (auto intro!: image_eqI[where x = "p - x - 1"])
qed simp
from _ this
have "?sum a = (\<Sum>i<p. ((b-a) ^ i / fact i) *\<^sub>R Df i a)"
by (rule sum.reindex_cong) (auto simp add: inj_on_def Dg_def one)
finally show c: ?case .
case 2 show ?case using c integral_unique
by (metis (lifting) add.commute diff_eq_eq integral_unique)
case 3 show ?case using c by force
qed
subsection \<open>Only need trivial subintervals if the interval itself is trivial\<close>
proposition division_of_nontrivial:
fixes \<D> :: "'a::euclidean_space set set"
assumes sdiv: "\<D> division_of (cbox a b)"
and cont0: "content (cbox a b) \<noteq> 0"
shows "{k. k \<in> \<D> \<and> content k \<noteq> 0} division_of (cbox a b)"
using sdiv
proof (induction "card \<D>" arbitrary: \<D> rule: less_induct)
case less
note \<D> = division_ofD[OF less.prems]
{
presume *: "{k \<in> \<D>. content k \<noteq> 0} \<noteq> \<D> \<Longrightarrow> ?case"
then show ?case
using less.prems by fastforce
}
assume noteq: "{k \<in> \<D>. content k \<noteq> 0} \<noteq> \<D>"
then obtain K c d where "K \<in> \<D>" and contk: "content K = 0" and keq: "K = cbox c d"
using \<D>(4) by blast
then have "card \<D> > 0"
unfolding card_gt_0_iff using less by auto
then have card: "card (\<D> - {K}) < card \<D>"
using less \<open>K \<in> \<D>\<close> by (simp add: \<D>(1))
have closed: "closed (\<Union>(\<D> - {K}))"
using less.prems by auto
have "x islimpt \<Union>(\<D> - {K})" if "x \<in> K" for x
unfolding islimpt_approachable
proof (intro allI impI)
fix e::real
assume "e > 0"
obtain i where i: "c\<bullet>i = d\<bullet>i" "i\<in>Basis"
using contk \<D>(3) [OF \<open>K \<in> \<D>\<close>] unfolding box_ne_empty keq
by (meson content_eq_0 dual_order.antisym)
then have xi: "x\<bullet>i = d\<bullet>i"
using \<open>x \<in> K\<close> unfolding keq mem_box by (metis antisym)
define y where "y = (\<Sum>j\<in>Basis. (if j = i then if c\<bullet>i \<le> (a\<bullet>i + b\<bullet>i)/2 then c\<bullet>i +
min e (b\<bullet>i - c\<bullet>i)/2 else c\<bullet>i - min e (c\<bullet>i - a\<bullet>i)/2 else x\<bullet>j) *\<^sub>R j)"
show "\<exists>x'\<in>\<Union>(\<D> - {K}). x' \<noteq> x \<and> dist x' x < e"
proof (intro bexI conjI)
have "d \<in> cbox c d"
using \<D>(3)[OF \<open>K \<in> \<D>\<close>] by (simp add: box_ne_empty(1) keq mem_box(2))
then have "d \<in> cbox a b"
using \<D>(2)[OF \<open>K \<in> \<D>\<close>] by (auto simp: keq)
then have di: "a \<bullet> i \<le> d \<bullet> i \<and> d \<bullet> i \<le> b \<bullet> i"
using \<open>i \<in> Basis\<close> mem_box(2) by blast
then have xyi: "y\<bullet>i \<noteq> x\<bullet>i"
unfolding y_def i xi using \<open>e > 0\<close> cont0 \<open>i \<in> Basis\<close>
by (auto simp: content_eq_0 elim!: ballE[of _ _ i])
then show "y \<noteq> x"
unfolding euclidean_eq_iff[where 'a='a] using i by auto
have "norm (y-x) \<le> (\<Sum>b\<in>Basis. \<bar>(y - x) \<bullet> b\<bar>)"
by (rule norm_le_l1)
also have "... = \<bar>(y - x) \<bullet> i\<bar> + (\<Sum>b \<in> Basis - {i}. \<bar>(y - x) \<bullet> b\<bar>)"
by (meson finite_Basis i(2) sum.remove)
also have "... < e + sum (\<lambda>i. 0) Basis"
proof (rule add_less_le_mono)
show "\<bar>(y-x) \<bullet> i\<bar> < e"
using di \<open>e > 0\<close> y_def i xi by (auto simp: inner_simps)
show "(\<Sum>i\<in>Basis - {i}. \<bar>(y-x) \<bullet> i\<bar>) \<le> (\<Sum>i\<in>Basis. 0)"
unfolding y_def by (auto simp: inner_simps)
qed
finally have "norm (y-x) < e + sum (\<lambda>i. 0) Basis" .
then show "dist y x < e"
unfolding dist_norm by auto
have "y \<notin> K"
unfolding keq mem_box using i(1) i(2) xi xyi by fastforce
moreover have "y \<in> \<Union>\<D>"
using subsetD[OF \<D>(2)[OF \<open>K \<in> \<D>\<close>] \<open>x \<in> K\<close>] \<open>e > 0\<close> di i
by (auto simp: \<D> mem_box y_def field_simps elim!: ballE[of _ _ i])
ultimately show "y \<in> \<Union>(\<D> - {K})" by auto
qed
qed
then have "K \<subseteq> \<Union>(\<D> - {K})"
using closed closed_limpt by blast
then have "\<Union>(\<D> - {K}) = cbox a b"
unfolding \<D>(6)[symmetric] by auto
then have "\<D> - {K} division_of cbox a b"
by (metis Diff_subset less.prems division_of_subset \<D>(6))
then have "{ka \<in> \<D> - {K}. content ka \<noteq> 0} division_of (cbox a b)"
using card less.hyps by blast
moreover have "{ka \<in> \<D> - {K}. content ka \<noteq> 0} = {K \<in> \<D>. content K \<noteq> 0}"
using contk by auto
ultimately show ?case by auto
qed
subsection \<open>Integrability on subintervals\<close>
lemma operative_integrableI:
fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach"
assumes "0 \<le> e"
shows "operative conj True (\<lambda>i. f integrable_on i)"
proof -
interpret comm_monoid conj True
by standard auto
have 1: "\<And>a b. box a b = {} \<Longrightarrow> f integrable_on cbox a b"
by (simp add: content_eq_0_interior integrable_on_null)
have 2: "\<And>a b c k.
\<lbrakk>k \<in> Basis;
f integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c};
f integrable_on cbox a b \<inter> {x. c \<le> x \<bullet> k}\<rbrakk>
\<Longrightarrow> f integrable_on cbox a b"
unfolding integrable_on_def by (auto intro!: has_integral_split)
show ?thesis
apply standard
using 1 2 by blast+
qed
lemma integrable_subinterval:
fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach"
assumes f: "f integrable_on cbox a b"
and cd: "cbox c d \<subseteq> cbox a b"
shows "f integrable_on cbox c d"
proof -
interpret operative conj True "\<lambda>i. f integrable_on i"
using order_refl by (rule operative_integrableI)
show ?thesis
proof (cases "cbox c d = {}")
case True
then show ?thesis
using division [symmetric] f by (auto simp: comm_monoid_set_F_and)
next
case False
then show ?thesis
by (metis cd comm_monoid_set_F_and division division_of_finite f partial_division_extend_1)
qed
qed
lemma integrable_subinterval_real:
fixes f :: "real \<Rightarrow> 'a::banach"
assumes "f integrable_on {a..b}"
and "{c..d} \<subseteq> {a..b}"
shows "f integrable_on {c..d}"
by (metis assms box_real(2) integrable_subinterval)
subsection \<open>Combining adjacent intervals in 1 dimension\<close>
lemma has_integral_combine:
fixes a b c :: real and j :: "'a::banach"
assumes "a \<le> c"
and "c \<le> b"
and ac: "(f has_integral i) {a..c}"
and cb: "(f has_integral j) {c..b}"
shows "(f has_integral (i + j)) {a..b}"
proof -
interpret operative_real "lift_option plus" "Some 0"
"\<lambda>i. if f integrable_on i then Some (integral i f) else None"
using operative_integralI by (rule operative_realI)
from \<open>a \<le> c\<close> \<open>c \<le> b\<close> ac cb coalesce_less_eq
have *: "lift_option (+)
(if f integrable_on {a..c} then Some (integral {a..c} f) else None)
(if f integrable_on {c..b} then Some (integral {c..b} f) else None) =
(if f integrable_on {a..b} then Some (integral {a..b} f) else None)"
by (auto simp: split: if_split_asm)
then have "f integrable_on cbox a b"
using ac cb by (auto split: if_split_asm)
with *
show ?thesis
using ac cb by (auto simp add: integrable_on_def integral_unique split: if_split_asm)
qed
lemma integral_combine:
fixes f :: "real \<Rightarrow> 'a::banach"
assumes "a \<le> c"
and "c \<le> b"
and ab: "f integrable_on {a..b}"
shows "integral {a..c} f + integral {c..b} f = integral {a..b} f"
proof -
have "(f has_integral integral {a..c} f) {a..c}"
using ab \<open>c \<le> b\<close> integrable_subinterval_real by fastforce
moreover
have "(f has_integral integral {c..b} f) {c..b}"
using ab \<open>a \<le> c\<close> integrable_subinterval_real by fastforce
ultimately have "(f has_integral integral {a..c} f + integral {c..b} f) {a..b}"
using \<open>a \<le> c\<close> \<open>c \<le> b\<close> has_integral_combine by blast
then show ?thesis
by (simp add: has_integral_integrable_integral)
qed
lemma integrable_combine:
fixes f :: "real \<Rightarrow> 'a::banach"
assumes "a \<le> c"
and "c \<le> b"
and "f integrable_on {a..c}"
and "f integrable_on {c..b}"
shows "f integrable_on {a..b}"
using assms
unfolding integrable_on_def
by (auto intro!: has_integral_combine)
lemma integral_minus_sets:
fixes f::"real \<Rightarrow> 'a::banach"
shows "c \<le> a \<Longrightarrow> c \<le> b \<Longrightarrow> f integrable_on {c .. max a b} \<Longrightarrow>
integral {c .. a} f - integral {c .. b} f =
(if a \<le> b then - integral {a .. b} f else integral {b .. a} f)"
using integral_combine[of c a b f] integral_combine[of c b a f]
by (auto simp: algebra_simps max_def)
lemma integral_minus_sets':
fixes f::"real \<Rightarrow> 'a::banach"
shows "c \<ge> a \<Longrightarrow> c \<ge> b \<Longrightarrow> f integrable_on {min a b .. c} \<Longrightarrow>
integral {a .. c} f - integral {b .. c} f =
(if a \<le> b then integral {a .. b} f else - integral {b .. a} f)"
using integral_combine[of b a c f] integral_combine[of a b c f]
by (auto simp: algebra_simps min_def)
subsection \<open>Reduce integrability to "local" integrability\<close>
lemma integrable_on_little_subintervals:
fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach"
assumes "\<forall>x\<in>cbox a b. \<exists>d>0. \<forall>u v. x \<in> cbox u v \<and> cbox u v \<subseteq> ball x d \<and> cbox u v \<subseteq> cbox a b \<longrightarrow>
f integrable_on cbox u v"
shows "f integrable_on cbox a b"
proof -
interpret operative conj True "\<lambda>i. f integrable_on i"
using order_refl by (rule operative_integrableI)
have "\<forall>x. \<exists>d>0. x\<in>cbox a b \<longrightarrow> (\<forall>u v. x \<in> cbox u v \<and> cbox u v \<subseteq> ball x d \<and> cbox u v \<subseteq> cbox a b \<longrightarrow>
f integrable_on cbox u v)"
using assms by (metis zero_less_one)
then obtain d where d: "\<And>x. 0 < d x"
"\<And>x u v. \<lbrakk>x \<in> cbox a b; x \<in> cbox u v; cbox u v \<subseteq> ball x (d x); cbox u v \<subseteq> cbox a b\<rbrakk>
\<Longrightarrow> f integrable_on cbox u v"
by metis
obtain p where p: "p tagged_division_of cbox a b" "(\<lambda>x. ball x (d x)) fine p"
using fine_division_exists[OF gauge_ball_dependent,of d a b] d(1) by blast
then have sndp: "snd ` p division_of cbox a b"
by (metis division_of_tagged_division)
have "f integrable_on k" if "(x, k) \<in> p" for x k
using tagged_division_ofD(2-4)[OF p(1) that] fineD[OF p(2) that] d[of x] by auto
then show ?thesis
unfolding division [symmetric, OF sndp] comm_monoid_set_F_and
by auto
qed
subsection \<open>Second FTC or existence of antiderivative\<close>
lemma integrable_const[intro]: "(\<lambda>x. c) integrable_on cbox a b"
unfolding integrable_on_def by blast
lemma integral_has_vector_derivative_continuous_at:
fixes f :: "real \<Rightarrow> 'a::banach"
assumes f: "f integrable_on {a..b}"
and x: "x \<in> {a..b} - S"
and "finite S"
and fx: "continuous (at x within ({a..b} - S)) f"
shows "((\<lambda>u. integral {a..u} f) has_vector_derivative f x) (at x within ({a..b} - S))"
proof -
let ?I = "\<lambda>a b. integral {a..b} f"
{ fix e::real
assume "e > 0"
obtain d where "d>0" and d: "\<And>x'. \<lbrakk>x' \<in> {a..b} - S; \<bar>x' - x\<bar> < d\<rbrakk> \<Longrightarrow> norm(f x' - f x) \<le> e"
using \<open>e>0\<close> fx by (auto simp: continuous_within_eps_delta dist_norm less_imp_le)
have "norm (integral {a..y} f - integral {a..x} f - (y-x) *\<^sub>R f x) \<le> e * \<bar>y - x\<bar>"
if y: "y \<in> {a..b} - S" and yx: "\<bar>y - x\<bar> < d" for y
proof (cases "y < x")
case False
have "f integrable_on {a..y}"
using f y by (simp add: integrable_subinterval_real)
then have Idiff: "?I a y - ?I a x = ?I x y"
using False x by (simp add: algebra_simps integral_combine)
have fux_int: "((\<lambda>u. f u - f x) has_integral integral {x..y} f - (y-x) *\<^sub>R f x) {x..y}"
apply (rule has_integral_diff)
using x y apply (auto intro: integrable_integral [OF integrable_subinterval_real [OF f]])
using has_integral_const_real [of "f x" x y] False
apply simp
done
have "\<And>xa. y - x < d \<Longrightarrow> (\<And>x'. a \<le> x' \<and> x' \<le> b \<and> x' \<notin> S \<Longrightarrow> \<bar>x' - x\<bar> < d \<Longrightarrow> norm (f x' - f x) \<le> e) \<Longrightarrow> 0 < e \<Longrightarrow> xa \<notin> S \<Longrightarrow> a \<le> x \<Longrightarrow> x \<notin> S \<Longrightarrow> y \<le> b \<Longrightarrow> y \<notin> S \<Longrightarrow> x \<le> xa \<Longrightarrow> xa \<le> y \<Longrightarrow> norm (f xa - f x) \<le> e"
using assms by auto
show ?thesis
using False
apply (simp add: abs_eq_content del: content_real_if measure_lborel_Icc)
apply (rule has_integral_bound_real[where f="(\<lambda>u. f u - f x)"])
using yx False d x y \<open>e>0\<close> assms by (auto simp: Idiff fux_int)
next
case True
have "f integrable_on {a..x}"
using f x by (simp add: integrable_subinterval_real)
then have Idiff: "?I a x - ?I a y = ?I y x"
using True x y by (simp add: algebra_simps integral_combine)
have fux_int: "((\<lambda>u. f u - f x) has_integral integral {y..x} f - (x - y) *\<^sub>R f x) {y..x}"
apply (rule has_integral_diff)
using x y apply (auto intro: integrable_integral [OF integrable_subinterval_real [OF f]])
using has_integral_const_real [of "f x" y x] True
apply simp
done
have "norm (integral {a..x} f - integral {a..y} f - (x - y) *\<^sub>R f x) \<le> e * \<bar>y - x\<bar>"
using True
apply (simp add: abs_eq_content del: content_real_if measure_lborel_Icc)
apply (rule has_integral_bound_real[where f="(\<lambda>u. f u - f x)"])
using yx True d x y \<open>e>0\<close> assms by (auto simp: Idiff fux_int)
then show ?thesis
by (simp add: algebra_simps norm_minus_commute)
qed
then have "\<exists>d>0. \<forall>y\<in>{a..b} - S. \<bar>y - x\<bar> < d \<longrightarrow> norm (integral {a..y} f - integral {a..x} f - (y-x) *\<^sub>R f x) \<le> e * \<bar>y - x\<bar>"
using \<open>d>0\<close> by blast
}
then show ?thesis
by (simp add: has_vector_derivative_def has_derivative_within_alt bounded_linear_scaleR_left)
qed
lemma integral_has_vector_derivative:
fixes f :: "real \<Rightarrow> 'a::banach"
assumes "continuous_on {a..b} f"
and "x \<in> {a..b}"
shows "((\<lambda>u. integral {a..u} f) has_vector_derivative f(x)) (at x within {a..b})"
using assms integral_has_vector_derivative_continuous_at [OF integrable_continuous_real]
by (fastforce simp: continuous_on_eq_continuous_within)
lemma integral_has_real_derivative:
assumes "continuous_on {a..b} g"
assumes "t \<in> {a..b}"
shows "((\<lambda>x. integral {a..x} g) has_real_derivative g t) (at t within {a..b})"
using integral_has_vector_derivative[of a b g t] assms
by (auto simp: has_field_derivative_iff_has_vector_derivative)
lemma antiderivative_continuous:
fixes q b :: real
assumes "continuous_on {a..b} f"
obtains g where "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_vector_derivative (f x::_::banach)) (at x within {a..b})"
using integral_has_vector_derivative[OF assms] by auto
subsection \<open>Combined fundamental theorem of calculus\<close>
lemma antiderivative_integral_continuous:
fixes f :: "real \<Rightarrow> 'a::banach"
assumes "continuous_on {a..b} f"
obtains g where "\<forall>u\<in>{a..b}. \<forall>v \<in> {a..b}. u \<le> v \<longrightarrow> (f has_integral (g v - g u)) {u..v}"
proof -
obtain g
where g: "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_vector_derivative f x) (at x within {a..b})"
using antiderivative_continuous[OF assms] by metis
have "(f has_integral g v - g u) {u..v}" if "u \<in> {a..b}" "v \<in> {a..b}" "u \<le> v" for u v
proof -
have "\<And>x. x \<in> cbox u v \<Longrightarrow> (g has_vector_derivative f x) (at x within cbox u v)"
by (metis atLeastAtMost_iff atLeastatMost_subset_iff box_real(2) g
has_vector_derivative_within_subset subsetCE that(1,2))
then show ?thesis
by (metis box_real(2) that(3) fundamental_theorem_of_calculus)
qed
then show ?thesis
using that by blast
qed
subsection \<open>General "twiddling" for interval-to-interval function image\<close>
lemma has_integral_twiddle:
assumes "0 < r"
and hg: "\<And>x. h(g x) = x"
and gh: "\<And>x. g(h x) = x"
and contg: "\<And>x. continuous (at x) g"
and g: "\<And>u v. \<exists>w z. g ` cbox u v = cbox w z"
and h: "\<And>u v. \<exists>w z. h ` cbox u v = cbox w z"
and r: "\<And>u v. content(g ` cbox u v) = r * content (cbox u v)"
and intfi: "(f has_integral i) (cbox a b)"
shows "((\<lambda>x. f(g x)) has_integral (1 / r) *\<^sub>R i) (h ` cbox a b)"
proof (cases "cbox a b = {}")
case True
then show ?thesis
using intfi by auto
next
case False
obtain w z where wz: "h ` cbox a b = cbox w z"
using h by blast
have inj: "inj g" "inj h"
using hg gh injI by metis+
from h obtain ha hb where h_eq: "h ` cbox a b = cbox ha hb" by blast
have "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of h ` cbox a b \<and> d fine p
\<longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e)"
if "e > 0" for e
proof -
have "e * r > 0" using that \<open>0 < r\<close> by simp
with intfi[unfolded has_integral]
obtain d where "gauge d"
and d: "\<And>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) - i) < e * r"
by metis
define d' where "d' x = g -` d (g x)" for x
show ?thesis
proof (rule_tac x=d' in exI, safe)
show "gauge d'"
using \<open>gauge d\<close> continuous_open_vimage[OF _ contg] by (auto simp: gauge_def d'_def)
next
fix p
assume ptag: "p tagged_division_of h ` cbox a b" and finep: "d' fine p"
note p = tagged_division_ofD[OF ptag]
have gab: "g y \<in> cbox a b" if "y \<in> K" "(x, K) \<in> p" for x y K
by (metis hg inj(2) inj_image_mem_iff p(3) subsetCE that that)
have gimp: "(\<lambda>(x,K). (g x, g ` K)) ` p tagged_division_of (cbox a b) \<and>
d fine (\<lambda>(x, k). (g x, g ` k)) ` p"
unfolding tagged_division_of
proof safe
show "finite ((\<lambda>(x, k). (g x, g ` k)) ` p)"
using ptag by auto
show "d fine (\<lambda>(x, k). (g x, g ` k)) ` p"
using finep unfolding fine_def d'_def by auto
next
fix x k
assume xk: "(x, k) \<in> p"
show "g x \<in> g ` k"
using p(2)[OF xk] by auto
show "\<exists>u v. g ` k = cbox u v"
using p(4)[OF xk] using assms(5-6) by auto
fix x' K' u
assume xk': "(x', K') \<in> p" and u: "u \<in> interior (g ` k)" "u \<in> interior (g ` K')"
have "interior k \<inter> interior K' \<noteq> {}"
proof
assume "interior k \<inter> interior K' = {}"
moreover have "u \<in> g ` (interior k \<inter> interior K')"
using interior_image_subset[OF \<open>inj g\<close> contg] u
unfolding image_Int[OF inj(1)] by blast
ultimately show False by blast
qed
then have same: "(x, k) = (x', K')"
using ptag xk' xk by blast
then show "g x = g x'"
by auto
show "g u \<in> g ` K'"if "u \<in> k" for u
using that same by auto
show "g u \<in> g ` k"if "u \<in> K'" for u
using that same by auto
next
fix x
assume "x \<in> cbox a b"
then have "h x \<in> \<Union>{k. \<exists>x. (x, k) \<in> p}"
using p(6) by auto
then obtain X y where "h x \<in> X" "(y, X) \<in> p" by blast
then show "x \<in> \<Union>{k. \<exists>x. (x, k) \<in> (\<lambda>(x, k). (g x, g ` k)) ` p}"
apply clarsimp
by (metis (no_types, lifting) assms(3) image_eqI pair_imageI)
qed (use gab in auto)
have *: "inj_on (\<lambda>(x, k). (g x, g ` k)) p"
using inj(1) unfolding inj_on_def by fastforce
have "(\<Sum>(x, k)\<in>(\<lambda>(x, k). (g x, g ` k)) ` p. content k *\<^sub>R f x) - i = r *\<^sub>R (\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - i" (is "?l = _")
using r
apply (simp only: algebra_simps add_left_cancel scaleR_right.sum)
apply (subst sum.reindex_bij_betw[symmetric, where h="\<lambda>(x, k). (g x, g ` k)" and S=p])
apply (auto intro!: * sum.cong simp: bij_betw_def dest!: p(4))
done
also have "\<dots> = r *\<^sub>R ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i)" (is "_ = ?r")
using \<open>0 < r\<close> by (auto simp: scaleR_diff_right)
finally have eq: "?l = ?r" .
show "norm ((\<Sum>(x,K)\<in>p. content K *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e"
using d[OF gimp] \<open>0 < r\<close> by (auto simp add: eq)
qed
qed
then show ?thesis
by (auto simp: h_eq has_integral)
qed
subsection \<open>Special case of a basic affine transformation\<close>
lemma AE_lborel_inner_neq:
assumes k: "k \<in> Basis"
shows "AE x in lborel. x \<bullet> k \<noteq> c"
proof -
interpret finite_product_sigma_finite "\<lambda>_. lborel" Basis
proof qed simp
have "emeasure lborel {x\<in>space lborel. x \<bullet> k = c} = emeasure (\<Pi>\<^sub>M j::'a\<in>Basis. lborel) (\<Pi>\<^sub>E j\<in>Basis. if j = k then {c} else UNIV)"
using k
by (auto simp add: lborel_eq[where 'a='a] emeasure_distr intro!: arg_cong2[where f=emeasure])
(auto simp: space_PiM PiE_iff extensional_def split: if_split_asm)
also have "\<dots> = (\<Prod>j\<in>Basis. emeasure lborel (if j = k then {c} else UNIV))"
by (intro measure_times) auto
also have "\<dots> = 0"
by (intro prod_zero bexI[OF _ k]) auto
finally show ?thesis
by (subst AE_iff_measurable[OF _ refl]) auto
qed
lemma content_image_stretch_interval:
fixes m :: "'a::euclidean_space \<Rightarrow> real"
defines "s f x \<equiv> (\<Sum>k::'a\<in>Basis. (f k * (x\<bullet>k)) *\<^sub>R k)"
shows "content (s m ` cbox a b) = \<bar>\<Prod>k\<in>Basis. m k\<bar> * content (cbox a b)"
proof cases
have s[measurable]: "s f \<in> borel \<rightarrow>\<^sub>M borel" for f
by (auto simp: s_def[abs_def])
assume m: "\<forall>k\<in>Basis. m k \<noteq> 0"
then have s_comp_s: "s (\<lambda>k. 1 / m k) \<circ> s m = id" "s m \<circ> s (\<lambda>k. 1 / m k) = id"
by (auto simp: s_def[abs_def] fun_eq_iff euclidean_representation)
then have "inv (s (\<lambda>k. 1 / m k)) = s m" "bij (s (\<lambda>k. 1 / m k))"
by (auto intro: inv_unique_comp o_bij)
then have eq: "s m ` cbox a b = s (\<lambda>k. 1 / m k) -` cbox a b"
using bij_vimage_eq_inv_image[OF \<open>bij (s (\<lambda>k. 1 / m k))\<close>, of "cbox a b"] by auto
show ?thesis
using m unfolding eq measure_def
by (subst lborel_affine_euclidean[where c=m and t=0])
(simp_all add: emeasure_density measurable_sets_borel[OF s] abs_prod nn_integral_cmult
s_def[symmetric] emeasure_distr vimage_comp s_comp_s enn2real_mult prod_nonneg)
next
assume "\<not> (\<forall>k\<in>Basis. m k \<noteq> 0)"
then obtain k where k: "k \<in> Basis" "m k = 0" by auto
then have [simp]: "(\<Prod>k\<in>Basis. m k) = 0"
by (intro prod_zero) auto
have "emeasure lborel {x\<in>space lborel. x \<in> s m ` cbox a b} = 0"
proof (rule emeasure_eq_0_AE)
show "AE x in lborel. x \<notin> s m ` cbox a b"
using AE_lborel_inner_neq[OF \<open>k\<in>Basis\<close>]
proof eventually_elim
show "x \<bullet> k \<noteq> 0 \<Longrightarrow> x \<notin> s m ` cbox a b " for x
using k by (auto simp: s_def[abs_def] cbox_def)
qed
qed
then show ?thesis
by (simp add: measure_def)
qed
lemma interval_image_affinity_interval:
"\<exists>u v. (\<lambda>x. m *\<^sub>R (x::'a::euclidean_space) + c) ` cbox a b = cbox u v"
unfolding image_affinity_cbox
by auto
lemma content_image_affinity_cbox:
"content((\<lambda>x::'a::euclidean_space. m *\<^sub>R x + c) ` cbox a b) =
\<bar>m\<bar> ^ DIM('a) * content (cbox a b)" (is "?l = ?r")
proof (cases "cbox a b = {}")
case True then show ?thesis by simp
next
case False
show ?thesis
proof (cases "m \<ge> 0")
case True
with \<open>cbox a b \<noteq> {}\<close> have "cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c) \<noteq> {}"
unfolding box_ne_empty
apply (intro ballI)
apply (erule_tac x=i in ballE)
apply (auto simp: inner_simps mult_left_mono)
done
moreover from True have *: "\<And>i. (m *\<^sub>R b + c) \<bullet> i - (m *\<^sub>R a + c) \<bullet> i = m *\<^sub>R (b-a) \<bullet> i"
by (simp add: inner_simps field_simps)
ultimately show ?thesis
by (simp add: image_affinity_cbox True content_cbox'
prod.distrib inner_diff_left)
next
case False
with \<open>cbox a b \<noteq> {}\<close> have "cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c) \<noteq> {}"
unfolding box_ne_empty
apply (intro ballI)
apply (erule_tac x=i in ballE)
apply (auto simp: inner_simps mult_left_mono)
done
moreover from False have *: "\<And>i. (m *\<^sub>R a + c) \<bullet> i - (m *\<^sub>R b + c) \<bullet> i = (-m) *\<^sub>R (b-a) \<bullet> i"
by (simp add: inner_simps field_simps)
ultimately show ?thesis using False
by (simp add: image_affinity_cbox content_cbox'
prod.distrib[symmetric] inner_diff_left flip: prod_constant)
qed
qed
lemma has_integral_affinity:
fixes a :: "'a::euclidean_space"
assumes "(f has_integral i) (cbox a b)"
and "m \<noteq> 0"
shows "((\<lambda>x. f(m *\<^sub>R x + c)) has_integral ((1 / (\<bar>m\<bar> ^ DIM('a))) *\<^sub>R i)) ((\<lambda>x. (1 / m) *\<^sub>R x + -((1 / m) *\<^sub>R c)) ` cbox a b)"
apply (rule has_integral_twiddle)
using assms
apply (safe intro!: interval_image_affinity_interval content_image_affinity_cbox)
apply (rule zero_less_power)
unfolding scaleR_right_distrib
apply auto
done
lemma integrable_affinity:
assumes "f integrable_on cbox a b"
and "m \<noteq> 0"
shows "(\<lambda>x. f(m *\<^sub>R x + c)) integrable_on ((\<lambda>x. (1 / m) *\<^sub>R x + -((1/m) *\<^sub>R c)) ` cbox a b)"
using assms
unfolding integrable_on_def
apply safe
apply (drule has_integral_affinity)
apply auto
done
lemmas has_integral_affinity01 = has_integral_affinity [of _ _ 0 "1::real", simplified]
lemma integrable_on_affinity:
assumes "m \<noteq> 0" "f integrable_on (cbox a b)"
shows "(\<lambda>x. f (m *\<^sub>R x + c)) integrable_on ((\<lambda>x. (1 / m) *\<^sub>R x - ((1 / m) *\<^sub>R c)) ` cbox a b)"
proof -
from assms obtain I where "(f has_integral I) (cbox a b)"
by (auto simp: integrable_on_def)
from has_integral_affinity[OF this assms(1), of c] show ?thesis
by (auto simp: integrable_on_def)
qed
lemma has_integral_cmul_iff:
assumes "c \<noteq> 0"
shows "((\<lambda>x. c *\<^sub>R f x) has_integral (c *\<^sub>R I)) A \<longleftrightarrow> (f has_integral I) A"
using assms has_integral_cmul[of f I A c]
has_integral_cmul[of "\<lambda>x. c *\<^sub>R f x" "c *\<^sub>R I" A "inverse c"] by (auto simp: field_simps)
lemma has_integral_affinity':
fixes a :: "'a::euclidean_space"
assumes "(f has_integral i) (cbox a b)" and "m > 0"
shows "((\<lambda>x. f(m *\<^sub>R x + c)) has_integral (i /\<^sub>R m ^ DIM('a)))
(cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m))"
proof (cases "cbox a b = {}")
case True
hence "(cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)) = {}"
using \<open>m > 0\<close> unfolding box_eq_empty by (auto simp: algebra_simps)
with True and assms show ?thesis by simp
next
case False
have "((\<lambda>x. f (m *\<^sub>R x + c)) has_integral (1 / \<bar>m\<bar> ^ DIM('a)) *\<^sub>R i)
((\<lambda>x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox a b)"
using assms by (intro has_integral_affinity) auto
also have "((\<lambda>x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox a b) =
((\<lambda>x. - ((1 / m) *\<^sub>R c) + x) ` (\<lambda>x. (1 / m) *\<^sub>R x) ` cbox a b)"
by (simp add: image_image algebra_simps)
also have "(\<lambda>x. (1 / m) *\<^sub>R x) ` cbox a b = cbox ((1 / m) *\<^sub>R a) ((1 / m) *\<^sub>R b)" using \<open>m > 0\<close> False
by (subst image_smult_cbox) simp_all
also have "(\<lambda>x. - ((1 / m) *\<^sub>R c) + x) ` \<dots> = cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)"
by (subst cbox_translation [symmetric]) (simp add: field_simps vector_add_divide_simps)
finally show ?thesis using \<open>m > 0\<close> by (simp add: field_simps)
qed
lemma has_integral_affinity_iff:
fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: real_normed_vector"
assumes "m > 0"
shows "((\<lambda>x. f (m *\<^sub>R x + c)) has_integral (I /\<^sub>R m ^ DIM('a)))
(cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)) \<longleftrightarrow>
(f has_integral I) (cbox a b)" (is "?lhs = ?rhs")
proof
assume ?lhs
from has_integral_affinity'[OF this, of "1 / m" "-c /\<^sub>R m"] and \<open>m > 0\<close>
show ?rhs by (simp add: vector_add_divide_simps) (simp add: field_simps)
next
assume ?rhs
from has_integral_affinity'[OF this, of m c] and \<open>m > 0\<close>
show ?lhs by simp
qed
subsection \<open>Special case of stretching coordinate axes separately\<close>
lemma has_integral_stretch:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
assumes "(f has_integral i) (cbox a b)"
and "\<forall>k\<in>Basis. m k \<noteq> 0"
shows "((\<lambda>x. f (\<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k)) has_integral
((1/ \<bar>prod m Basis\<bar>) *\<^sub>R i)) ((\<lambda>x. (\<Sum>k\<in>Basis. (1 / m k * (x\<bullet>k))*\<^sub>R k)) ` cbox a b)"
apply (rule has_integral_twiddle[where f=f])
unfolding zero_less_abs_iff content_image_stretch_interval
unfolding image_stretch_interval empty_as_interval euclidean_eq_iff[where 'a='a]
using assms
by auto
lemma has_integral_stretch_real:
fixes f :: "real \<Rightarrow> 'b::real_normed_vector"
assumes "(f has_integral i) {a..b}" and "m \<noteq> 0"
shows "((\<lambda>x. f (m * x)) has_integral (1 / \<bar>m\<bar>) *\<^sub>R i) ((\<lambda>x. x / m) ` {a..b})"
using has_integral_stretch [of f i a b "\<lambda>b. m"] assms by simp
lemma integrable_stretch:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
assumes "f integrable_on cbox a b"
and "\<forall>k\<in>Basis. m k \<noteq> 0"
shows "(\<lambda>x::'a. f (\<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k)) integrable_on
((\<lambda>x. \<Sum>k\<in>Basis. (1 / m k * (x\<bullet>k))*\<^sub>R k) ` cbox a b)"
using assms unfolding integrable_on_def
by (force dest: has_integral_stretch)
lemma vec_lambda_eq_sum:
shows "(\<chi> k. f k (x $ k)) = (\<Sum>k\<in>Basis. (f (axis_index k) (x \<bullet> k)) *\<^sub>R k)"
apply (simp add: Basis_vec_def cart_eq_inner_axis UNION_singleton_eq_range sum.reindex axis_eq_axis inj_on_def)
apply (simp add: vec_eq_iff axis_def if_distrib cong: if_cong)
done
lemma has_integral_stretch_cart:
fixes m :: "'n::finite \<Rightarrow> real"
assumes f: "(f has_integral i) (cbox a b)" and m: "\<And>k. m k \<noteq> 0"
shows "((\<lambda>x. f(\<chi> k. m k * x$k)) has_integral i /\<^sub>R \<bar>prod m UNIV\<bar>)
((\<lambda>x. \<chi> k. x$k / m k) ` (cbox a b))"
proof -
have *: "\<forall>k:: real^'n \<in> Basis. m (axis_index k) \<noteq> 0"
using axis_index by (simp add: m)
have eqp: "(\<Prod>k:: real^'n \<in> Basis. m (axis_index k)) = prod m UNIV"
by (simp add: Basis_vec_def UNION_singleton_eq_range prod.reindex axis_eq_axis inj_on_def)
show ?thesis
using has_integral_stretch [OF f *] vec_lambda_eq_sum [where f="\<lambda>i x. m i * x"] vec_lambda_eq_sum [where f="\<lambda>i x. x / m i"]
by (simp add: field_simps eqp)
qed
lemma image_stretch_interval_cart:
fixes m :: "'n::finite \<Rightarrow> real"
shows "(\<lambda>x. \<chi> k. m k * x$k) ` cbox a b =
(if cbox a b = {} then {}
else cbox (\<chi> k. min (m k * a$k) (m k * b$k)) (\<chi> k. max (m k * a$k) (m k * b$k)))"
proof -
have *: "(\<Sum>k\<in>Basis. min (m (axis_index k) * (a \<bullet> k)) (m (axis_index k) * (b \<bullet> k)) *\<^sub>R k)
= (\<chi> k. min (m k * a $ k) (m k * b $ k))"
"(\<Sum>k\<in>Basis. max (m (axis_index k) * (a \<bullet> k)) (m (axis_index k) * (b \<bullet> k)) *\<^sub>R k)
= (\<chi> k. max (m k * a $ k) (m k * b $ k))"
apply (simp_all add: Basis_vec_def cart_eq_inner_axis UNION_singleton_eq_range sum.reindex axis_eq_axis inj_on_def)
apply (simp_all add: vec_eq_iff axis_def if_distrib cong: if_cong)
done
show ?thesis
by (simp add: vec_lambda_eq_sum [where f="\<lambda>i x. m i * x"] image_stretch_interval eq_cbox *)
qed
subsection \<open>even more special cases\<close>
lemma uminus_interval_vector[simp]:
fixes a b :: "'a::euclidean_space"
shows "uminus ` cbox a b = cbox (-b) (-a)"
apply safe
apply (simp add: mem_box(2))
apply (rule_tac x="-x" in image_eqI)
apply (auto simp add: mem_box)
done
lemma has_integral_reflect_lemma[intro]:
assumes "(f has_integral i) (cbox a b)"
shows "((\<lambda>x. f(-x)) has_integral i) (cbox (-b) (-a))"
using has_integral_affinity[OF assms, of "-1" 0]
by auto
lemma has_integral_reflect_lemma_real[intro]:
assumes "(f has_integral i) {a..b::real}"
shows "((\<lambda>x. f(-x)) has_integral i) {-b .. -a}"
using assms
unfolding box_real[symmetric]
by (rule has_integral_reflect_lemma)
lemma has_integral_reflect[simp]:
"((\<lambda>x. f (-x)) has_integral i) (cbox (-b) (-a)) \<longleftrightarrow> (f has_integral i) (cbox a b)"
by (auto dest: has_integral_reflect_lemma)
lemma has_integral_reflect_real[simp]:
fixes a b::real
shows "((\<lambda>x. f (-x)) has_integral i) {-b..-a} \<longleftrightarrow> (f has_integral i) {a..b}"
by (metis has_integral_reflect interval_cbox)
lemma integrable_reflect[simp]: "(\<lambda>x. f(-x)) integrable_on cbox (-b) (-a) \<longleftrightarrow> f integrable_on cbox a b"
unfolding integrable_on_def by auto
lemma integrable_reflect_real[simp]: "(\<lambda>x. f(-x)) integrable_on {-b .. -a} \<longleftrightarrow> f integrable_on {a..b::real}"
unfolding box_real[symmetric]
by (rule integrable_reflect)
lemma integral_reflect[simp]: "integral (cbox (-b) (-a)) (\<lambda>x. f (-x)) = integral (cbox a b) f"
unfolding integral_def by auto
lemma integral_reflect_real[simp]: "integral {-b .. -a} (\<lambda>x. f (-x)) = integral {a..b::real} f"
unfolding box_real[symmetric]
by (rule integral_reflect)
subsection \<open>Stronger form of FCT; quite a tedious proof\<close>
lemma split_minus[simp]: "(\<lambda>(x, k). f x k) x - (\<lambda>(x, k). g x k) x = (\<lambda>(x, k). f x k - g x k) x"
by (simp add: split_def)
theorem fundamental_theorem_of_calculus_interior:
fixes f :: "real \<Rightarrow> 'a::real_normed_vector"
assumes "a \<le> b"
and contf: "continuous_on {a..b} f"
and derf: "\<And>x. x \<in> {a <..< b} \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
shows "(f' has_integral (f b - f a)) {a..b}"
proof (cases "a = b")
case True
then have *: "cbox a b = {b}" "f b - f a = 0"
by (auto simp add: order_antisym)
with True show ?thesis by auto
next
case False
with \<open>a \<le> b\<close> have ab: "a < b" by arith
show ?thesis
unfolding has_integral_factor_content_real
proof (intro allI impI)
fix e :: real
assume e: "e > 0"
then have eba8: "(e * (b-a)) / 8 > 0"
using ab by (auto simp add: field_simps)
note derf_exp = derf[unfolded has_vector_derivative_def has_derivative_at_alt]
have bounded: "\<And>x. x \<in> {a<..<b} \<Longrightarrow> bounded_linear (\<lambda>u. u *\<^sub>R f' x)"
using derf_exp by simp
have "\<forall>x \<in> box a b. \<exists>d>0. \<forall>y. norm (y-x) < d \<longrightarrow> norm (f y - f x - (y-x) *\<^sub>R f' x) \<le> e/2 * norm (y-x)"
(is "\<forall>x \<in> box a b. ?Q x")
proof
fix x assume x: "x \<in> box a b"
show "?Q x"
apply (rule allE [where x="e/2", OF derf_exp [THEN conjunct2, of x]])
using x e by auto
qed
from this [unfolded bgauge_existence_lemma]
obtain d where d: "\<And>x. 0 < d x"
"\<And>x y. \<lbrakk>x \<in> box a b; norm (y-x) < d x\<rbrakk>
\<Longrightarrow> norm (f y - f x - (y-x) *\<^sub>R f' x) \<le> e/2 * norm (y-x)"
by metis
have "bounded (f ` cbox a b)"
using compact_cbox assms by (auto simp: compact_imp_bounded compact_continuous_image)
then obtain B
where "0 < B" and B: "\<And>x. x \<in> f ` cbox a b \<Longrightarrow> norm x \<le> B"
unfolding bounded_pos by metis
obtain da where "0 < da"
and da: "\<And>c. \<lbrakk>a \<le> c; {a..c} \<subseteq> {a..b}; {a..c} \<subseteq> ball a da\<rbrakk>
\<Longrightarrow> norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \<le> (e * (b-a)) / 4"
proof -
have "continuous (at a within {a..b}) f"
using contf continuous_on_eq_continuous_within by force
with eba8 obtain k where "0 < k"
and k: "\<And>x. \<lbrakk>x \<in> {a..b}; 0 < norm (x-a); norm (x-a) < k\<rbrakk> \<Longrightarrow> norm (f x - f a) < e * (b-a) / 8"
unfolding continuous_within Lim_within dist_norm by metis
obtain l where l: "0 < l" "norm (l *\<^sub>R f' a) \<le> e * (b-a) / 8"
proof (cases "f' a = 0")
case True with ab e that show ?thesis by auto
next
case False
then show ?thesis
apply (rule_tac l="(e * (b-a)) / 8 / norm (f' a)" in that)
using ab e apply (auto simp add: field_simps)
done
qed
have "norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \<le> e * (b-a) / 4"
if "a \<le> c" "{a..c} \<subseteq> {a..b}" and bmin: "{a..c} \<subseteq> ball a (min k l)" for c
proof -
have minkl: "\<bar>a - x\<bar> < min k l" if "x \<in> {a..c}" for x
using bmin dist_real_def that by auto
then have lel: "\<bar>c - a\<bar> \<le> \<bar>l\<bar>"
using that by force
have "norm ((c - a) *\<^sub>R f' a - (f c - f a)) \<le> norm ((c - a) *\<^sub>R f' a) + norm (f c - f a)"
by (rule norm_triangle_ineq4)
also have "\<dots> \<le> e * (b-a) / 8 + e * (b-a) / 8"
proof (rule add_mono)
have "norm ((c - a) *\<^sub>R f' a) \<le> norm (l *\<^sub>R f' a)"
by (auto intro: mult_right_mono [OF lel])
also have "... \<le> e * (b-a) / 8"
by (rule l)
finally show "norm ((c - a) *\<^sub>R f' a) \<le> e * (b-a) / 8" .
next
have "norm (f c - f a) < e * (b-a) / 8"
proof (cases "a = c")
case True then show ?thesis
using eba8 by auto
next
case False show ?thesis
by (rule k) (use minkl \<open>a \<le> c\<close> that False in auto)
qed
then show "norm (f c - f a) \<le> e * (b-a) / 8" by simp
qed
finally show "norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \<le> e * (b-a) / 4"
unfolding content_real[OF \<open>a \<le> c\<close>] by auto
qed
then show ?thesis
by (rule_tac da="min k l" in that) (auto simp: l \<open>0 < k\<close>)
qed
obtain db where "0 < db"
and db: "\<And>c. \<lbrakk>c \<le> b; {c..b} \<subseteq> {a..b}; {c..b} \<subseteq> ball b db\<rbrakk>
\<Longrightarrow> norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \<le> (e * (b-a)) / 4"
proof -
have "continuous (at b within {a..b}) f"
using contf continuous_on_eq_continuous_within by force
with eba8 obtain k
where "0 < k"
and k: "\<And>x. \<lbrakk>x \<in> {a..b}; 0 < norm(x-b); norm(x-b) < k\<rbrakk>
\<Longrightarrow> norm (f b - f x) < e * (b-a) / 8"
unfolding continuous_within Lim_within dist_norm norm_minus_commute by metis
obtain l where l: "0 < l" "norm (l *\<^sub>R f' b) \<le> (e * (b-a)) / 8"
proof (cases "f' b = 0")
case True thus ?thesis
using ab e that by auto
next
case False then show ?thesis
apply (rule_tac l="(e * (b-a)) / 8 / norm (f' b)" in that)
using ab e by (auto simp add: field_simps)
qed
have "norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \<le> e * (b-a) / 4"
if "c \<le> b" "{c..b} \<subseteq> {a..b}" and bmin: "{c..b} \<subseteq> ball b (min k l)" for c
proof -
have minkl: "\<bar>b - x\<bar> < min k l" if "x \<in> {c..b}" for x
using bmin dist_real_def that by auto
then have lel: "\<bar>b - c\<bar> \<le> \<bar>l\<bar>"
using that by force
have "norm ((b - c) *\<^sub>R f' b - (f b - f c)) \<le> norm ((b - c) *\<^sub>R f' b) + norm (f b - f c)"
by (rule norm_triangle_ineq4)
also have "\<dots> \<le> e * (b-a) / 8 + e * (b-a) / 8"
proof (rule add_mono)
have "norm ((b - c) *\<^sub>R f' b) \<le> norm (l *\<^sub>R f' b)"
by (auto intro: mult_right_mono [OF lel])
also have "... \<le> e * (b-a) / 8"
by (rule l)
finally show "norm ((b - c) *\<^sub>R f' b) \<le> e * (b-a) / 8" .
next
have "norm (f b - f c) < e * (b-a) / 8"
proof (cases "b = c")
case True with eba8 show ?thesis
by auto
next
case False show ?thesis
by (rule k) (use minkl \<open>c \<le> b\<close> that False in auto)
qed
then show "norm (f b - f c) \<le> e * (b-a) / 8" by simp
qed
finally show "norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \<le> e * (b-a) / 4"
unfolding content_real[OF \<open>c \<le> b\<close>] by auto
qed
then show ?thesis
by (rule_tac db="min k l" in that) (auto simp: l \<open>0 < k\<close>)
qed
let ?d = "(\<lambda>x. ball x (if x=a then da else if x=b then db else d x))"
show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {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 {a..b})"
proof (rule exI, safe)
show "gauge ?d"
using ab \<open>db > 0\<close> \<open>da > 0\<close> d(1) by (auto intro: gauge_ball_dependent)
next
fix p
assume ptag: "p tagged_division_of {a..b}" and fine: "?d fine p"
let ?A = "{t. fst t \<in> {a, b}}"
note p = tagged_division_ofD[OF ptag]
have pA: "p = (p \<inter> ?A) \<union> (p - ?A)" "finite (p \<inter> ?A)" "finite (p - ?A)" "(p \<inter> ?A) \<inter> (p - ?A) = {}"
using ptag fine by auto
have le_xz: "\<And>w x y z::real. y \<le> z/2 \<Longrightarrow> w - x \<le> z/2 \<Longrightarrow> w + y \<le> x + z"
by arith
have non: False if xk: "(x,K) \<in> p" and "x \<noteq> a" "x \<noteq> b"
and less: "e * (Sup K - Inf K)/2 < norm (content K *\<^sub>R f' x - (f (Sup K) - f (Inf K)))"
for x K
proof -
obtain u v where k: "K = cbox u v"
using p(4) xk by blast
then have "u \<le> v" and uv: "{u, v} \<subseteq> cbox u v"
using p(2)[OF xk] by auto
then have result: "e * (v - u)/2 < norm ((v - u) *\<^sub>R f' x - (f v - f u))"
using less[unfolded k box_real interval_bounds_real content_real] by auto
then have "x \<in> box a b"
using p(2) p(3) \<open>x \<noteq> a\<close> \<open>x \<noteq> b\<close> xk by fastforce
with d have *: "\<And>y. norm (y-x) < d x
\<Longrightarrow> norm (f y - f x - (y-x) *\<^sub>R f' x) \<le> e/2 * norm (y-x)"
by metis
have xd: "norm (u - x) < d x" "norm (v - x) < d x"
using fineD[OF fine xk] \<open>x \<noteq> a\<close> \<open>x \<noteq> b\<close> uv
by (auto simp add: k subset_eq dist_commute dist_real_def)
have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) =
norm ((f u - f x - (u - x) *\<^sub>R f' x) - (f v - f x - (v - x) *\<^sub>R f' x))"
by (rule arg_cong[where f=norm]) (auto simp: scaleR_left.diff)
also have "\<dots> \<le> e/2 * norm (u - x) + e/2 * norm (v - x)"
by (metis norm_triangle_le_diff add_mono * xd)
also have "\<dots> \<le> e/2 * norm (v - u)"
using p(2)[OF xk] by (auto simp add: field_simps k)
also have "\<dots> < norm ((v - u) *\<^sub>R f' x - (f v - f u))"
using result by (simp add: \<open>u \<le> v\<close>)
finally have "e * (v - u)/2 < e * (v - u)/2"
using uv by auto
then show False by auto
qed
have "norm (\<Sum>(x, K)\<in>p - ?A. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K)))
\<le> (\<Sum>(x, K)\<in>p - ?A. norm (content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))))"
by (auto intro: sum_norm_le)
also have "... \<le> (\<Sum>n\<in>p - ?A. e * (case n of (x, k) \<Rightarrow> Sup k - Inf k)/2)"
using non by (fastforce intro: sum_mono)
finally have I: "norm (\<Sum>(x, k)\<in>p - ?A.
content k *\<^sub>R f' x - (f (Sup k) - f (Inf k)))
\<le> (\<Sum>n\<in>p - ?A. e * (case n of (x, k) \<Rightarrow> Sup k - Inf k))/2"
by (simp add: sum_divide_distrib)
have II: "norm (\<Sum>(x, k)\<in>p \<inter> ?A. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) -
(\<Sum>n\<in>p \<inter> ?A. e * (case n of (x, k) \<Rightarrow> Sup k - Inf k))
\<le> (\<Sum>n\<in>p - ?A. e * (case n of (x, k) \<Rightarrow> Sup k - Inf k))/2"
proof -
have ge0: "0 \<le> e * (Sup k - Inf k)" if xkp: "(x, k) \<in> p \<inter> ?A" for x k
proof -
obtain u v where uv: "k = cbox u v"
by (meson Int_iff xkp p(4))
with p(2) that uv have "cbox u v \<noteq> {}"
by blast
then show "0 \<le> e * ((Sup k) - (Inf k))"
unfolding uv using e by (auto simp add: field_simps)
qed
let ?B = "\<lambda>x. {t \<in> p. fst t = x \<and> content (snd t) \<noteq> 0}"
let ?C = "{t \<in> p. fst t \<in> {a, b} \<and> content (snd t) \<noteq> 0}"
have "norm (\<Sum>(x, k)\<in>p \<inter> {t. fst t \<in> {a, b}}. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \<le> e * (b-a)/2"
proof -
have *: "\<And>S f e. sum f S = sum f (p \<inter> ?C) \<Longrightarrow> norm (sum f (p \<inter> ?C)) \<le> e \<Longrightarrow> norm (sum f S) \<le> e"
by auto
have 1: "content K *\<^sub>R (f' x) - (f ((Sup K)) - f ((Inf K))) = 0"
if "(x,K) \<in> p \<inter> {t. fst t \<in> {a, b}} - p \<inter> ?C" for x K
proof -
have xk: "(x,K) \<in> p" and k0: "content K = 0"
using that by auto
then obtain u v where uv: "K = cbox u v"
using p(4) by blast
then have "u = v"
using xk k0 p(2) by force
then show "content K *\<^sub>R (f' x) - (f ((Sup K)) - f ((Inf K))) = 0"
using xk unfolding uv by auto
qed
have 2: "norm(\<Sum>(x,K)\<in>p \<inter> ?C. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \<le> e * (b-a)/2"
proof -
have norm_le: "norm (sum f S) \<le> e"
if "\<And>x y. \<lbrakk>x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> x = y" "\<And>x. x \<in> S \<Longrightarrow> norm (f x) \<le> e" "e > 0"
for S f and e :: real
proof (cases "S = {}")
case True
with that show ?thesis by auto
next
case False then obtain x where "x \<in> S"
by auto
then have "S = {x}"
using that(1) by auto
then show ?thesis
using \<open>x \<in> S\<close> that(2) by auto
qed
have *: "p \<inter> ?C = ?B a \<union> ?B b"
by blast
then have "norm (\<Sum>(x,K)\<in>p \<inter> ?C. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) =
norm (\<Sum>(x,K) \<in> ?B a \<union> ?B b. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K)))"
by simp
also have "... = norm ((\<Sum>(x,K) \<in> ?B a. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) +
(\<Sum>(x,K) \<in> ?B b. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))))"
apply (subst sum.union_disjoint)
using p(1) ab e by auto
also have "... \<le> e * (b - a) / 4 + e * (b - a) / 4"
proof (rule norm_triangle_le [OF add_mono])
have pa: "\<exists>v. k = cbox a v \<and> a \<le> v" if "(a, k) \<in> p" for k
using p(2) p(3) p(4) that by fastforce
show "norm (\<Sum>(x,K) \<in> ?B a. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \<le> e * (b - a) / 4"
proof (intro norm_le; clarsimp)
fix K K'
assume K: "(a, K) \<in> p" "(a, K') \<in> p" and ne0: "content K \<noteq> 0" "content K' \<noteq> 0"
with pa obtain v v' where v: "K = cbox a v" "a \<le> v" and v': "K' = cbox a v'" "a \<le> v'"
by blast
let ?v = "min v v'"
have "box a ?v \<subseteq> K \<inter> K'"
unfolding v v' by (auto simp add: mem_box)
then have "interior (box a (min v v')) \<subseteq> interior K \<inter> interior K'"
using interior_Int interior_mono by blast
moreover have "(a + ?v)/2 \<in> box a ?v"
using ne0 unfolding v v' content_eq_0 not_le
by (auto simp add: mem_box)
ultimately have "(a + ?v)/2 \<in> interior K \<inter> interior K'"
unfolding interior_open[OF open_box] by auto
then show "K = K'"
using p(5)[OF K] by auto
next
fix K
assume K: "(a, K) \<in> p" and ne0: "content K \<noteq> 0"
show "norm (content c *\<^sub>R f' a - (f (Sup c) - f (Inf c))) * 4 \<le> e * (b-a)"
if "(a, c) \<in> p" and ne0: "content c \<noteq> 0" for c
proof -
obtain v where v: "c = cbox a v" and "a \<le> v"
using pa[OF \<open>(a, c) \<in> p\<close>] by metis
then have "a \<in> {a..v}" "v \<le> b"
using p(3)[OF \<open>(a, c) \<in> p\<close>] by auto
moreover have "{a..v} \<subseteq> ball a da"
using fineD[OF \<open>?d fine p\<close> \<open>(a, c) \<in> p\<close>] by (simp add: v split: if_split_asm)
ultimately show ?thesis
unfolding v interval_bounds_real[OF \<open>a \<le> v\<close>] box_real
using da \<open>a \<le> v\<close> by auto
qed
qed (use ab e in auto)
next
have pb: "\<exists>v. k = cbox v b \<and> b \<ge> v" if "(b, k) \<in> p" for k
using p(2) p(3) p(4) that by fastforce
show "norm (\<Sum>(x,K) \<in> ?B b. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \<le> e * (b - a) / 4"
proof (intro norm_le; clarsimp)
fix K K'
assume K: "(b, K) \<in> p" "(b, K') \<in> p" and ne0: "content K \<noteq> 0" "content K' \<noteq> 0"
with pb obtain v v' where v: "K = cbox v b" "v \<le> b" and v': "K' = cbox v' b" "v' \<le> b"
by blast
let ?v = "max v v'"
have "box ?v b \<subseteq> K \<inter> K'"
unfolding v v' by (auto simp: mem_box)
then have "interior (box (max v v') b) \<subseteq> interior K \<inter> interior K'"
using interior_Int interior_mono by blast
moreover have " ((b + ?v)/2) \<in> box ?v b"
using ne0 unfolding v v' content_eq_0 not_le by (auto simp: mem_box)
ultimately have "((b + ?v)/2) \<in> interior K \<inter> interior K'"
unfolding interior_open[OF open_box] by auto
then show "K = K'"
using p(5)[OF K] by auto
next
fix K
assume K: "(b, K) \<in> p" and ne0: "content K \<noteq> 0"
show "norm (content c *\<^sub>R f' b - (f (Sup c) - f (Inf c))) * 4 \<le> e * (b-a)"
if "(b, c) \<in> p" and ne0: "content c \<noteq> 0" for c
proof -
obtain v where v: "c = cbox v b" and "v \<le> b"
using \<open>(b, c) \<in> p\<close> pb by blast
then have "v \<ge> a""b \<in> {v.. b}"
using p(3)[OF \<open>(b, c) \<in> p\<close>] by auto
moreover have "{v..b} \<subseteq> ball b db"
using fineD[OF \<open>?d fine p\<close> \<open>(b, c) \<in> p\<close>] box_real(2) v False by force
ultimately show ?thesis
using db v by auto
qed
qed (use ab e in auto)
qed
also have "... = e * (b-a)/2"
by simp
finally show "norm (\<Sum>(x,k)\<in>p \<inter> ?C.
content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \<le> e * (b-a)/2" .
qed
show "norm (\<Sum>(x, k)\<in>p \<inter> ?A. content k *\<^sub>R f' x - (f ((Sup k)) - f ((Inf k)))) \<le> e * (b-a)/2"
apply (rule * [OF sum.mono_neutral_right[OF pA(2)]])
using 1 2 by (auto simp: split_paired_all)
qed
also have "... = (\<Sum>n\<in>p. e * (case n of (x, k) \<Rightarrow> Sup k - Inf k))/2"
unfolding sum_distrib_left[symmetric]
apply (subst additive_tagged_division_1[OF \<open>a \<le> b\<close> ptag])
by auto
finally have norm_le: "norm (\<Sum>(x,K)\<in>p \<inter> {t. fst t \<in> {a, b}}. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K)))
\<le> (\<Sum>n\<in>p. e * (case n of (x, K) \<Rightarrow> Sup K - Inf K))/2" .
have le2: "\<And>x s1 s2::real. 0 \<le> s1 \<Longrightarrow> x \<le> (s1 + s2)/2 \<Longrightarrow> x - s1 \<le> s2/2"
by auto
show ?thesis
apply (rule le2 [OF sum_nonneg])
using ge0 apply force
unfolding sum.union_disjoint[OF pA(2-), symmetric] pA(1)[symmetric]
by (metis norm_le)
qed
note * = additive_tagged_division_1[OF assms(1) ptag, symmetric]
have "norm (\<Sum>(x,K)\<in>p \<inter> ?A \<union> (p - ?A). content K *\<^sub>R f' x - (f (Sup K) - f (Inf K)))
\<le> e * (\<Sum>(x,K)\<in>p \<inter> ?A \<union> (p - ?A). Sup K - Inf K)"
unfolding sum_distrib_left
unfolding sum.union_disjoint[OF pA(2-)]
using le_xz norm_triangle_le I II by blast
then
show "norm ((\<Sum>(x,K)\<in>p. content K *\<^sub>R f' x) - (f b - f a)) \<le> e * content {a..b}"
by (simp only: content_real[OF \<open>a \<le> b\<close>] *[of "\<lambda>x. x"] *[of f] sum_subtractf[symmetric] split_minus pA(1) [symmetric])
qed
qed
qed
subsection \<open>Stronger form with finite number of exceptional points\<close>
lemma fundamental_theorem_of_calculus_interior_strong:
fixes f :: "real \<Rightarrow> 'a::banach"
assumes "finite S"
and "a \<le> b" "\<And>x. x \<in> {a <..< b} - S \<Longrightarrow> (f has_vector_derivative f'(x)) (at x)"
and "continuous_on {a .. b} f"
shows "(f' has_integral (f b - f a)) {a .. b}"
using assms
proof (induction arbitrary: a b)
case empty
then show ?case
using fundamental_theorem_of_calculus_interior by force
next
case (insert x S)
show ?case
proof (cases "x \<in> {a<..<b}")
case False then show ?thesis
using insert by blast
next
case True then have "a < x" "x < b"
by auto
have "(f' has_integral f x - f a) {a..x}" "(f' has_integral f b - f x) {x..b}"
using \<open>continuous_on {a..b} f\<close> \<open>a < x\<close> \<open>x < b\<close> continuous_on_subset by (force simp: intro!: insert)+
then have "(f' has_integral f x - f a + (f b - f x)) {a..b}"
using \<open>a < x\<close> \<open>x < b\<close> has_integral_combine less_imp_le by blast
then show ?thesis
by simp
qed
qed
corollary fundamental_theorem_of_calculus_strong:
fixes f :: "real \<Rightarrow> 'a::banach"
assumes "finite S"
and "a \<le> b"
and vec: "\<And>x. x \<in> {a..b} - S \<Longrightarrow> (f has_vector_derivative f'(x)) (at x)"
and "continuous_on {a..b} f"
shows "(f' has_integral (f b - f a)) {a..b}"
by (rule fundamental_theorem_of_calculus_interior_strong [OF \<open>finite S\<close>]) (force simp: assms)+
proposition indefinite_integral_continuous_left:
fixes f:: "real \<Rightarrow> 'a::banach"
assumes intf: "f integrable_on {a..b}" and "a < c" "c \<le> b" "e > 0"
obtains d where "d > 0"
and "\<forall>t. c - d < t \<and> t \<le> c \<longrightarrow> norm (integral {a..c} f - integral {a..t} f) < e"
proof -
obtain w where "w > 0" and w: "\<And>t. \<lbrakk>c - w < t; t < c\<rbrakk> \<Longrightarrow> norm (f c) * norm(c - t) < e/3"
proof (cases "f c = 0")
case False
hence e3: "0 < e/3 / norm (f c)" using \<open>e>0\<close> by simp
moreover have "norm (f c) * norm (c - t) < e/3"
if "t < c" and "c - e/3 / norm (f c) < t" for t
proof -
have "norm (c - t) < e/3 / norm (f c)"
using that by auto
then show "norm (f c) * norm (c - t) < e/3"
by (metis e3 mult.commute norm_not_less_zero pos_less_divide_eq zero_less_divide_iff)
qed
ultimately show ?thesis
using that by auto
next
case True then show ?thesis
using \<open>e > 0\<close> that by auto
qed
let ?SUM = "\<lambda>p. (\<Sum>(x,K) \<in> p. content K *\<^sub>R f x)"
have e3: "e/3 > 0"
using \<open>e > 0\<close> by auto
have "f integrable_on {a..c}"
apply (rule integrable_subinterval_real[OF intf])
using \<open>a < c\<close> \<open>c \<le> b\<close> by auto
then obtain d1 where "gauge d1" and d1:
"\<And>p. \<lbrakk>p tagged_division_of {a..c}; d1 fine p\<rbrakk> \<Longrightarrow> norm (?SUM p - integral {a..c} f) < e/3"
using integrable_integral has_integral_real e3 by metis
define d where [abs_def]: "d x = ball x w \<inter> d1 x" for x
have "gauge d"
unfolding d_def using \<open>w > 0\<close> \<open>gauge d1\<close> by auto
then obtain k where "0 < k" and k: "ball c k \<subseteq> d c"
by (meson gauge_def open_contains_ball)
let ?d = "min k (c - a)/2"
show thesis
proof (intro that[of ?d] allI impI, safe)
show "?d > 0"
using \<open>0 < k\<close> \<open>a < c\<close> by auto
next
fix t
assume t: "c - ?d < t" "t \<le> c"
show "norm (integral ({a..c}) f - integral ({a..t}) f) < e"
proof (cases "t < c")
case False with \<open>t \<le> c\<close> show ?thesis
by (simp add: \<open>e > 0\<close>)
next
case True
have "f integrable_on {a..t}"
apply (rule integrable_subinterval_real[OF intf])
using \<open>t < c\<close> \<open>c \<le> b\<close> by auto
then obtain d2 where d2: "gauge d2"
"\<And>p. p tagged_division_of {a..t} \<and> d2 fine p \<Longrightarrow> norm (?SUM p - integral {a..t} f) < e/3"
using integrable_integral has_integral_real e3 by metis
define d3 where "d3 x = (if x \<le> t then d1 x \<inter> d2 x else d1 x)" for x
have "gauge d3"
using \<open>gauge d1\<close> \<open>gauge d2\<close> unfolding d3_def gauge_def by auto
then obtain p where ptag: "p tagged_division_of {a..t}" and pfine: "d3 fine p"
by (metis box_real(2) fine_division_exists)
note p' = tagged_division_ofD[OF ptag]
have pt: "(x,K)\<in>p \<Longrightarrow> x \<le> t" for x K
by (meson atLeastAtMost_iff p'(2) p'(3) subsetCE)
with pfine have "d2 fine p"
unfolding fine_def d3_def by fastforce
then have d2_fin: "norm (?SUM p - integral {a..t} f) < e/3"
using d2(2) ptag by auto
have eqs: "{a..c} \<inter> {x. x \<le> t} = {a..t}" "{a..c} \<inter> {x. x \<ge> t} = {t..c}"
using t by (auto simp add: field_simps)
have "p \<union> {(c, {t..c})} tagged_division_of {a..c}"
apply (rule tagged_division_Un_interval_real[of _ _ _ 1 "t"])
using \<open>t \<le> c\<close> by (auto simp: eqs ptag tagged_division_of_self_real)
moreover
have "d1 fine p \<union> {(c, {t..c})}"
unfolding fine_def
proof safe
fix x K y
assume "(x,K) \<in> p" and "y \<in> K" then show "y \<in> d1 x"
by (metis Int_iff d3_def subsetD fineD pfine)
next
fix x assume "x \<in> {t..c}"
then have "dist c x < k"
using t(1) by (auto simp add: field_simps dist_real_def)
with k show "x \<in> d1 c"
unfolding d_def by auto
qed
ultimately have d1_fin: "norm (?SUM(p \<union> {(c, {t..c})}) - integral {a..c} f) < e/3"
using d1 by metis
have SUMEQ: "?SUM(p \<union> {(c, {t..c})}) = (c - t) *\<^sub>R f c + ?SUM p"
proof -
have "?SUM(p \<union> {(c, {t..c})}) = (content{t..c} *\<^sub>R f c) + ?SUM p"
proof (subst sum.union_disjoint)
show "p \<inter> {(c, {t..c})} = {}"
using \<open>t < c\<close> pt by force
qed (use p'(1) in auto)
also have "... = (c - t) *\<^sub>R f c + ?SUM p"
using \<open>t \<le> c\<close> by auto
finally show ?thesis .
qed
have "c - k < t"
using \<open>k>0\<close> t(1) by (auto simp add: field_simps)
moreover have "k \<le> w"
proof (rule ccontr)
assume "\<not> k \<le> w"
then have "c + (k + w) / 2 \<notin> d c"
by (auto simp add: field_simps not_le not_less dist_real_def d_def)
then have "c + (k + w) / 2 \<notin> ball c k"
using k by blast
then show False
using \<open>0 < w\<close> \<open>\<not> k \<le> w\<close> dist_real_def by auto
qed
ultimately have cwt: "c - w < t"
by (auto simp add: field_simps)
have eq: "integral {a..c} f - integral {a..t} f = -(((c - t) *\<^sub>R f c + ?SUM p) -
integral {a..c} f) + (?SUM p - integral {a..t} f) + (c - t) *\<^sub>R f c"
by auto
have "norm (integral {a..c} f - integral {a..t} f) < e/3 + e/3 + e/3"
unfolding eq
proof (intro norm_triangle_lt add_strict_mono)
show "norm (- ((c - t) *\<^sub>R f c + ?SUM p - integral {a..c} f)) < e/3"
by (metis SUMEQ d1_fin norm_minus_cancel)
show "norm (?SUM p - integral {a..t} f) < e/3"
using d2_fin by blast
show "norm ((c - t) *\<^sub>R f c) < e/3"
using w cwt \<open>t < c\<close> by simp (simp add: field_simps)
qed
then show ?thesis by simp
qed
qed
qed
lemma indefinite_integral_continuous_right:
fixes f :: "real \<Rightarrow> 'a::banach"
assumes "f integrable_on {a..b}"
and "a \<le> c"
and "c < b"
and "e > 0"
obtains d where "0 < d"
and "\<forall>t. c \<le> t \<and> t < c + d \<longrightarrow> norm (integral {a..c} f - integral {a..t} f) < e"
proof -
have intm: "(\<lambda>x. f (- x)) integrable_on {-b .. -a}" "- b < - c" "- c \<le> - a"
using assms by auto
from indefinite_integral_continuous_left[OF intm \<open>e>0\<close>]
obtain d where "0 < d"
and d: "\<And>t. \<lbrakk>- c - d < t; t \<le> -c\<rbrakk>
\<Longrightarrow> norm (integral {- b..- c} (\<lambda>x. f (-x)) - integral {- b..t} (\<lambda>x. f (-x))) < e"
by metis
let ?d = "min d (b - c)"
show ?thesis
proof (intro that[of "?d"] allI impI)
show "0 < ?d"
using \<open>0 < d\<close> \<open>c < b\<close> by auto
fix t :: real
assume t: "c \<le> t \<and> t < c + ?d"
have *: "integral {a..c} f = integral {a..b} f - integral {c..b} f"
"integral {a..t} f = integral {a..b} f - integral {t..b} f"
apply (simp_all only: algebra_simps)
using assms t by (auto simp: integral_combine)
have "(- c) - d < (- t)" "- t \<le> - c"
using t by auto
from d[OF this] show "norm (integral {a..c} f - integral {a..t} f) < e"
by (auto simp add: algebra_simps norm_minus_commute *)
qed
qed
lemma indefinite_integral_continuous_1:
fixes f :: "real \<Rightarrow> 'a::banach"
assumes "f integrable_on {a..b}"
shows "continuous_on {a..b} (\<lambda>x. integral {a..x} f)"
proof -
have "\<exists>d>0. \<forall>x'\<in>{a..b}. dist x' x < d \<longrightarrow> dist (integral {a..x'} f) (integral {a..x} f) < e"
if x: "x \<in> {a..b}" and "e > 0" for x e :: real
proof (cases "a = b")
case True
with that show ?thesis by force
next
case False
with x have "a < b" by force
with x consider "x = a" | "x = b" | "a < x" "x < b"
by force
then show ?thesis
proof cases
case 1 show ?thesis
apply (rule indefinite_integral_continuous_right [OF assms _ \<open>a < b\<close> \<open>e > 0\<close>], force)
using \<open>x = a\<close> apply (force simp: dist_norm algebra_simps)
done
next
case 2 show ?thesis
apply (rule indefinite_integral_continuous_left [OF assms \<open>a < b\<close> _ \<open>e > 0\<close>], force)
using \<open>x = b\<close> apply (force simp: dist_norm norm_minus_commute algebra_simps)
done
next
case 3
obtain d1 where "0 < d1"
and d1: "\<And>t. \<lbrakk>x - d1 < t; t \<le> x\<rbrakk> \<Longrightarrow> norm (integral {a..x} f - integral {a..t} f) < e"
using 3 by (auto intro: indefinite_integral_continuous_left [OF assms \<open>a < x\<close> _ \<open>e > 0\<close>])
obtain d2 where "0 < d2"
and d2: "\<And>t. \<lbrakk>x \<le> t; t < x + d2\<rbrakk> \<Longrightarrow> norm (integral {a..x} f - integral {a..t} f) < e"
using 3 by (auto intro: indefinite_integral_continuous_right [OF assms _ \<open>x < b\<close> \<open>e > 0\<close>])
show ?thesis
proof (intro exI ballI conjI impI)
show "0 < min d1 d2"
using \<open>0 < d1\<close> \<open>0 < d2\<close> by simp
show "dist (integral {a..y} f) (integral {a..x} f) < e"
if "y \<in> {a..b}" "dist y x < min d1 d2" for y
proof (cases "y < x")
case True
with that d1 show ?thesis by (auto simp: dist_commute dist_norm)
next
case False
with that d2 show ?thesis
by (auto simp: dist_commute dist_norm)
qed
qed
qed
qed
then show ?thesis
by (auto simp: continuous_on_iff)
qed
lemma indefinite_integral_continuous_1':
fixes f::"real \<Rightarrow> 'a::banach"
assumes "f integrable_on {a..b}"
shows "continuous_on {a..b} (\<lambda>x. integral {x..b} f)"
proof -
have "integral {a..b} f - integral {a..x} f = integral {x..b} f" if "x \<in> {a..b}" for x
using integral_combine[OF _ _ assms, of x] that
by (auto simp: algebra_simps)
with _ show ?thesis
by (rule continuous_on_eq) (auto intro!: continuous_intros indefinite_integral_continuous_1 assms)
qed
theorem integral_has_vector_derivative':
fixes f :: "real \<Rightarrow> 'b::banach"
assumes "continuous_on {a..b} f"
and "x \<in> {a..b}"
shows "((\<lambda>u. integral {u..b} f) has_vector_derivative - f x) (at x within {a..b})"
proof -
have *: "integral {x..b} f = integral {a .. b} f - integral {a .. x} f" if "a \<le> x" "x \<le> b" for x
using integral_combine[of a x b for x, OF that integrable_continuous_real[OF assms(1)]]
by (simp add: algebra_simps)
show ?thesis
using \<open>x \<in> _\<close> *
by (rule has_vector_derivative_transform)
(auto intro!: derivative_eq_intros assms integral_has_vector_derivative)
qed
lemma integral_has_real_derivative':
assumes "continuous_on {a..b} g"
assumes "t \<in> {a..b}"
shows "((\<lambda>x. integral {x..b} g) has_real_derivative -g t) (at t within {a..b})"
using integral_has_vector_derivative'[OF assms]
by (auto simp: has_field_derivative_iff_has_vector_derivative)
subsection \<open>This doesn't directly involve integration, but that gives an easy proof\<close>
lemma has_derivative_zero_unique_strong_interval:
fixes f :: "real \<Rightarrow> 'a::banach"
assumes "finite k"
and contf: "continuous_on {a..b} f"
and "f a = y"
and fder: "\<And>x. x \<in> {a..b} - k \<Longrightarrow> (f has_derivative (\<lambda>h. 0)) (at x within {a..b})"
and x: "x \<in> {a..b}"
shows "f x = y"
proof -
have "a \<le> b" "a \<le> x"
using assms by auto
have "((\<lambda>x. 0::'a) has_integral f x - f a) {a..x}"
proof (rule fundamental_theorem_of_calculus_interior_strong[OF \<open>finite k\<close> \<open>a \<le> x\<close>]; clarify?)
have "{a..x} \<subseteq> {a..b}"
using x by auto
then show "continuous_on {a..x} f"
by (rule continuous_on_subset[OF contf])
show "(f has_vector_derivative 0) (at z)" if z: "z \<in> {a<..<x}" and notin: "z \<notin> k" for z
unfolding has_vector_derivative_def
proof (simp add: at_within_open[OF z, symmetric])
show "(f has_derivative (\<lambda>x. 0)) (at z within {a<..<x})"
by (rule has_derivative_within_subset [OF fder]) (use x z notin in auto)
qed
qed
from has_integral_unique[OF has_integral_0 this]
show ?thesis
unfolding assms by auto
qed
subsection \<open>Generalize a bit to any convex set\<close>
lemma has_derivative_zero_unique_strong_convex:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
assumes "convex S" "finite K"
and contf: "continuous_on S f"
and "c \<in> S" "f c = y"
and derf: "\<And>x. x \<in> S - K \<Longrightarrow> (f has_derivative (\<lambda>h. 0)) (at x within S)"
and "x \<in> S"
shows "f x = y"
proof (cases "x = c")
case True with \<open>f c = y\<close> show ?thesis
by blast
next
case False
let ?\<phi> = "\<lambda>u. (1 - u) *\<^sub>R c + u *\<^sub>R x"
have contf': "continuous_on {0 ..1} (f \<circ> ?\<phi>)"
proof (rule continuous_intros continuous_on_subset[OF contf])+
show "(\<lambda>u. (1 - u) *\<^sub>R c + u *\<^sub>R x) ` {0..1} \<subseteq> S"
using \<open>convex S\<close> \<open>x \<in> S\<close> \<open>c \<in> S\<close> by (auto simp add: convex_alt algebra_simps)
qed
have "t = u" if "?\<phi> t = ?\<phi> u" for t u
proof -
from that have "(t - u) *\<^sub>R x = (t - u) *\<^sub>R c"
by (auto simp add: algebra_simps)
then show ?thesis
using \<open>x \<noteq> c\<close> by auto
qed
then have eq: "(SOME t. ?\<phi> t = ?\<phi> u) = u" for u
by blast
then have "(?\<phi> -` K) \<subseteq> (\<lambda>z. SOME t. ?\<phi> t = z) ` K"
by (clarsimp simp: image_iff) (metis (no_types) eq)
then have fin: "finite (?\<phi> -` K)"
by (rule finite_surj[OF \<open>finite K\<close>])
have derf': "((\<lambda>u. f (?\<phi> u)) has_derivative (\<lambda>h. 0)) (at t within {0..1})"
if "t \<in> {0..1} - {t. ?\<phi> t \<in> K}" for t
proof -
have df: "(f has_derivative (\<lambda>h. 0)) (at (?\<phi> t) within ?\<phi> ` {0..1})"
apply (rule has_derivative_within_subset [OF derf])
using \<open>convex S\<close> \<open>x \<in> S\<close> \<open>c \<in> S\<close> that by (auto simp add: convex_alt algebra_simps)
have "(f \<circ> ?\<phi> has_derivative (\<lambda>x. 0) \<circ> (\<lambda>z. (0 - z *\<^sub>R c) + z *\<^sub>R x)) (at t within {0..1})"
by (rule derivative_eq_intros df | simp)+
then show ?thesis
unfolding o_def .
qed
have "(f \<circ> ?\<phi>) 1 = y"
apply (rule has_derivative_zero_unique_strong_interval[OF fin contf'])
unfolding o_def using \<open>f c = y\<close> derf' by auto
then show ?thesis
by auto
qed
text \<open>Also to any open connected set with finite set of exceptions. Could
generalize to locally convex set with limpt-free set of exceptions.\<close>
lemma has_derivative_zero_unique_strong_connected:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
assumes "connected S"
and "open S"
and "finite K"
and contf: "continuous_on S f"
and "c \<in> S"
and "f c = y"
and derf: "\<And>x. x \<in> S - K \<Longrightarrow> (f has_derivative (\<lambda>h. 0)) (at x within S)"
and "x \<in> S"
shows "f x = y"
proof -
have "\<exists>e>0. ball x e \<subseteq> (S \<inter> f -` {f x})" if "x \<in> S" for x
proof -
obtain e where "0 < e" and e: "ball x e \<subseteq> S"
using \<open>x \<in> S\<close> \<open>open S\<close> open_contains_ball by blast
have "ball x e \<subseteq> {u \<in> S. f u \<in> {f x}}"
proof safe
fix y
assume y: "y \<in> ball x e"
then show "y \<in> S"
using e by auto
show "f y = f x"
proof (rule has_derivative_zero_unique_strong_convex[OF convex_ball \<open>finite K\<close>])
show "continuous_on (ball x e) f"
using contf continuous_on_subset e by blast
show "(f has_derivative (\<lambda>h. 0)) (at u within ball x e)"
if "u \<in> ball x e - K" for u
by (metis Diff_iff contra_subsetD derf e has_derivative_within_subset that)
qed (use y e \<open>0 < e\<close> in auto)
qed
then show "\<exists>e>0. ball x e \<subseteq> (S \<inter> f -` {f x})"
using \<open>0 < e\<close> by blast
qed
then have "openin (top_of_set S) (S \<inter> f -` {y})"
by (auto intro!: open_openin_trans[OF \<open>open S\<close>] simp: open_contains_ball)
moreover have "closedin (top_of_set S) (S \<inter> f -` {y})"
by (force intro!: continuous_closedin_preimage [OF contf])
ultimately have "(S \<inter> f -` {y}) = {} \<or> (S \<inter> f -` {y}) = S"
using \<open>connected S\<close> by (simp add: connected_clopen)
then show ?thesis
using \<open>x \<in> S\<close> \<open>f c = y\<close> \<open>c \<in> S\<close> by auto
qed
lemma has_derivative_zero_connected_constant:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
assumes "connected S"
and "open S"
and "finite k"
and "continuous_on S f"
and "\<forall>x\<in>(S - k). (f has_derivative (\<lambda>h. 0)) (at x within S)"
obtains c where "\<And>x. x \<in> S \<Longrightarrow> f(x) = c"
proof (cases "S = {}")
case True
then show ?thesis
by (metis empty_iff that)
next
case False
then obtain c where "c \<in> S"
by (metis equals0I)
then show ?thesis
by (metis has_derivative_zero_unique_strong_connected assms that)
qed
lemma DERIV_zero_connected_constant:
fixes f :: "'a::{real_normed_field,euclidean_space} \<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)
subsection \<open>Integrating characteristic function of an interval\<close>
lemma has_integral_restrict_open_subinterval:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
assumes intf: "(f has_integral i) (cbox c d)"
and cb: "cbox c d \<subseteq> cbox a b"
shows "((\<lambda>x. if x \<in> box c d then f x else 0) has_integral i) (cbox a b)"
proof (cases "cbox c d = {}")
case True
then have "box c d = {}"
by (metis bot.extremum_uniqueI box_subset_cbox)
then show ?thesis
using True intf by auto
next
case False
then obtain p where pdiv: "p division_of cbox a b" and inp: "cbox c d \<in> p"
using cb partial_division_extend_1 by blast
define g where [abs_def]: "g x = (if x \<in>box c d then f x else 0)" for x
interpret operative "lift_option plus" "Some (0 :: 'b)"
"\<lambda>i. if g integrable_on i then Some (integral i g) else None"
by (fact operative_integralI)
note operat = division [OF pdiv, symmetric]
show ?thesis
proof (cases "(g has_integral i) (cbox a b)")
case True then show ?thesis
by (simp add: g_def)
next
case False
have iterate:"F (\<lambda>i. if g integrable_on i then Some (integral i g) else None) (p - {cbox c d}) = Some 0"
proof (intro neutral ballI)
fix x
assume x: "x \<in> p - {cbox c d}"
then have "x \<in> p"
by auto
then obtain u v where uv: "x = cbox u v"
using pdiv by blast
have "interior x \<inter> interior (cbox c d) = {}"
using pdiv inp x by blast
then have "(g has_integral 0) x"
unfolding uv using has_integral_spike_interior[where f="\<lambda>x. 0"]
by (metis (no_types, lifting) disjoint_iff_not_equal g_def has_integral_0_eq interior_cbox)
then show "(if g integrable_on x then Some (integral x g) else None) = Some 0"
by auto
qed
interpret comm_monoid_set "lift_option plus" "Some (0 :: 'b)"
by (intro comm_monoid_set.intro comm_monoid_lift_option add.comm_monoid_axioms)
have intg: "g integrable_on cbox c d"
using integrable_spike_interior[where f=f]
by (meson g_def has_integral_integrable intf)
moreover have "integral (cbox c d) g = i"
proof (rule has_integral_unique[OF has_integral_spike_interior intf])
show "\<And>x. x \<in> box c d \<Longrightarrow> f x = g x"
by (auto simp: g_def)
show "(g has_integral integral (cbox c d) g) (cbox c d)"
by (rule integrable_integral[OF intg])
qed
ultimately have "F (\<lambda>A. if g integrable_on A then Some (integral A g) else None) p = Some i"
by (metis (full_types, lifting) division_of_finite inp iterate pdiv remove right_neutral)
then
have "(g has_integral i) (cbox a b)"
by (metis integrable_on_def integral_unique operat option.inject option.simps(3))
with False show ?thesis
by blast
qed
qed
lemma has_integral_restrict_closed_subinterval:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
assumes "(f has_integral i) (cbox c d)"
and "cbox c d \<subseteq> cbox a b"
shows "((\<lambda>x. if x \<in> cbox c d then f x else 0) has_integral i) (cbox a b)"
proof -
note has_integral_restrict_open_subinterval[OF assms]
note * = has_integral_spike[OF negligible_frontier_interval _ this]
show ?thesis
by (rule *[of c d]) (use box_subset_cbox[of c d] in auto)
qed
lemma has_integral_restrict_closed_subintervals_eq:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
assumes "cbox c d \<subseteq> cbox a b"
shows "((\<lambda>x. if x \<in> cbox c d then f x else 0) has_integral i) (cbox a b) \<longleftrightarrow> (f has_integral i) (cbox c d)"
(is "?l = ?r")
proof (cases "cbox c d = {}")
case False
let ?g = "\<lambda>x. if x \<in> cbox c d then f x else 0"
show ?thesis
proof
assume ?l
then have "?g integrable_on cbox c d"
using assms has_integral_integrable integrable_subinterval by blast
then have "f integrable_on cbox c d"
by (rule integrable_eq) auto
moreover then have "i = integral (cbox c d) f"
by (meson \<open>((\<lambda>x. if x \<in> cbox c d then f x else 0) has_integral i) (cbox a b)\<close> assms has_integral_restrict_closed_subinterval has_integral_unique integrable_integral)
ultimately show ?r by auto
next
assume ?r then show ?l
by (rule has_integral_restrict_closed_subinterval[OF _ assms])
qed
qed auto
text \<open>Hence we can apply the limit process uniformly to all integrals.\<close>
lemma has_integral':
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
shows "(f has_integral i) S \<longleftrightarrow>
(\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<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 - i) < e))"
(is "?l \<longleftrightarrow> (\<forall>e>0. ?r e)")
proof (cases "\<exists>a b. S = cbox a b")
case False then show ?thesis
by (simp add: has_integral_alt)
next
case True
then obtain a b where S: "S = cbox a b"
by blast
obtain B where " 0 < B" and B: "\<And>x. x \<in> cbox a b \<Longrightarrow> norm x \<le> B"
using bounded_cbox[unfolded bounded_pos] by blast
show ?thesis
proof safe
fix e :: real
assume ?l and "e > 0"
have "((\<lambda>x. if x \<in> S then f x else 0) has_integral i) (cbox c d)"
if "ball 0 (B+1) \<subseteq> cbox c d" for c d
unfolding S using B that
by (force intro: \<open>?l\<close>[unfolded S] has_integral_restrict_closed_subinterval)
then show "?r e"
apply (rule_tac x="B+1" in exI)
using \<open>B>0\<close> \<open>e>0\<close> by force
next
assume as: "\<forall>e>0. ?r e"
then obtain C
where C: "\<And>a b. ball 0 C \<subseteq> cbox a b \<Longrightarrow>
\<exists>z. ((\<lambda>x. if x \<in> S then f x else 0) has_integral z) (cbox a b)"
by (meson zero_less_one)
define c :: 'n where "c = (\<Sum>i\<in>Basis. (- max B C) *\<^sub>R i)"
define d :: 'n where "d = (\<Sum>i\<in>Basis. max B C *\<^sub>R i)"
have "c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i" if "norm x \<le> B" "i \<in> Basis" for x i
using that and Basis_le_norm[OF \<open>i\<in>Basis\<close>, of x]
by (auto simp add: field_simps sum_negf c_def d_def)
then have c_d: "cbox a b \<subseteq> cbox c d"
by (meson B mem_box(2) subsetI)
have "c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i"
if x: "norm (0 - x) < C" and i: "i \<in> Basis" for x i
using Basis_le_norm[OF i, of x] x i by (auto simp: sum_negf c_def d_def)
then have "ball 0 C \<subseteq> cbox c d"
by (auto simp: mem_box dist_norm)
with C obtain y where y: "(f has_integral y) (cbox a b)"
using c_d has_integral_restrict_closed_subintervals_eq S by blast
have "y = i"
proof (rule ccontr)
assume "y \<noteq> i"
then have "0 < norm (y - i)"
by auto
from as[rule_format,OF this]
obtain C where C: "\<And>a b. ball 0 C \<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-i) < norm (y-i)"
by auto
define c :: 'n where "c = (\<Sum>i\<in>Basis. (- max B C) *\<^sub>R i)"
define d :: 'n where "d = (\<Sum>i\<in>Basis. max B C *\<^sub>R i)"
have "c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i"
if "norm x \<le> B" and "i \<in> Basis" for x i
using that Basis_le_norm[of i x] by (auto simp add: field_simps sum_negf c_def d_def)
then have c_d: "cbox a b \<subseteq> cbox c d"
by (simp add: B mem_box(2) subset_eq)
have "c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i" if "norm (0 - x) < C" and "i \<in> Basis" for x i
using Basis_le_norm[of i x] that by (auto simp: sum_negf c_def d_def)
then have "ball 0 C \<subseteq> cbox c d"
by (auto simp: mem_box dist_norm)
with C obtain z where z: "(f has_integral z) (cbox a b)" "norm (z-i) < norm (y-i)"
using has_integral_restrict_closed_subintervals_eq[OF c_d] S by blast
moreover then have "z = y"
by (blast intro: has_integral_unique[OF _ y])
ultimately show False
by auto
qed
then show ?l
using y by (auto simp: S)
qed
qed
lemma has_integral_le:
fixes f :: "'n::euclidean_space \<Rightarrow> real"
assumes fg: "(f has_integral i) S" "(g has_integral j) S"
and le: "\<And>x. x \<in> S \<Longrightarrow> f x \<le> g x"
shows "i \<le> j"
using has_integral_component_le[OF _ fg, of 1] le by auto
lemma integral_le:
fixes f :: "'n::euclidean_space \<Rightarrow> real"
assumes "f integrable_on S"
and "g integrable_on S"
and "\<And>x. x \<in> S \<Longrightarrow> f x \<le> g x"
shows "integral S f \<le> integral S g"
by (rule has_integral_le[OF assms(1,2)[unfolded has_integral_integral] assms(3)])
lemma has_integral_nonneg:
fixes f :: "'n::euclidean_space \<Rightarrow> real"
assumes "(f has_integral i) S"
and "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f x"
shows "0 \<le> i"
using has_integral_component_nonneg[of 1 f i S]
unfolding o_def
using assms
by auto
lemma integral_nonneg:
fixes f :: "'n::euclidean_space \<Rightarrow> real"
assumes f: "f integrable_on S" and 0: "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f x"
shows "0 \<le> integral S f"
by (rule has_integral_nonneg[OF f[unfolded has_integral_integral] 0])
text \<open>Hence a general restriction property.\<close>
lemma has_integral_restrict [simp]:
fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: banach"
assumes "S \<subseteq> T"
shows "((\<lambda>x. if x \<in> S then f x else 0) has_integral i) T \<longleftrightarrow> (f has_integral i) S"
proof -
have *: "\<And>x. (if x \<in> T then if x \<in> S then f x else 0 else 0) = (if x\<in>S then f x else 0)"
using assms by auto
show ?thesis
apply (subst(2) has_integral')
apply (subst has_integral')
apply (simp add: *)
done
qed
corollary has_integral_restrict_UNIV:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
shows "((\<lambda>x. if x \<in> s then f x else 0) has_integral i) UNIV \<longleftrightarrow> (f has_integral i) s"
by auto
lemma has_integral_restrict_Int:
fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: banach"
shows "((\<lambda>x. if x \<in> S then f x else 0) has_integral i) T \<longleftrightarrow> (f has_integral i) (S \<inter> T)"
proof -
have "((\<lambda>x. if x \<in> T then if x \<in> S then f x else 0 else 0) has_integral i) UNIV =
((\<lambda>x. if x \<in> S \<inter> T then f x else 0) has_integral i) UNIV"
by (rule has_integral_cong) auto
then show ?thesis
using has_integral_restrict_UNIV by fastforce
qed
lemma integral_restrict_Int:
fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: banach"
shows "integral T (\<lambda>x. if x \<in> S then f x else 0) = integral (S \<inter> T) f"
by (metis (no_types, lifting) has_integral_cong has_integral_restrict_Int integrable_integral integral_unique not_integrable_integral)
lemma integrable_restrict_Int:
fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: banach"
shows "(\<lambda>x. if x \<in> S then f x else 0) integrable_on T \<longleftrightarrow> f integrable_on (S \<inter> T)"
using has_integral_restrict_Int by fastforce
lemma has_integral_on_superset:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes f: "(f has_integral i) S"
and "\<And>x. x \<notin> S \<Longrightarrow> f x = 0"
and "S \<subseteq> T"
shows "(f has_integral i) T"
proof -
have "(\<lambda>x. if x \<in> S then f x else 0) = (\<lambda>x. if x \<in> T then f x else 0)"
using assms by fastforce
with f show ?thesis
by (simp only: has_integral_restrict_UNIV [symmetric, of f])
qed
lemma integrable_on_superset:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes "f integrable_on S"
and "\<And>x. x \<notin> S \<Longrightarrow> f x = 0"
and "S \<subseteq> t"
shows "f integrable_on t"
using assms
unfolding integrable_on_def
by (auto intro:has_integral_on_superset)
lemma integral_restrict_UNIV:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
shows "integral UNIV (\<lambda>x. if x \<in> S then f x else 0) = integral S f"
by (simp add: integral_restrict_Int)
lemma integrable_restrict_UNIV:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
shows "(\<lambda>x. if x \<in> s then f x else 0) integrable_on UNIV \<longleftrightarrow> f integrable_on s"
unfolding integrable_on_def
by auto
lemma has_integral_subset_component_le:
fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
assumes k: "k \<in> Basis"
and as: "S \<subseteq> T" "(f has_integral i) S" "(f has_integral j) T" "\<And>x. x\<in>T \<Longrightarrow> 0 \<le> f(x)\<bullet>k"
shows "i\<bullet>k \<le> j\<bullet>k"
proof -
have "((\<lambda>x. if x \<in> S then f x else 0) has_integral i) UNIV"
"((\<lambda>x. if x \<in> T then f x else 0) has_integral j) UNIV"
by (simp_all add: assms)
then show ?thesis
apply (rule has_integral_component_le[OF k])
using as by auto
qed
subsection\<open>Integrals on set differences\<close>
lemma has_integral_setdiff:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
assumes S: "(f has_integral i) S" and T: "(f has_integral j) T"
and neg: "negligible (T - S)"
shows "(f has_integral (i - j)) (S - T)"
proof -
show ?thesis
unfolding has_integral_restrict_UNIV [symmetric, of f]
proof (rule has_integral_spike [OF neg])
have eq: "(\<lambda>x. (if x \<in> S then f x else 0) - (if x \<in> T then f x else 0)) =
(\<lambda>x. if x \<in> T - S then - f x else if x \<in> S - T then f x else 0)"
by (force simp add: )
have "((\<lambda>x. if x \<in> S then f x else 0) has_integral i) UNIV"
"((\<lambda>x. if x \<in> T then f x else 0) has_integral j) UNIV"
using S T has_integral_restrict_UNIV by auto
from has_integral_diff [OF this]
show "((\<lambda>x. if x \<in> T - S then - f x else if x \<in> S - T then f x else 0)
has_integral i-j) UNIV"
by (simp add: eq)
qed force
qed
lemma integral_setdiff:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
assumes "f integrable_on S" "f integrable_on T" "negligible(T - S)"
shows "integral (S - T) f = integral S f - integral T f"
by (rule integral_unique) (simp add: assms has_integral_setdiff integrable_integral)
lemma integrable_setdiff:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
assumes "(f has_integral i) S" "(f has_integral j) T" "negligible (T - S)"
shows "f integrable_on (S - T)"
using has_integral_setdiff [OF assms]
by (simp add: has_integral_iff)
lemma negligible_setdiff [simp]: "T \<subseteq> S \<Longrightarrow> negligible (T - S)"
by (metis Diff_eq_empty_iff negligible_empty)
lemma negligible_on_intervals: "negligible s \<longleftrightarrow> (\<forall>a b. negligible(s \<inter> cbox a b))" (is "?l \<longleftrightarrow> ?r")
proof
assume ?r
show ?l
unfolding negligible_def
proof safe
fix a b
show "(indicator s has_integral 0) (cbox a b)"
apply (rule has_integral_negligible[OF \<open>?r\<close>[rule_format,of a b]])
unfolding indicator_def
apply auto
done
qed
qed (simp add: negligible_Int)
lemma negligible_translation:
assumes "negligible S"
shows "negligible ((+) c ` S)"
proof -
have inj: "inj ((+) c)"
by simp
show ?thesis
using assms
proof (clarsimp simp: negligible_def)
fix a b
assume "\<forall>x y. (indicator S has_integral 0) (cbox x y)"
then have *: "(indicator S has_integral 0) (cbox (a-c) (b-c))"
by (meson Diff_iff assms has_integral_negligible indicator_simps(2))
have eq: "indicator ((+) c ` S) = (\<lambda>x. indicator S (x - c))"
by (force simp add: indicator_def)
show "(indicator ((+) c ` S) has_integral 0) (cbox a b)"
using has_integral_affinity [OF *, of 1 "-c"]
cbox_translation [of "c" "-c+a" "-c+b"]
by (simp add: eq) (simp add: ac_simps)
qed
qed
lemma negligible_translation_rev:
assumes "negligible ((+) c ` S)"
shows "negligible S"
by (metis negligible_translation [OF assms, of "-c"] translation_galois)
lemma has_integral_spike_set_eq:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes "negligible {x \<in> S - T. f x \<noteq> 0}" "negligible {x \<in> T - S. f x \<noteq> 0}"
shows "(f has_integral y) S \<longleftrightarrow> (f has_integral y) T"
proof -
have "((\<lambda>x. if x \<in> S then f x else 0) has_integral y) UNIV =
((\<lambda>x. if x \<in> T then f x else 0) has_integral y) UNIV"
proof (rule has_integral_spike_eq)
show "negligible ({x \<in> S - T. f x \<noteq> 0} \<union> {x \<in> T - S. f x \<noteq> 0})"
by (rule negligible_Un [OF assms])
qed auto
then show ?thesis
by (simp add: has_integral_restrict_UNIV)
qed
corollary integral_spike_set:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes "negligible {x \<in> S - T. f x \<noteq> 0}" "negligible {x \<in> T - S. f x \<noteq> 0}"
shows "integral S f = integral T f"
using has_integral_spike_set_eq [OF assms]
by (metis eq_integralD integral_unique)
lemma integrable_spike_set:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes f: "f integrable_on S" and neg: "negligible {x \<in> S - T. f x \<noteq> 0}" "negligible {x \<in> T - S. f x \<noteq> 0}"
shows "f integrable_on T"
using has_integral_spike_set_eq [OF neg] f by blast
lemma integrable_spike_set_eq:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes "negligible ((S - T) \<union> (T - S))"
shows "f integrable_on S \<longleftrightarrow> f integrable_on T"
by (blast intro: integrable_spike_set assms negligible_subset)
lemma integrable_on_insert_iff: "f integrable_on (insert x X) \<longleftrightarrow> f integrable_on X"
for f::"_ \<Rightarrow> 'a::banach"
by (rule integrable_spike_set_eq) (auto simp: insert_Diff_if)
lemma has_integral_interior:
fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: banach"
shows "negligible(frontier S) \<Longrightarrow> (f has_integral y) (interior S) \<longleftrightarrow> (f has_integral y) S"
by (rule has_integral_spike_set_eq [OF empty_imp_negligible negligible_subset])
(use interior_subset in \<open>auto simp: frontier_def closure_def\<close>)
lemma has_integral_closure:
fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: banach"
shows "negligible(frontier S) \<Longrightarrow> (f has_integral y) (closure S) \<longleftrightarrow> (f has_integral y) S"
by (rule has_integral_spike_set_eq [OF negligible_subset empty_imp_negligible]) (auto simp: closure_Un_frontier )
lemma has_integral_open_interval:
fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: banach"
shows "(f has_integral y) (box a b) \<longleftrightarrow> (f has_integral y) (cbox a b)"
unfolding interior_cbox [symmetric]
by (metis frontier_cbox has_integral_interior negligible_frontier_interval)
lemma integrable_on_open_interval:
fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: banach"
shows "f integrable_on box a b \<longleftrightarrow> f integrable_on cbox a b"
by (simp add: has_integral_open_interval integrable_on_def)
lemma integral_open_interval:
fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: banach"
shows "integral(box a b) f = integral(cbox a b) f"
by (metis has_integral_integrable_integral has_integral_open_interval not_integrable_integral)
subsection \<open>More lemmas that are useful later\<close>
lemma has_integral_subset_le:
fixes f :: "'n::euclidean_space \<Rightarrow> real"
assumes "s \<subseteq> t"
and "(f has_integral i) s"
and "(f has_integral j) t"
and "\<forall>x\<in>t. 0 \<le> f x"
shows "i \<le> j"
using has_integral_subset_component_le[OF _ assms(1), of 1 f i j]
using assms
by auto
lemma integral_subset_component_le:
fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
assumes "k \<in> Basis"
and "s \<subseteq> t"
and "f integrable_on s"
and "f integrable_on t"
and "\<forall>x \<in> t. 0 \<le> f x \<bullet> k"
shows "(integral s f)\<bullet>k \<le> (integral t f)\<bullet>k"
apply (rule has_integral_subset_component_le)
using assms
apply auto
done
lemma integral_subset_le:
fixes f :: "'n::euclidean_space \<Rightarrow> real"
assumes "s \<subseteq> t"
and "f integrable_on s"
and "f integrable_on t"
and "\<forall>x \<in> t. 0 \<le> f x"
shows "integral s f \<le> integral t f"
apply (rule has_integral_subset_le)
using assms
apply auto
done
lemma has_integral_alt':
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
shows "(f has_integral i) s \<longleftrightarrow>
(\<forall>a b. (\<lambda>x. if x \<in> s then f x else 0) integrable_on cbox a b) \<and>
(\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0) - i) < e)"
(is "?l = ?r")
proof
assume rhs: ?r
show ?l
proof (subst has_integral', intro allI impI)
fix e::real
assume "e > 0"
from rhs[THEN conjunct2,rule_format,OF this]
show "\<exists>B>0. \<forall>a b. ball 0 B \<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 - i) < e)"
apply (rule ex_forward)
using rhs by blast
qed
next
let ?\<Phi> = "\<lambda>e a b. \<exists>z. ((\<lambda>x. if x \<in> s then f x else 0) has_integral z) (cbox a b) \<and> norm (z - i) < e"
assume ?l
then have lhs: "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> ?\<Phi> e a b" if "e > 0" for e
using that has_integral'[of f] by auto
let ?f = "\<lambda>x. if x \<in> s then f x else 0"
show ?r
proof (intro conjI allI impI)
fix a b :: 'n
from lhs[OF zero_less_one]
obtain B where "0 < B" and B: "\<And>a b. ball 0 B \<subseteq> cbox a b \<Longrightarrow> ?\<Phi> 1 a b"
by blast
let ?a = "\<Sum>i\<in>Basis. min (a\<bullet>i) (-B) *\<^sub>R i::'n"
let ?b = "\<Sum>i\<in>Basis. max (b\<bullet>i) B *\<^sub>R i::'n"
show "?f integrable_on cbox a b"
proof (rule integrable_subinterval[of _ ?a ?b])
have "?a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> ?b \<bullet> i" if "norm (0 - x) < B" "i \<in> Basis" for x i
using Basis_le_norm[of i x] that by (auto simp add:field_simps)
then have "ball 0 B \<subseteq> cbox ?a ?b"
by (auto simp: mem_box dist_norm)
then show "?f integrable_on cbox ?a ?b"
unfolding integrable_on_def using B by blast
show "cbox a b \<subseteq> cbox ?a ?b"
by (force simp: mem_box)
qed
fix e :: real
assume "e > 0"
with lhs show "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0) - i) < e"
by (metis (no_types, lifting) has_integral_integrable_integral)
qed
qed
subsection \<open>Continuity of the integral (for a 1-dimensional interval)\<close>
lemma integrable_alt:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
shows "f integrable_on s \<longleftrightarrow>
(\<forall>a b. (\<lambda>x. if x \<in> s then f x else 0) integrable_on cbox a b) \<and>
(\<forall>e>0. \<exists>B>0. \<forall>a b c d. ball 0 B \<subseteq> cbox a b \<and> ball 0 B \<subseteq> cbox c d \<longrightarrow>
norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0) -
integral (cbox c d) (\<lambda>x. if x \<in> s then f x else 0)) < e)"
(is "?l = ?r")
proof
let ?F = "\<lambda>x. if x \<in> s then f x else 0"
assume ?l
then obtain y where intF: "\<And>a b. ?F integrable_on cbox a b"
and y: "\<And>e. 0 < e \<Longrightarrow>
\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> norm (integral (cbox a b) ?F - y) < e"
unfolding integrable_on_def has_integral_alt'[of f] by auto
show ?r
proof (intro conjI allI impI intF)
fix e::real
assume "e > 0"
then have "e/2 > 0"
by auto
obtain B where "0 < B"
and B: "\<And>a b. ball 0 B \<subseteq> cbox a b \<Longrightarrow> norm (integral (cbox a b) ?F - y) < e/2"
using \<open>0 < e/2\<close> y by blast
show "\<exists>B>0. \<forall>a b c d. ball 0 B \<subseteq> cbox a b \<and> ball 0 B \<subseteq> cbox c d \<longrightarrow>
norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e"
proof (intro conjI exI impI allI, rule \<open>0 < B\<close>)
fix a b c d::'n
assume sub: "ball 0 B \<subseteq> cbox a b \<and> ball 0 B \<subseteq> cbox c d"
show "norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e"
using sub by (auto intro: norm_triangle_half_l dest: B)
qed
qed
next
let ?F = "\<lambda>x. if x \<in> s then f x else 0"
assume rhs: ?r
let ?cube = "\<lambda>n. cbox (\<Sum>i\<in>Basis. - real n *\<^sub>R i::'n) (\<Sum>i\<in>Basis. real n *\<^sub>R i)"
have "Cauchy (\<lambda>n. integral (?cube n) ?F)"
unfolding Cauchy_def
proof (intro allI impI)
fix e::real
assume "e > 0"
with rhs obtain B where "0 < B"
and B: "\<And>a b c d. ball 0 B \<subseteq> cbox a b \<and> ball 0 B \<subseteq> cbox c d
\<Longrightarrow> norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e"
by blast
obtain N where N: "B \<le> real N"
using real_arch_simple by blast
have "ball 0 B \<subseteq> ?cube n" if n: "n \<ge> N" for n
proof -
have "sum ((*\<^sub>R) (- real n)) Basis \<bullet> i \<le> x \<bullet> i \<and>
x \<bullet> i \<le> sum ((*\<^sub>R) (real n)) Basis \<bullet> i"
if "norm x < B" "i \<in> Basis" for x i::'n
using Basis_le_norm[of i x] n N that by (auto simp add: field_simps sum_negf)
then show ?thesis
by (auto simp: mem_box dist_norm)
qed
then show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (integral (?cube m) ?F) (integral (?cube n) ?F) < e"
by (fastforce simp add: dist_norm intro!: B)
qed
then obtain i where i: "(\<lambda>n. integral (?cube n) ?F) \<longlonglongrightarrow> i"
using convergent_eq_Cauchy by blast
have "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> norm (integral (cbox a b) ?F - i) < e"
if "e > 0" for e
proof -
have *: "e/2 > 0" using that by auto
then obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> norm (i - integral (?cube n) ?F) < e/2"
using i[THEN LIMSEQ_D, simplified norm_minus_commute] by meson
obtain B where "0 < B"
and B: "\<And>a b c d. \<lbrakk>ball 0 B \<subseteq> cbox a b; ball 0 B \<subseteq> cbox c d\<rbrakk> \<Longrightarrow>
norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e/2"
using rhs * by meson
let ?B = "max (real N) B"
show ?thesis
proof (intro exI conjI allI impI)
show "0 < ?B"
using \<open>B > 0\<close> by auto
fix a b :: 'n
assume "ball 0 ?B \<subseteq> cbox a b"
moreover obtain n where n: "max (real N) B \<le> real n"
using real_arch_simple by blast
moreover have "ball 0 B \<subseteq> ?cube n"
proof
fix x :: 'n
assume x: "x \<in> ball 0 B"
have "\<lbrakk>norm (0 - x) < B; i \<in> Basis\<rbrakk>
\<Longrightarrow> sum ((*\<^sub>R) (-n)) Basis \<bullet> i\<le> x \<bullet> i \<and> x \<bullet> i \<le> sum ((*\<^sub>R) n) Basis \<bullet> i" for i
using Basis_le_norm[of i x] n by (auto simp add: field_simps sum_negf)
then show "x \<in> ?cube n"
using x by (auto simp: mem_box dist_norm)
qed
ultimately show "norm (integral (cbox a b) ?F - i) < e"
using norm_triangle_half_l [OF B N] by force
qed
qed
then show ?l unfolding integrable_on_def has_integral_alt'[of f]
using rhs by blast
qed
lemma integrable_altD:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes "f integrable_on s"
shows "\<And>a b. (\<lambda>x. if x \<in> s then f x else 0) integrable_on cbox a b"
and "\<And>e. e > 0 \<Longrightarrow> \<exists>B>0. \<forall>a b c d. ball 0 B \<subseteq> cbox a b \<and> ball 0 B \<subseteq> cbox c d \<longrightarrow>
norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0) - integral (cbox c d) (\<lambda>x. if x \<in> s then f x else 0)) < e"
using assms[unfolded integrable_alt[of f]] by auto
lemma integrable_alt_subset:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
shows
"f integrable_on S \<longleftrightarrow>
(\<forall>a b. (\<lambda>x. if x \<in> S then f x else 0) integrable_on cbox a b) \<and>
(\<forall>e>0. \<exists>B>0. \<forall>a b c d.
ball 0 B \<subseteq> cbox a b \<and> cbox a b \<subseteq> cbox c d
\<longrightarrow> norm(integral (cbox a b) (\<lambda>x. if x \<in> S then f x else 0) -
integral (cbox c d) (\<lambda>x. if x \<in> S then f x else 0)) < e)"
(is "_ = ?rhs")
proof -
let ?g = "\<lambda>x. if x \<in> S then f x else 0"
have "f integrable_on S \<longleftrightarrow>
(\<forall>a b. ?g integrable_on cbox a b) \<and>
(\<forall>e>0. \<exists>B>0. \<forall>a b c d. ball 0 B \<subseteq> cbox a b \<and> ball 0 B \<subseteq> cbox c d \<longrightarrow>
norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e)"
by (rule integrable_alt)
also have "\<dots> = ?rhs"
proof -
{ fix e :: "real"
assume e: "\<And>e. e>0 \<Longrightarrow> \<exists>B>0. \<forall>a b c d. ball 0 B \<subseteq> cbox a b \<and> cbox a b \<subseteq> cbox c d \<longrightarrow>
norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e"
and "e > 0"
obtain B where "B > 0"
and B: "\<And>a b c d. \<lbrakk>ball 0 B \<subseteq> cbox a b; cbox a b \<subseteq> cbox c d\<rbrakk> \<Longrightarrow>
norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e/2"
using \<open>e > 0\<close> e [of "e/2"] by force
have "\<exists>B>0. \<forall>a b c d.
ball 0 B \<subseteq> cbox a b \<and> ball 0 B \<subseteq> cbox c d \<longrightarrow>
norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e"
proof (intro exI allI conjI impI)
fix a b c d :: "'a"
let ?\<alpha> = "\<Sum>i\<in>Basis. max (a \<bullet> i) (c \<bullet> i) *\<^sub>R i"
let ?\<beta> = "\<Sum>i\<in>Basis. min (b \<bullet> i) (d \<bullet> i) *\<^sub>R i"
show "norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e"
if ball: "ball 0 B \<subseteq> cbox a b \<and> ball 0 B \<subseteq> cbox c d"
proof -
have B': "norm (integral (cbox a b \<inter> cbox c d) ?g - integral (cbox x y) ?g) < e/2"
if "cbox a b \<inter> cbox c d \<subseteq> cbox x y" for x y
using B [of ?\<alpha> ?\<beta> x y] ball that by (simp add: Int_interval [symmetric])
show ?thesis
using B' [of a b] B' [of c d] norm_triangle_half_r by blast
qed
qed (use \<open>B > 0\<close> in auto)}
then show ?thesis
by force
qed
finally show ?thesis .
qed
lemma integrable_on_subcbox:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes intf: "f integrable_on S"
and sub: "cbox a b \<subseteq> S"
shows "f integrable_on cbox a b"
proof -
have "(\<lambda>x. if x \<in> S then f x else 0) integrable_on cbox a b"
by (simp add: intf integrable_altD(1))
then show ?thesis
by (metis (mono_tags) sub integrable_restrict_Int le_inf_iff order_refl subset_antisym)
qed
subsection \<open>A straddling criterion for integrability\<close>
lemma integrable_straddle_interval:
fixes f :: "'n::euclidean_space \<Rightarrow> real"
assumes "\<And>e. e>0 \<Longrightarrow> \<exists>g h i j. (g has_integral i) (cbox a b) \<and> (h has_integral j) (cbox a b) \<and>
\<bar>i - j\<bar> < e \<and> (\<forall>x\<in>cbox a b. (g x) \<le> f x \<and> f x \<le> h x)"
shows "f integrable_on cbox a b"
proof -
have "\<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>
\<bar>(\<Sum>(x,K)\<in>p1. content K *\<^sub>R f x) - (\<Sum>(x,K)\<in>p2. content K *\<^sub>R f x)\<bar> < e)"
if "e > 0" for e
proof -
have e: "e/3 > 0"
using that by auto
then obtain g h i j where ij: "\<bar>i - j\<bar> < e/3"
and "(g has_integral i) (cbox a b)"
and "(h has_integral j) (cbox a b)"
and fgh: "\<And>x. x \<in> cbox a b \<Longrightarrow> g x \<le> f x \<and> f x \<le> h x"
using assms real_norm_def by metis
then obtain d1 d2 where "gauge d1" "gauge d2"
and d1: "\<And>p. \<lbrakk>p tagged_division_of cbox a b; d1 fine p\<rbrakk> \<Longrightarrow>
\<bar>(\<Sum>(x,K)\<in>p. content K *\<^sub>R g x) - i\<bar> < e/3"
and d2: "\<And>p. \<lbrakk>p tagged_division_of cbox a b; d2 fine p\<rbrakk> \<Longrightarrow>
\<bar>(\<Sum>(x,K) \<in> p. content K *\<^sub>R h x) - j\<bar> < e/3"
by (metis e has_integral real_norm_def)
have "\<bar>(\<Sum>(x,K) \<in> p1. content K *\<^sub>R f x) - (\<Sum>(x,K) \<in> p2. content K *\<^sub>R f x)\<bar> < e"
if p1: "p1 tagged_division_of cbox a b" and 11: "d1 fine p1" and 21: "d2 fine p1"
and p2: "p2 tagged_division_of cbox a b" and 12: "d1 fine p2" and 22: "d2 fine p2" for p1 p2
proof -
have *: "\<And>g1 g2 h1 h2 f1 f2.
\<lbrakk>\<bar>g2 - i\<bar> < e/3; \<bar>g1 - i\<bar> < e/3; \<bar>h2 - j\<bar> < e/3; \<bar>h1 - j\<bar> < e/3;
g1 - h2 \<le> f1 - f2; f1 - f2 \<le> h1 - g2\<rbrakk>
\<Longrightarrow> \<bar>f1 - f2\<bar> < e"
using \<open>e > 0\<close> ij by arith
have 0: "(\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p1. content k *\<^sub>R g x) \<ge> 0"
"0 \<le> (\<Sum>(x, k)\<in>p2. content k *\<^sub>R h x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)"
"(\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R g x) \<ge> 0"
"0 \<le> (\<Sum>(x, k)\<in>p1. content k *\<^sub>R h x) - (\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x)"
unfolding sum_subtractf[symmetric]
apply (auto intro!: sum_nonneg)
apply (meson fgh measure_nonneg mult_left_mono tag_in_interval that sum_nonneg)+
done
show ?thesis
proof (rule *)
show "\<bar>(\<Sum>(x,K) \<in> p2. content K *\<^sub>R g x) - i\<bar> < e/3"
by (rule d1[OF p2 12])
show "\<bar>(\<Sum>(x,K) \<in> p1. content K *\<^sub>R g x) - i\<bar> < e/3"
by (rule d1[OF p1 11])
show "\<bar>(\<Sum>(x,K) \<in> p2. content K *\<^sub>R h x) - j\<bar> < e/3"
by (rule d2[OF p2 22])
show "\<bar>(\<Sum>(x,K) \<in> p1. content K *\<^sub>R h x) - j\<bar> < e/3"
by (rule d2[OF p1 21])
qed (use 0 in auto)
qed
then show ?thesis
by (rule_tac x="\<lambda>x. d1 x \<inter> d2 x" in exI)
(auto simp: fine_Int intro: \<open>gauge d1\<close> \<open>gauge d2\<close> d1 d2)
qed
then show ?thesis
by (simp add: integrable_Cauchy)
qed
lemma integrable_straddle:
fixes f :: "'n::euclidean_space \<Rightarrow> real"
assumes "\<And>e. e>0 \<Longrightarrow> \<exists>g h i j. (g has_integral i) s \<and> (h has_integral j) s \<and>
\<bar>i - j\<bar> < e \<and> (\<forall>x\<in>s. g x \<le> f x \<and> f x \<le> h x)"
shows "f integrable_on s"
proof -
let ?fs = "(\<lambda>x. if x \<in> s then f x else 0)"
have "?fs integrable_on cbox a b" for a b
proof (rule integrable_straddle_interval)
fix e::real
assume "e > 0"
then have *: "e/4 > 0"
by auto
with assms obtain g h i j where g: "(g has_integral i) s" and h: "(h has_integral j) s"
and ij: "\<bar>i - j\<bar> < e/4"
and fgh: "\<And>x. x \<in> s \<Longrightarrow> g x \<le> f x \<and> f x \<le> h x"
by metis
let ?gs = "(\<lambda>x. if x \<in> s then g x else 0)"
let ?hs = "(\<lambda>x. if x \<in> s then h x else 0)"
obtain Bg where Bg: "\<And>a b. ball 0 Bg \<subseteq> cbox a b \<Longrightarrow> \<bar>integral (cbox a b) ?gs - i\<bar> < e/4"
and int_g: "\<And>a b. ?gs integrable_on cbox a b"
using g * unfolding has_integral_alt' real_norm_def by meson
obtain Bh where
Bh: "\<And>a b. ball 0 Bh \<subseteq> cbox a b \<Longrightarrow> \<bar>integral (cbox a b) ?hs - j\<bar> < e/4"
and int_h: "\<And>a b. ?hs integrable_on cbox a b"
using h * unfolding has_integral_alt' real_norm_def by meson
define c where "c = (\<Sum>i\<in>Basis. min (a\<bullet>i) (- (max Bg Bh)) *\<^sub>R i)"
define d where "d = (\<Sum>i\<in>Basis. max (b\<bullet>i) (max Bg Bh) *\<^sub>R i)"
have "\<lbrakk>norm (0 - x) < Bg; i \<in> Basis\<rbrakk> \<Longrightarrow> c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i" for x i
using Basis_le_norm[of i x] unfolding c_def d_def by auto
then have ballBg: "ball 0 Bg \<subseteq> cbox c d"
by (auto simp: mem_box dist_norm)
have "\<lbrakk>norm (0 - x) < Bh; i \<in> Basis\<rbrakk> \<Longrightarrow> c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i" for x i
using Basis_le_norm[of i x] unfolding c_def d_def by auto
then have ballBh: "ball 0 Bh \<subseteq> cbox c d"
by (auto simp: mem_box dist_norm)
have ab_cd: "cbox a b \<subseteq> cbox c d"
by (auto simp: c_def d_def subset_box_imp)
have **: "\<And>ch cg ag ah::real. \<lbrakk>\<bar>ah - ag\<bar> \<le> \<bar>ch - cg\<bar>; \<bar>cg - i\<bar> < e/4; \<bar>ch - j\<bar> < e/4\<rbrakk>
\<Longrightarrow> \<bar>ag - ah\<bar> < e"
using ij by arith
show "\<exists>g h i j. (g has_integral i) (cbox a b) \<and> (h has_integral j) (cbox a b) \<and> \<bar>i - j\<bar> < e \<and>
(\<forall>x\<in>cbox a b. g x \<le> (if x \<in> s then f x else 0) \<and>
(if x \<in> s then f x else 0) \<le> h x)"
proof (intro exI ballI conjI)
have eq: "\<And>x f g. (if x \<in> s then f x else 0) - (if x \<in> s then g x else 0) =
(if x \<in> s then f x - g x else (0::real))"
by auto
have int_hg: "(\<lambda>x. if x \<in> s then h x - g x else 0) integrable_on cbox a b"
"(\<lambda>x. if x \<in> s then h x - g x else 0) integrable_on cbox c d"
by (metis (no_types) integrable_diff g h has_integral_integrable integrable_altD(1))+
show "(?gs has_integral integral (cbox a b) ?gs) (cbox a b)"
"(?hs has_integral integral (cbox a b) ?hs) (cbox a b)"
by (intro integrable_integral int_g int_h)+
then have "integral (cbox a b) ?gs \<le> integral (cbox a b) ?hs"
apply (rule has_integral_le)
using fgh by force
then have "0 \<le> integral (cbox a b) ?hs - integral (cbox a b) ?gs"
by simp
then have "\<bar>integral (cbox a b) ?hs - integral (cbox a b) ?gs\<bar>
\<le> \<bar>integral (cbox c d) ?hs - integral (cbox c d) ?gs\<bar>"
apply (simp add: integral_diff [symmetric] int_g int_h)
apply (subst abs_of_nonneg[OF integral_nonneg[OF integrable_diff, OF int_h int_g]])
using fgh apply (force simp: eq intro!: integral_subset_le [OF ab_cd int_hg])+
done
then show "\<bar>integral (cbox a b) ?gs - integral (cbox a b) ?hs\<bar> < e"
apply (rule **)
apply (rule Bg ballBg Bh ballBh)+
done
show "\<And>x. x \<in> cbox a b \<Longrightarrow> ?gs x \<le> ?fs x" "\<And>x. x \<in> cbox a b \<Longrightarrow> ?fs x \<le> ?hs x"
using fgh by auto
qed
qed
then have int_f: "?fs integrable_on cbox a b" for a b
by simp
have "\<exists>B>0. \<forall>a b c d.
ball 0 B \<subseteq> cbox a b \<and> ball 0 B \<subseteq> cbox c d \<longrightarrow>
abs (integral (cbox a b) ?fs - integral (cbox c d) ?fs) < e"
if "0 < e" for e
proof -
have *: "e/3 > 0"
using that by auto
with assms obtain g h i j where g: "(g has_integral i) s" and h: "(h has_integral j) s"
and ij: "\<bar>i - j\<bar> < e/3"
and fgh: "\<And>x. x \<in> s \<Longrightarrow> g x \<le> f x \<and> f x \<le> h x"
by metis
let ?gs = "(\<lambda>x. if x \<in> s then g x else 0)"
let ?hs = "(\<lambda>x. if x \<in> s then h x else 0)"
obtain Bg where "Bg > 0"
and Bg: "\<And>a b. ball 0 Bg \<subseteq> cbox a b \<Longrightarrow> \<bar>integral (cbox a b) ?gs - i\<bar> < e/3"
and int_g: "\<And>a b. ?gs integrable_on cbox a b"
using g * unfolding has_integral_alt' real_norm_def by meson
obtain Bh where "Bh > 0"
and Bh: "\<And>a b. ball 0 Bh \<subseteq> cbox a b \<Longrightarrow> \<bar>integral (cbox a b) ?hs - j\<bar> < e/3"
and int_h: "\<And>a b. ?hs integrable_on cbox a b"
using h * unfolding has_integral_alt' real_norm_def by meson
{ fix a b c d :: 'n
assume as: "ball 0 (max Bg Bh) \<subseteq> cbox a b" "ball 0 (max Bg Bh) \<subseteq> cbox c d"
have **: "ball 0 Bg \<subseteq> ball (0::'n) (max Bg Bh)" "ball 0 Bh \<subseteq> ball (0::'n) (max Bg Bh)"
by auto
have *: "\<And>ga gc ha hc fa fc. \<lbrakk>\<bar>ga - i\<bar> < e/3; \<bar>gc - i\<bar> < e/3; \<bar>ha - j\<bar> < e/3;
\<bar>hc - j\<bar> < e/3; ga \<le> fa; fa \<le> ha; gc \<le> fc; fc \<le> hc\<rbrakk> \<Longrightarrow>
\<bar>fa - fc\<bar> < e"
using ij by arith
have "abs (integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0) - integral (cbox c d)
(\<lambda>x. if x \<in> s then f x else 0)) < e"
proof (rule *)
show "\<bar>integral (cbox a b) ?gs - i\<bar> < e/3"
using "**" Bg as by blast
show "\<bar>integral (cbox c d) ?gs - i\<bar> < e/3"
using "**" Bg as by blast
show "\<bar>integral (cbox a b) ?hs - j\<bar> < e/3"
using "**" Bh as by blast
show "\<bar>integral (cbox c d) ?hs - j\<bar> < e/3"
using "**" Bh as by blast
qed (use int_f int_g int_h fgh in \<open>simp_all add: integral_le\<close>)
}
then show ?thesis
apply (rule_tac x="max Bg Bh" in exI)
using \<open>Bg > 0\<close> by auto
qed
then show ?thesis
unfolding integrable_alt[of f] real_norm_def by (blast intro: int_f)
qed
subsection \<open>Adding integrals over several sets\<close>
lemma has_integral_Un:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes f: "(f has_integral i) S" "(f has_integral j) T"
and neg: "negligible (S \<inter> T)"
shows "(f has_integral (i + j)) (S \<union> T)"
unfolding has_integral_restrict_UNIV[symmetric, of f]
proof (rule has_integral_spike[OF neg])
let ?f = "\<lambda>x. (if x \<in> S then f x else 0) + (if x \<in> T then f x else 0)"
show "(?f has_integral i + j) UNIV"
by (simp add: f has_integral_add)
qed auto
lemma integral_Un [simp]:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes "f integrable_on S" "f integrable_on T" "negligible (S \<inter> T)"
shows "integral (S \<union> T) f = integral S f + integral T f"
by (simp add: has_integral_Un assms integrable_integral integral_unique)
lemma integrable_Un:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b :: banach"
assumes "negligible (A \<inter> B)" "f integrable_on A" "f integrable_on B"
shows "f integrable_on (A \<union> B)"
proof -
from assms obtain y z where "(f has_integral y) A" "(f has_integral z) B"
by (auto simp: integrable_on_def)
from has_integral_Un[OF this assms(1)] show ?thesis by (auto simp: integrable_on_def)
qed
lemma integrable_Un':
fixes f :: "'a::euclidean_space \<Rightarrow> 'b :: banach"
assumes "f integrable_on A" "f integrable_on B" "negligible (A \<inter> B)" "C = A \<union> B"
shows "f integrable_on C"
using integrable_Un[of A B f] assms by simp
lemma has_integral_Union:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes \<T>: "finite \<T>"
and int: "\<And>S. S \<in> \<T> \<Longrightarrow> (f has_integral (i S)) S"
and neg: "pairwise (\<lambda>S S'. negligible (S \<inter> S')) \<T>"
shows "(f has_integral (sum i \<T>)) (\<Union>\<T>)"
proof -
let ?\<U> = "((\<lambda>(a,b). a \<inter> b) ` {(a,b). a \<in> \<T> \<and> b \<in> {y. y \<in> \<T> \<and> a \<noteq> y}})"
have "((\<lambda>x. if x \<in> \<Union>\<T> then f x else 0) has_integral sum i \<T>) UNIV"
proof (rule has_integral_spike)
show "negligible (\<Union>?\<U>)"
proof (rule negligible_Union)
have "finite (\<T> \<times> \<T>)"
by (simp add: \<T>)
moreover have "{(a, b). a \<in> \<T> \<and> b \<in> {y \<in> \<T>. a \<noteq> y}} \<subseteq> \<T> \<times> \<T>"
by auto
ultimately show "finite ?\<U>"
by (blast intro: finite_subset[of _ "\<T> \<times> \<T>"])
show "\<And>t. t \<in> ?\<U> \<Longrightarrow> negligible t"
using neg unfolding pairwise_def by auto
qed
next
show "(if x \<in> \<Union>\<T> then f x else 0) = (\<Sum>A\<in>\<T>. if x \<in> A then f x else 0)"
if "x \<in> UNIV - (\<Union>?\<U>)" for x
proof clarsimp
fix S assume "S \<in> \<T>" "x \<in> S"
moreover then have "\<forall>b\<in>\<T>. x \<in> b \<longleftrightarrow> b = S"
using that by blast
ultimately show "f x = (\<Sum>A\<in>\<T>. if x \<in> A then f x else 0)"
by (simp add: sum.delta[OF \<T>])
qed
next
show "((\<lambda>x. \<Sum>A\<in>\<T>. if x \<in> A then f x else 0) has_integral (\<Sum>A\<in>\<T>. i A)) UNIV"
apply (rule has_integral_sum [OF \<T>])
using int by (simp add: has_integral_restrict_UNIV)
qed
then show ?thesis
using has_integral_restrict_UNIV by blast
qed
text \<open>In particular adding integrals over a division, maybe not of an interval.\<close>
lemma has_integral_combine_division:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes "\<D> division_of S"
and "\<And>k. k \<in> \<D> \<Longrightarrow> (f has_integral (i k)) k"
shows "(f has_integral (sum i \<D>)) S"
proof -
note \<D> = division_ofD[OF assms(1)]
have neg: "negligible (S \<inter> s')" if "S \<in> \<D>" "s' \<in> \<D>" "S \<noteq> s'" for S s'
proof -
obtain a c b \<D> where obt: "S = cbox a b" "s' = cbox c \<D>"
by (meson \<open>S \<in> \<D>\<close> \<open>s' \<in> \<D>\<close> \<D>(4))
from \<D>(5)[OF that] show ?thesis
unfolding obt interior_cbox
by (metis (no_types, lifting) Diff_empty Int_interval box_Int_box negligible_frontier_interval)
qed
show ?thesis
unfolding \<D>(6)[symmetric]
by (auto intro: \<D> neg assms has_integral_Union pairwiseI)
qed
lemma integral_combine_division_bottomup:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes "\<D> division_of S" "\<And>k. k \<in> \<D> \<Longrightarrow> f integrable_on k"
shows "integral S f = sum (\<lambda>i. integral i f) \<D>"
apply (rule integral_unique)
by (meson assms has_integral_combine_division has_integral_integrable_integral)
lemma has_integral_combine_division_topdown:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes f: "f integrable_on S"
and \<D>: "\<D> division_of K"
and "K \<subseteq> S"
shows "(f has_integral (sum (\<lambda>i. integral i f) \<D>)) K"
proof -
have "f integrable_on L" if "L \<in> \<D>" for L
proof -
have "L \<subseteq> S"
using \<open>K \<subseteq> S\<close> \<D> that by blast
then show "f integrable_on L"
using that by (metis (no_types) f \<D> division_ofD(4) integrable_on_subcbox)
qed
then show ?thesis
by (meson \<D> has_integral_combine_division has_integral_integrable_integral)
qed
lemma integral_combine_division_topdown:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes "f integrable_on S"
and "\<D> division_of S"
shows "integral S f = sum (\<lambda>i. integral i f) \<D>"
apply (rule integral_unique)
apply (rule has_integral_combine_division_topdown)
using assms
apply auto
done
lemma integrable_combine_division:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes \<D>: "\<D> division_of S"
and f: "\<And>i. i \<in> \<D> \<Longrightarrow> f integrable_on i"
shows "f integrable_on S"
using f unfolding integrable_on_def by (metis has_integral_combine_division[OF \<D>])
lemma integrable_on_subdivision:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes \<D>: "\<D> division_of i"
and f: "f integrable_on S"
and "i \<subseteq> S"
shows "f integrable_on i"
proof -
have "f integrable_on i" if "i \<in> \<D>" for i
proof -
have "i \<subseteq> S"
using assms that by auto
then show "f integrable_on i"
using that by (metis (no_types) \<D> f division_ofD(4) integrable_on_subcbox)
qed
then show ?thesis
using \<D> integrable_combine_division by blast
qed
subsection \<open>Also tagged divisions\<close>
lemma has_integral_combine_tagged_division:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes "p tagged_division_of S"
and "\<forall>(x,k) \<in> p. (f has_integral (i k)) k"
shows "(f has_integral (\<Sum>(x,k)\<in>p. i k)) S"
proof -
have *: "(f has_integral (\<Sum>k\<in>snd`p. integral k f)) S"
using assms(2)
apply (intro has_integral_combine_division)
apply (auto simp: has_integral_integral[symmetric] intro: division_of_tagged_division[OF assms(1)])
apply auto
done
also have "(\<Sum>k\<in>snd`p. integral k f) = (\<Sum>(x, k)\<in>p. integral k f)"
by (intro sum.over_tagged_division_lemma[OF assms(1), symmetric] integral_null)
(simp add: content_eq_0_interior)
finally show ?thesis
using assms by (auto simp add: has_integral_iff intro!: sum.cong)
qed
lemma integral_combine_tagged_division_bottomup:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes "p tagged_division_of (cbox a b)"
and "\<forall>(x,k)\<in>p. f integrable_on k"
shows "integral (cbox a b) f = sum (\<lambda>(x,k). integral k f) p"
apply (rule integral_unique)
apply (rule has_integral_combine_tagged_division[OF assms(1)])
using assms(2)
apply auto
done
lemma has_integral_combine_tagged_division_topdown:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes f: "f integrable_on cbox a b"
and p: "p tagged_division_of (cbox a b)"
shows "(f has_integral (sum (\<lambda>(x,K). integral K f) p)) (cbox a b)"
proof -
have "(f has_integral integral K f) K" if "(x,K) \<in> p" for x K
by (metis assms integrable_integral integrable_on_subcbox tagged_division_ofD(3,4) that)
then show ?thesis
by (metis assms case_prodI2 has_integral_integrable_integral integral_combine_tagged_division_bottomup)
qed
lemma integral_combine_tagged_division_topdown:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes "f integrable_on cbox a b"
and "p tagged_division_of (cbox a b)"
shows "integral (cbox a b) f = sum (\<lambda>(x,k). integral k f) p"
apply (rule integral_unique [OF has_integral_combine_tagged_division_topdown])
using assms apply auto
done
subsection \<open>Henstock's lemma\<close>
lemma Henstock_lemma_part1:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes intf: "f integrable_on cbox a b"
and "e > 0"
and "gauge d"
and less_e: "\<And>p. \<lbrakk>p tagged_division_of (cbox a b); d fine p\<rbrakk> \<Longrightarrow>
norm (sum (\<lambda>(x,K). content K *\<^sub>R f x) p - integral(cbox a b) f) < e"
and p: "p tagged_partial_division_of (cbox a b)" "d fine p"
shows "norm (sum (\<lambda>(x,K). content K *\<^sub>R f x - integral K f) p) \<le> e" (is "?lhs \<le> e")
proof (rule field_le_epsilon)
fix k :: real
assume "k > 0"
let ?SUM = "\<lambda>p. (\<Sum>(x,K) \<in> p. content K *\<^sub>R f x)"
note p' = tagged_partial_division_ofD[OF p(1)]
have "\<Union>(snd ` p) \<subseteq> cbox a b"
using p'(3) by fastforce
then obtain q where q: "snd ` p \<subseteq> q" and qdiv: "q division_of cbox a b"
by (meson p(1) partial_division_extend_interval partial_division_of_tagged_division)
note q' = division_ofD[OF qdiv]
define r where "r = q - snd ` p"
have "snd ` p \<inter> r = {}"
unfolding r_def by auto
have "finite r"
using q' unfolding r_def by auto
have "\<exists>p. p tagged_division_of i \<and> d fine p \<and>
norm (?SUM p - integral i f) < k / (real (card r) + 1)"
if "i\<in>r" for i
proof -
have gt0: "k / (real (card r) + 1) > 0" using \<open>k > 0\<close> by simp
have i: "i \<in> q"
using that unfolding r_def by auto
then obtain u v where uv: "i = cbox u v"
using q'(4) by blast
then have "cbox u v \<subseteq> cbox a b"
using i q'(2) by auto
then have "f integrable_on cbox u v"
by (rule integrable_subinterval[OF intf])
with integrable_integral[OF this, unfolded has_integral[of f]]
obtain dd where "gauge dd" and dd:
"\<And>\<D>. \<lbrakk>\<D> tagged_division_of cbox u v; dd fine \<D>\<rbrakk> \<Longrightarrow>
norm (?SUM \<D> - integral (cbox u v) f) < k / (real (card r) + 1)"
using gt0 by auto
with gauge_Int[OF \<open>gauge d\<close> \<open>gauge dd\<close>]
obtain qq where qq: "qq tagged_division_of cbox u v" "(\<lambda>x. d x \<inter> dd x) fine qq"
using fine_division_exists by blast
with dd[of qq] show ?thesis
by (auto simp: fine_Int uv)
qed
then obtain qq where qq: "\<And>i. i \<in> r \<Longrightarrow> qq i tagged_division_of i \<and>
d fine qq i \<and> norm (?SUM (qq i) - integral i f) < k / (real (card r) + 1)"
by metis
let ?p = "p \<union> \<Union>(qq ` r)"
have "norm (?SUM ?p - integral (cbox a b) f) < e"
proof (rule less_e)
show "d fine ?p"
by (metis (mono_tags, hide_lams) qq fine_Un fine_Union imageE p(2))
note ptag = tagged_partial_division_of_Union_self[OF p(1)]
have "p \<union> \<Union>(qq ` r) tagged_division_of \<Union>(snd ` p) \<union> \<Union>r"
proof (rule tagged_division_Un[OF ptag tagged_division_Union [OF \<open>finite r\<close>]])
show "\<And>i. i \<in> r \<Longrightarrow> qq i tagged_division_of i"
using qq by auto
show "\<And>i1 i2. \<lbrakk>i1 \<in> r; i2 \<in> r; i1 \<noteq> i2\<rbrakk> \<Longrightarrow> interior i1 \<inter> interior i2 = {}"
by (simp add: q'(5) r_def)
show "interior (\<Union>(snd ` p)) \<inter> interior (\<Union>r) = {}"
proof (rule Int_interior_Union_intervals [OF \<open>finite r\<close>])
show "open (interior (\<Union>(snd ` p)))"
by blast
show "\<And>T. T \<in> r \<Longrightarrow> \<exists>a b. T = cbox a b"
by (simp add: q'(4) r_def)
have "finite (snd ` p)"
by (simp add: p'(1))
then show "\<And>T. T \<in> r \<Longrightarrow> interior (\<Union>(snd ` p)) \<inter> interior T = {}"
apply (subst Int_commute)
apply (rule Int_interior_Union_intervals)
using r_def q'(5) q(1) apply auto
by (simp add: p'(4))
qed
qed
moreover have "\<Union>(snd ` p) \<union> \<Union>r = cbox a b" and "{qq i |i. i \<in> r} = qq ` r"
using qdiv q unfolding Union_Un_distrib[symmetric] r_def by auto
ultimately show "?p tagged_division_of (cbox a b)"
by fastforce
qed
then have "norm (?SUM p + (?SUM (\<Union>(qq ` r))) - integral (cbox a b) f) < e"
proof (subst sum.union_inter_neutral[symmetric, OF \<open>finite p\<close>], safe)
show "content L *\<^sub>R f x = 0" if "(x, L) \<in> p" "(x, L) \<in> qq K" "K \<in> r" for x K L
proof -
obtain u v where uv: "L = cbox u v"
using \<open>(x,L) \<in> p\<close> p'(4) by blast
have "L \<subseteq> K"
using qq[OF that(3)] tagged_division_ofD(3) \<open>(x,L) \<in> qq K\<close> by metis
have "L \<in> snd ` p"
using \<open>(x,L) \<in> p\<close> image_iff by fastforce
then have "L \<in> q" "K \<in> q" "L \<noteq> K"
using that(1,3) q(1) unfolding r_def by auto
with q'(5) have "interior L = {}"
using interior_mono[OF \<open>L \<subseteq> K\<close>] by blast
then show "content L *\<^sub>R f x = 0"
unfolding uv content_eq_0_interior[symmetric] by auto
qed
show "finite (\<Union>(qq ` r))"
by (meson finite_UN qq \<open>finite r\<close> tagged_division_of_finite)
qed
moreover have "content M *\<^sub>R f x = 0"
if x: "(x,M) \<in> qq K" "(x,M) \<in> qq L" and KL: "qq K \<noteq> qq L" and r: "K \<in> r" "L \<in> r"
for x M K L
proof -
note kl = tagged_division_ofD(3,4)[OF qq[THEN conjunct1]]
obtain u v where uv: "M = cbox u v"
using \<open>(x, M) \<in> qq L\<close> \<open>L \<in> r\<close> kl(2) by blast
have empty: "interior (K \<inter> L) = {}"
by (metis DiffD1 interior_Int q'(5) r_def KL r)
have "interior M = {}"
by (metis (no_types, lifting) Int_assoc empty inf.absorb_iff2 interior_Int kl(1) subset_empty x r)
then show "content M *\<^sub>R f x = 0"
unfolding uv content_eq_0_interior[symmetric]
by auto
qed
ultimately have "norm (?SUM p + sum ?SUM (qq ` r) - integral (cbox a b) f) < e"
apply (subst (asm) sum.Union_comp)
using qq by (force simp: split_paired_all)+
moreover have "content M *\<^sub>R f x = 0"
if "K \<in> r" "L \<in> r" "K \<noteq> L" "qq K = qq L" "(x, M) \<in> qq K" for K L x M
using tagged_division_ofD(6) qq that by (metis (no_types, lifting))
ultimately have less_e: "norm (?SUM p + sum (?SUM \<circ> qq) r - integral (cbox a b) f) < e"
apply (subst (asm) sum.reindex_nontrivial [OF \<open>finite r\<close>])
apply (auto simp: split_paired_all sum.neutral)
done
have norm_le: "norm (cp - ip) \<le> e + k"
if "norm ((cp + cr) - i) < e" "norm (cr - ir) < k" "ip + ir = i"
for ir ip i cr cp::'a
proof -
from that show ?thesis
using norm_triangle_le[of "cp + cr - i" "- (cr - ir)"]
unfolding that(3)[symmetric] norm_minus_cancel
by (auto simp add: algebra_simps)
qed
have "?lhs = norm (?SUM p - (\<Sum>(x, k)\<in>p. integral k f))"
unfolding split_def sum_subtractf ..
also have "\<dots> \<le> e + k"
proof (rule norm_le[OF less_e])
have lessk: "k * real (card r) / (1 + real (card r)) < k"
using \<open>k>0\<close> by (auto simp add: field_simps)
have "norm (sum (?SUM \<circ> qq) r - (\<Sum>k\<in>r. integral k f)) \<le> (\<Sum>x\<in>r. k / (real (card r) + 1))"
unfolding sum_subtractf[symmetric] by (force dest: qq intro!: sum_norm_le)
also have "... < k"
by (simp add: lessk add.commute mult.commute)
finally show "norm (sum (?SUM \<circ> qq) r - (\<Sum>k\<in>r. integral k f)) < k" .
next
from q(1) have [simp]: "snd ` p \<union> q = q" by auto
have "integral l f = 0"
if inp: "(x, l) \<in> p" "(y, m) \<in> p" and ne: "(x, l) \<noteq> (y, m)" and "l = m" for x l y m
proof -
obtain u v where uv: "l = cbox u v"
using inp p'(4) by blast
have "content (cbox u v) = 0"
unfolding content_eq_0_interior using that p(1) uv by auto
then show ?thesis
using uv by blast
qed
then have "(\<Sum>(x, k)\<in>p. integral k f) = (\<Sum>k\<in>snd ` p. integral k f)"
apply (subst sum.reindex_nontrivial [OF \<open>finite p\<close>])
unfolding split_paired_all split_def by auto
then show "(\<Sum>(x, k)\<in>p. integral k f) + (\<Sum>k\<in>r. integral k f) = integral (cbox a b) f"
unfolding integral_combine_division_topdown[OF intf qdiv] r_def
using q'(1) p'(1) sum.union_disjoint [of "snd ` p" "q - snd ` p", symmetric]
by simp
qed
finally show "?lhs \<le> e + k" .
qed
lemma Henstock_lemma_part2:
fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
assumes fed: "f integrable_on cbox a b" "e > 0" "gauge d"
and less_e: "\<And>\<D>. \<lbrakk>\<D> tagged_division_of (cbox a b); d fine \<D>\<rbrakk> \<Longrightarrow>
norm (sum (\<lambda>(x,k). content k *\<^sub>R f x) \<D> - integral (cbox a b) f) < e"
and tag: "p tagged_partial_division_of (cbox a b)"
and "d fine p"
shows "sum (\<lambda>(x,k). norm (content k *\<^sub>R f x - integral k f)) p \<le> 2 * real (DIM('n)) * e"
proof -
have "finite p"
using tag by blast
then show ?thesis
unfolding split_def
proof (rule sum_norm_allsubsets_bound)
fix Q
assume Q: "Q \<subseteq> p"
then have fine: "d fine Q"
by (simp add: \<open>d fine p\<close> fine_subset)
show "norm (\<Sum>x\<in>Q. content (snd x) *\<^sub>R f (fst x) - integral (snd x) f) \<le> e"
apply (rule Henstock_lemma_part1[OF fed less_e, unfolded split_def])
using Q tag tagged_partial_division_subset apply (force simp add: fine)+
done
qed
qed
lemma Henstock_lemma:
fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
assumes intf: "f integrable_on cbox a b"
and "e > 0"
obtains \<gamma> where "gauge \<gamma>"
and "\<And>p. \<lbrakk>p tagged_partial_division_of (cbox a b); \<gamma> fine p\<rbrakk> \<Longrightarrow>
sum (\<lambda>(x,k). norm(content k *\<^sub>R f x - integral k f)) p < e"
proof -
have *: "e/(2 * (real DIM('n) + 1)) > 0" using \<open>e > 0\<close> by simp
with integrable_integral[OF intf, unfolded has_integral]
obtain \<gamma> where "gauge \<gamma>"
and \<gamma>: "\<And>\<D>. \<lbrakk>\<D> tagged_division_of cbox a b; \<gamma> fine \<D>\<rbrakk> \<Longrightarrow>
norm ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f x) - integral (cbox a b) f)
< e/(2 * (real DIM('n) + 1))"
by metis
show thesis
proof (rule that [OF \<open>gauge \<gamma>\<close>])
fix p
assume p: "p tagged_partial_division_of cbox a b" "\<gamma> fine p"
have "(\<Sum>(x,K)\<in>p. norm (content K *\<^sub>R f x - integral K f))
\<le> 2 * real DIM('n) * (e/(2 * (real DIM('n) + 1)))"
using Henstock_lemma_part2[OF intf * \<open>gauge \<gamma>\<close> \<gamma> p] by metis
also have "... < e"
using \<open>e > 0\<close> by (auto simp add: field_simps)
finally
show "(\<Sum>(x,K)\<in>p. norm (content K *\<^sub>R f x - integral K f)) < e" .
qed
qed
subsection \<open>Monotone convergence (bounded interval first)\<close>
lemma bounded_increasing_convergent:
fixes f :: "nat \<Rightarrow> real"
shows "\<lbrakk>bounded (range f); \<And>n. f n \<le> f (Suc n)\<rbrakk> \<Longrightarrow> \<exists>l. f \<longlonglongrightarrow> l"
using Bseq_mono_convergent[of f] incseq_Suc_iff[of f]
by (auto simp: image_def Bseq_eq_bounded convergent_def incseq_def)
lemma monotone_convergence_interval:
fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> real"
assumes intf: "\<And>k. (f k) integrable_on cbox a b"
and le: "\<And>k x. x \<in> cbox a b \<Longrightarrow> (f k x) \<le> f (Suc k) x"
and fg: "\<And>x. x \<in> cbox a b \<Longrightarrow> ((\<lambda>k. f k x) \<longlongrightarrow> g x) sequentially"
and bou: "bounded (range (\<lambda>k. integral (cbox a b) (f k)))"
shows "g integrable_on cbox a b \<and> ((\<lambda>k. integral (cbox a b) (f k)) \<longlongrightarrow> integral (cbox a b) g) sequentially"
proof (cases "content (cbox a b) = 0")
case True then show ?thesis
by auto
next
case False
have fg1: "(f k x) \<le> (g x)" if x: "x \<in> cbox a b" for x k
proof -
have "\<forall>\<^sub>F j in sequentially. f k x \<le> f j x"
apply (rule eventually_sequentiallyI [of k])
using le x apply (force intro: transitive_stepwise_le)
done
then show "f k x \<le> g x"
using tendsto_lowerbound [OF fg] x trivial_limit_sequentially by blast
qed
have int_inc: "\<And>n. integral (cbox a b) (f n) \<le> integral (cbox a b) (f (Suc n))"
by (metis integral_le intf le)
then obtain i where i: "(\<lambda>k. integral (cbox a b) (f k)) \<longlonglongrightarrow> i"
using bounded_increasing_convergent bou by blast
have "\<And>k. \<forall>\<^sub>F x in sequentially. integral (cbox a b) (f k) \<le> integral (cbox a b) (f x)"
unfolding eventually_sequentially
by (force intro: transitive_stepwise_le int_inc)
then have i': "\<And>k. (integral(cbox a b) (f k)) \<le> i"
using tendsto_le [OF trivial_limit_sequentially i] by blast
have "(g has_integral i) (cbox a b)"
unfolding has_integral real_norm_def
proof clarify
fix e::real
assume e: "e > 0"
have "\<And>k. (\<exists>\<gamma>. gauge \<gamma> \<and> (\<forall>\<D>. \<D> tagged_division_of (cbox a b) \<and> \<gamma> fine \<D> \<longrightarrow>
abs ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f k x) - integral (cbox a b) (f k)) < e/2 ^ (k + 2)))"
using intf e by (auto simp: has_integral_integral has_integral)
then obtain c where c: "\<And>x. gauge (c x)"
"\<And>x \<D>. \<lbrakk>\<D> tagged_division_of cbox a b; c x fine \<D>\<rbrakk> \<Longrightarrow>
abs ((\<Sum>(u,K)\<in>\<D>. content K *\<^sub>R f x u) - integral (cbox a b) (f x))
< e/2 ^ (x + 2)"
by metis
have "\<exists>r. \<forall>k\<ge>r. 0 \<le> i - (integral (cbox a b) (f k)) \<and> i - (integral (cbox a b) (f k)) < e/4"
proof -
have "e/4 > 0"
using e by auto
show ?thesis
using LIMSEQ_D [OF i \<open>e/4 > 0\<close>] i' by auto
qed
then obtain r where r: "\<And>k. r \<le> k \<Longrightarrow> 0 \<le> i - integral (cbox a b) (f k)"
"\<And>k. r \<le> k \<Longrightarrow> i - integral (cbox a b) (f k) < e/4"
by metis
have "\<exists>n\<ge>r. \<forall>k\<ge>n. 0 \<le> (g x) - (f k x) \<and> (g x) - (f k x) < e/(4 * content(cbox a b))"
if "x \<in> cbox a b" for x
proof -
have "e/(4 * content (cbox a b)) > 0"
by (simp add: False content_lt_nz e)
with fg that LIMSEQ_D
obtain N where "\<forall>n\<ge>N. norm (f n x - g x) < e/(4 * content (cbox a b))"
by metis
then show "\<exists>n\<ge>r.
\<forall>k\<ge>n.
0 \<le> g x - f k x \<and>
g x - f k x
< e/(4 * content (cbox a b))"
apply (rule_tac x="N + r" in exI)
using fg1[OF that] apply (auto simp add: field_simps)
done
qed
then obtain m where r_le_m: "\<And>x. x \<in> cbox a b \<Longrightarrow> r \<le> m x"
and m: "\<And>x k. \<lbrakk>x \<in> cbox a b; m x \<le> k\<rbrakk>
\<Longrightarrow> 0 \<le> g x - f k x \<and> g x - f k x < e/(4 * content (cbox a b))"
by metis
define d where "d x = c (m x) x" for x
show "\<exists>\<gamma>. gauge \<gamma> \<and>
(\<forall>\<D>. \<D> tagged_division_of cbox a b \<and>
\<gamma> fine \<D> \<longrightarrow> abs ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R g x) - i) < e)"
proof (rule exI, safe)
show "gauge d"
using c(1) unfolding gauge_def d_def by auto
next
fix \<D>
assume ptag: "\<D> tagged_division_of (cbox a b)" and "d fine \<D>"
note p'=tagged_division_ofD[OF ptag]
obtain s where s: "\<And>x. x \<in> \<D> \<Longrightarrow> m (fst x) \<le> s"
by (metis finite_imageI finite_nat_set_iff_bounded_le p'(1) rev_image_eqI)
have *: "\<bar>a - d\<bar> < e" if "\<bar>a - b\<bar> \<le> e/4" "\<bar>b - c\<bar> < e/2" "\<bar>c - d\<bar> < e/4" for a b c d
using that norm_triangle_lt[of "a - b" "b - c" "3* e/4"]
norm_triangle_lt[of "a - b + (b - c)" "c - d" e]
by (auto simp add: algebra_simps)
show "\<bar>(\<Sum>(x, k)\<in>\<D>. content k *\<^sub>R g x) - i\<bar> < e"
proof (rule *)
have "\<bar>(\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R g x) - (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f (m x) x)\<bar>
\<le> (\<Sum>i\<in>\<D>. \<bar>(case i of (x, K) \<Rightarrow> content K *\<^sub>R g x) - (case i of (x, K) \<Rightarrow> content K *\<^sub>R f (m x) x)\<bar>)"
by (metis (mono_tags) sum_subtractf sum_abs)
also have "... \<le> (\<Sum>(x, k)\<in>\<D>. content k * (e/(4 * content (cbox a b))))"
proof (rule sum_mono, simp add: split_paired_all)
fix x K
assume xk: "(x,K) \<in> \<D>"
with ptag have x: "x \<in> cbox a b"
by blast
then have "abs (content K * (g x - f (m x) x)) \<le> content K * (e/(4 * content (cbox a b)))"
by (metis m[OF x] mult_nonneg_nonneg abs_of_nonneg less_eq_real_def measure_nonneg mult_left_mono order_refl)
then show "\<bar>content K * g x - content K * f (m x) x\<bar> \<le> content K * e/(4 * content (cbox a b))"
by (simp add: algebra_simps)
qed
also have "... = (e/(4 * content (cbox a b))) * (\<Sum>(x, k)\<in>\<D>. content k)"
by (simp add: sum_distrib_left sum_divide_distrib split_def mult.commute)
also have "... \<le> e/4"
by (metis False additive_content_tagged_division [OF ptag] nonzero_mult_divide_mult_cancel_right order_refl times_divide_eq_left)
finally show "\<bar>(\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R g x) - (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f (m x) x)\<bar> \<le> e/4" .
next
have "norm ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f (m x) x) - (\<Sum>(x,K)\<in>\<D>. integral K (f (m x))))
\<le> norm (\<Sum>j = 0..s. \<Sum>(x,K)\<in>{xk \<in> \<D>. m (fst xk) = j}. content K *\<^sub>R f (m x) x - integral K (f (m x)))"
apply (subst sum.group)
using s by (auto simp: sum_subtractf split_def p'(1))
also have "\<dots> < e/2"
proof -
have "norm (\<Sum>j = 0..s. \<Sum>(x, k)\<in>{xk \<in> \<D>. m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x)))
\<le> (\<Sum>i = 0..s. e/2 ^ (i + 2))"
proof (rule sum_norm_le)
fix t
assume "t \<in> {0..s}"
have "norm (\<Sum>(x,k)\<in>{xk \<in> \<D>. m (fst xk) = t}. content k *\<^sub>R f (m x) x - integral k (f (m x))) =
norm (\<Sum>(x,k)\<in>{xk \<in> \<D>. m (fst xk) = t}. content k *\<^sub>R f t x - integral k (f t))"
by (force intro!: sum.cong arg_cong[where f=norm])
also have "... \<le> e/2 ^ (t + 2)"
proof (rule Henstock_lemma_part1 [OF intf])
show "{xk \<in> \<D>. m (fst xk) = t} tagged_partial_division_of cbox a b"
apply (rule tagged_partial_division_subset[of \<D>])
using ptag by (auto simp: tagged_division_of_def)
show "c t fine {xk \<in> \<D>. m (fst xk) = t}"
using \<open>d fine \<D>\<close> by (auto simp: fine_def d_def)
qed (use c e in auto)
finally show "norm (\<Sum>(x,K)\<in>{xk \<in> \<D>. m (fst xk) = t}. content K *\<^sub>R f (m x) x -
integral K (f (m x))) \<le> e/2 ^ (t + 2)" .
qed
also have "... = (e/2/2) * (\<Sum>i = 0..s. (1/2) ^ i)"
by (simp add: sum_distrib_left field_simps)
also have "\<dots> < e/2"
by (simp add: sum_gp mult_strict_left_mono[OF _ e])
finally show "norm (\<Sum>j = 0..s. \<Sum>(x, k)\<in>{xk \<in> \<D>.
m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x))) < e/2" .
qed
finally show "\<bar>(\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f (m x) x) - (\<Sum>(x,K)\<in>\<D>. integral K (f (m x)))\<bar> < e/2"
by simp
next
have comb: "integral (cbox a b) (f y) = (\<Sum>(x, k)\<in>\<D>. integral k (f y))" for y
using integral_combine_tagged_division_topdown[OF intf ptag] by metis
have f_le: "\<And>y m n. \<lbrakk>y \<in> cbox a b; n\<ge>m\<rbrakk> \<Longrightarrow> f m y \<le> f n y"
using le by (auto intro: transitive_stepwise_le)
have "(\<Sum>(x, k)\<in>\<D>. integral k (f r)) \<le> (\<Sum>(x, K)\<in>\<D>. integral K (f (m x)))"
proof (rule sum_mono, simp add: split_paired_all)
fix x K
assume xK: "(x, K) \<in> \<D>"
show "integral K (f r) \<le> integral K (f (m x))"
proof (rule integral_le)
show "f r integrable_on K"
by (metis integrable_on_subcbox intf p'(3) p'(4) xK)
show "f (m x) integrable_on K"
by (metis elementary_interval integrable_on_subdivision intf p'(3) p'(4) xK)
show "f r y \<le> f (m x) y" if "y \<in> K" for y
using that r_le_m[of x] p'(2-3)[OF xK] f_le by auto
qed
qed
moreover have "(\<Sum>(x, K)\<in>\<D>. integral K (f (m x))) \<le> (\<Sum>(x, k)\<in>\<D>. integral k (f s))"
proof (rule sum_mono, simp add: split_paired_all)
fix x K
assume xK: "(x, K) \<in> \<D>"
show "integral K (f (m x)) \<le> integral K (f s)"
proof (rule integral_le)
show "f (m x) integrable_on K"
by (metis elementary_interval integrable_on_subdivision intf p'(3) p'(4) xK)
show "f s integrable_on K"
by (metis integrable_on_subcbox intf p'(3) p'(4) xK)
show "f (m x) y \<le> f s y" if "y \<in> K" for y
using that s xK f_le p'(3) by fastforce
qed
qed
moreover have "0 \<le> i - integral (cbox a b) (f r)" "i - integral (cbox a b) (f r) < e/4"
using r by auto
ultimately show "\<bar>(\<Sum>(x,K)\<in>\<D>. integral K (f (m x))) - i\<bar> < e/4"
using comb i'[of s] by auto
qed
qed
qed
with i integral_unique show ?thesis
by blast
qed
lemma monotone_convergence_increasing:
fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> real"
assumes int_f: "\<And>k. (f k) integrable_on S"
and "\<And>k x. x \<in> S \<Longrightarrow> (f k x) \<le> (f (Suc k) x)"
and fg: "\<And>x. x \<in> S \<Longrightarrow> ((\<lambda>k. f k x) \<longlongrightarrow> g x) sequentially"
and bou: "bounded (range (\<lambda>k. integral S (f k)))"
shows "g integrable_on S \<and> ((\<lambda>k. integral S (f k)) \<longlongrightarrow> integral S g) sequentially"
proof -
have lem: "g integrable_on S \<and> ((\<lambda>k. integral S (f k)) \<longlongrightarrow> integral S g) sequentially"
if f0: "\<And>k x. x \<in> S \<Longrightarrow> 0 \<le> f k x"
and int_f: "\<And>k. (f k) integrable_on S"
and le: "\<And>k x. x \<in> S \<Longrightarrow> f k x \<le> f (Suc k) x"
and lim: "\<And>x. x \<in> S \<Longrightarrow> ((\<lambda>k. f k x) \<longlongrightarrow> g x) sequentially"
and bou: "bounded (range(\<lambda>k. integral S (f k)))"
for f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> real" and g S
proof -
have fg: "(f k x) \<le> (g x)" if "x \<in> S" for x k
apply (rule tendsto_lowerbound [OF lim [OF that]])
apply (rule eventually_sequentiallyI [of k])
using le by (force intro: transitive_stepwise_le that)+
obtain i where i: "(\<lambda>k. integral S (f k)) \<longlonglongrightarrow> i"
using bounded_increasing_convergent [OF bou] le int_f integral_le by blast
have i': "(integral S (f k)) \<le> i" for k
proof -
have "\<And>k. \<And>x. x \<in> S \<Longrightarrow> \<forall>n\<ge>k. f k x \<le> f n x"
using le by (force intro: transitive_stepwise_le)
then show ?thesis
using tendsto_lowerbound [OF i eventually_sequentiallyI trivial_limit_sequentially]
by (meson int_f integral_le)
qed
let ?f = "(\<lambda>k x. if x \<in> S then f k x else 0)"
let ?g = "(\<lambda>x. if x \<in> S then g x else 0)"
have int: "?f k integrable_on cbox a b" for a b k
by (simp add: int_f integrable_altD(1))
have int': "\<And>k a b. f k integrable_on cbox a b \<inter> S"
using int by (simp add: Int_commute integrable_restrict_Int)
have g: "?g integrable_on cbox a b \<and>
(\<lambda>k. integral (cbox a b) (?f k)) \<longlonglongrightarrow> integral (cbox a b) ?g" for a b
proof (rule monotone_convergence_interval)
have "norm (integral (cbox a b) (?f k)) \<le> norm (integral S (f k))" for k
proof -
have "0 \<le> integral (cbox a b) (?f k)"
by (metis (no_types) integral_nonneg Int_iff f0 inf_commute integral_restrict_Int int')
moreover have "0 \<le> integral S (f k)"
by (simp add: integral_nonneg f0 int_f)
moreover have "integral (S \<inter> cbox a b) (f k) \<le> integral S (f k)"
by (metis f0 inf_commute int' int_f integral_subset_le le_inf_iff order_refl)
ultimately show ?thesis
by (simp add: integral_restrict_Int)
qed
moreover obtain B where "\<And>x. x \<in> range (\<lambda>k. integral S (f k)) \<Longrightarrow> norm x \<le> B"
using bou unfolding bounded_iff by blast
ultimately show "bounded (range (\<lambda>k. integral (cbox a b) (?f k)))"
unfolding bounded_iff by (blast intro: order_trans)
qed (use int le lim in auto)
moreover have "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> norm (integral (cbox a b) ?g - i) < e"
if "0 < e" for e
proof -
have "e/4>0"
using that by auto
with LIMSEQ_D [OF i] obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> norm (integral S (f n) - i) < e/4"
by metis
with int_f[of N, unfolded has_integral_integral has_integral_alt'[of "f N"]]
obtain B where "0 < B" and B:
"\<And>a b. ball 0 B \<subseteq> cbox a b \<Longrightarrow> norm (integral (cbox a b) (?f N) - integral S (f N)) < e/4"
by (meson \<open>0 < e/4\<close>)
have "norm (integral (cbox a b) ?g - i) < e" if ab: "ball 0 B \<subseteq> cbox a b" for a b
proof -
obtain M where M: "\<And>n. n \<ge> M \<Longrightarrow> abs (integral (cbox a b) (?f n) - integral (cbox a b) ?g) < e/2"
using \<open>e > 0\<close> g by (fastforce simp add: dest!: LIMSEQ_D [where r = "e/2"])
have *: "\<And>\<alpha> \<beta> g. \<lbrakk>\<bar>\<alpha> - i\<bar> < e/2; \<bar>\<beta> - g\<bar> < e/2; \<alpha> \<le> \<beta>; \<beta> \<le> i\<rbrakk> \<Longrightarrow> \<bar>g - i\<bar> < e"
unfolding real_inner_1_right by arith
show "norm (integral (cbox a b) ?g - i) < e"
unfolding real_norm_def
proof (rule *)
show "\<bar>integral (cbox a b) (?f N) - i\<bar> < e/2"
proof (rule abs_triangle_half_l)
show "\<bar>integral (cbox a b) (?f N) - integral S (f N)\<bar> < e/2/2"
using B[OF ab] by simp
show "abs (i - integral S (f N)) < e/2/2"
using N by (simp add: abs_minus_commute)
qed
show "\<bar>integral (cbox a b) (?f (M + N)) - integral (cbox a b) ?g\<bar> < e/2"
by (metis le_add1 M[of "M + N"])
show "integral (cbox a b) (?f N) \<le> integral (cbox a b) (?f (M + N))"
proof (intro ballI integral_le[OF int int])
fix x assume "x \<in> cbox a b"
have "(f m x) \<le> (f n x)" if "x \<in> S" "n \<ge> m" for m n
apply (rule transitive_stepwise_le [OF \<open>n \<ge> m\<close> order_refl])
using dual_order.trans apply blast
by (simp add: le \<open>x \<in> S\<close>)
then show "(?f N)x \<le> (?f (M+N))x"
by auto
qed
have "integral (cbox a b \<inter> S) (f (M + N)) \<le> integral S (f (M + N))"
by (metis Int_lower1 f0 inf_commute int' int_f integral_subset_le)
then have "integral (cbox a b) (?f (M + N)) \<le> integral S (f (M + N))"
by (metis (no_types) inf_commute integral_restrict_Int)
also have "... \<le> i"
using i'[of "M + N"] by auto
finally show "integral (cbox a b) (?f (M + N)) \<le> i" .
qed
qed
then show ?thesis
using \<open>0 < B\<close> by blast
qed
ultimately have "(g has_integral i) S"
unfolding has_integral_alt' by auto
then show ?thesis
using has_integral_integrable_integral i integral_unique by metis
qed
have sub: "\<And>k. integral S (\<lambda>x. f k x - f 0 x) = integral S (f k) - integral S (f 0)"
by (simp add: integral_diff int_f)
have *: "\<And>x m n. x \<in> S \<Longrightarrow> n\<ge>m \<Longrightarrow> f m x \<le> f n x"
using assms(2) by (force intro: transitive_stepwise_le)
have gf: "(\<lambda>x. g x - f 0 x) integrable_on S \<and> ((\<lambda>k. integral S (\<lambda>x. f (Suc k) x - f 0 x)) \<longlongrightarrow>
integral S (\<lambda>x. g x - f 0 x)) sequentially"
proof (rule lem)
show "\<And>k. (\<lambda>x. f (Suc k) x - f 0 x) integrable_on S"
by (simp add: integrable_diff int_f)
show "(\<lambda>k. f (Suc k) x - f 0 x) \<longlonglongrightarrow> g x - f 0 x" if "x \<in> S" for x
proof -
have "(\<lambda>n. f (Suc n) x) \<longlonglongrightarrow> g x"
using LIMSEQ_ignore_initial_segment[OF fg[OF \<open>x \<in> S\<close>], of 1] by simp
then show ?thesis
by (simp add: tendsto_diff)
qed
show "bounded (range (\<lambda>k. integral S (\<lambda>x. f (Suc k) x - f 0 x)))"
proof -
obtain B where B: "\<And>k. norm (integral S (f k)) \<le> B"
using bou by (auto simp: bounded_iff)
then have "norm (integral S (\<lambda>x. f (Suc k) x - f 0 x))
\<le> B + norm (integral S (f 0))" for k
unfolding sub by (meson add_le_cancel_right norm_triangle_le_diff)
then show ?thesis
unfolding bounded_iff by blast
qed
qed (use * in auto)
then have "(\<lambda>x. integral S (\<lambda>xa. f (Suc x) xa - f 0 xa) + integral S (f 0))
\<longlonglongrightarrow> integral S (\<lambda>x. g x - f 0 x) + integral S (f 0)"
by (auto simp add: tendsto_add)
moreover have "(\<lambda>x. g x - f 0 x + f 0 x) integrable_on S"
using gf integrable_add int_f [of 0] by metis
ultimately show ?thesis
by (simp add: integral_diff int_f LIMSEQ_imp_Suc sub)
qed
lemma has_integral_monotone_convergence_increasing:
fixes f :: "nat \<Rightarrow> 'a::euclidean_space \<Rightarrow> real"
assumes f: "\<And>k. (f k has_integral x k) s"
assumes "\<And>k x. x \<in> s \<Longrightarrow> f k x \<le> f (Suc k) x"
assumes "\<And>x. x \<in> s \<Longrightarrow> (\<lambda>k. f k x) \<longlonglongrightarrow> g x"
assumes "x \<longlonglongrightarrow> x'"
shows "(g has_integral x') s"
proof -
have x_eq: "x = (\<lambda>i. integral s (f i))"
by (simp add: integral_unique[OF f])
then have x: "range(\<lambda>k. integral s (f k)) = range x"
by auto
have *: "g integrable_on s \<and> (\<lambda>k. integral s (f k)) \<longlonglongrightarrow> integral s g"
proof (intro monotone_convergence_increasing allI ballI assms)
show "bounded (range(\<lambda>k. integral s (f k)))"
using x convergent_imp_bounded assms by metis
qed (use f in auto)
then have "integral s g = x'"
by (intro LIMSEQ_unique[OF _ \<open>x \<longlonglongrightarrow> x'\<close>]) (simp add: x_eq)
with * show ?thesis
by (simp add: has_integral_integral)
qed
lemma monotone_convergence_decreasing:
fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> real"
assumes intf: "\<And>k. (f k) integrable_on S"
and le: "\<And>k x. x \<in> S \<Longrightarrow> f (Suc k) x \<le> f k x"
and fg: "\<And>x. x \<in> S \<Longrightarrow> ((\<lambda>k. f k x) \<longlongrightarrow> g x) sequentially"
and bou: "bounded (range(\<lambda>k. integral S (f k)))"
shows "g integrable_on S \<and> (\<lambda>k. integral S (f k)) \<longlonglongrightarrow> integral S g"
proof -
have *: "range(\<lambda>k. integral S (\<lambda>x. - f k x)) = (*\<^sub>R) (- 1) ` (range(\<lambda>k. integral S (f k)))"
by force
have "(\<lambda>x. - g x) integrable_on S \<and> (\<lambda>k. integral S (\<lambda>x. - f k x)) \<longlonglongrightarrow> integral S (\<lambda>x. - g x)"
proof (rule monotone_convergence_increasing)
show "\<And>k. (\<lambda>x. - f k x) integrable_on S"
by (blast intro: integrable_neg intf)
show "\<And>k x. x \<in> S \<Longrightarrow> - f k x \<le> - f (Suc k) x"
by (simp add: le)
show "\<And>x. x \<in> S \<Longrightarrow> (\<lambda>k. - f k x) \<longlonglongrightarrow> - g x"
by (simp add: fg tendsto_minus)
show "bounded (range(\<lambda>k. integral S (\<lambda>x. - f k x)))"
using "*" bou bounded_scaling by auto
qed
then show ?thesis
by (force dest: integrable_neg tendsto_minus)
qed
lemma integral_norm_bound_integral:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes int_f: "f integrable_on S"
and int_g: "g integrable_on S"
and le_g: "\<And>x. x \<in> S \<Longrightarrow> norm (f x) \<le> g x"
shows "norm (integral S f) \<le> integral S g"
proof -
have norm: "norm \<eta> \<le> y + e"
if "norm \<zeta> \<le> x" and "\<bar>x - y\<bar> < e/2" and "norm (\<zeta> - \<eta>) < e/2"
for e x y and \<zeta> \<eta> :: 'a
proof -
have "norm (\<eta> - \<zeta>) < e/2"
by (metis norm_minus_commute that(3))
moreover have "x \<le> y + e/2"
using that(2) by linarith
ultimately show ?thesis
using that(1) le_less_trans[OF norm_triangle_sub[of \<eta> \<zeta>]] by (auto simp: less_imp_le)
qed
have lem: "norm (integral(cbox a b) f) \<le> integral (cbox a b) g"
if f: "f integrable_on cbox a b"
and g: "g integrable_on cbox a b"
and nle: "\<And>x. x \<in> cbox a b \<Longrightarrow> norm (f x) \<le> g x"
for f :: "'n \<Rightarrow> 'a" and g a b
proof (rule field_le_epsilon)
fix e :: real
assume "e > 0"
then have e: "e/2 > 0"
by auto
with integrable_integral[OF f,unfolded has_integral[of f]]
obtain \<gamma> where \<gamma>: "gauge \<gamma>"
"\<And>\<D>. \<D> tagged_division_of cbox a b \<and> \<gamma> fine \<D>
\<Longrightarrow> norm ((\<Sum>(x, k)\<in>\<D>. content k *\<^sub>R f x) - integral (cbox a b) f) < e/2"
by meson
moreover
from integrable_integral[OF g,unfolded has_integral[of g]] e
obtain \<delta> where \<delta>: "gauge \<delta>"
"\<And>\<D>. \<D> tagged_division_of cbox a b \<and> \<delta> fine \<D>
\<Longrightarrow> norm ((\<Sum>(x, k)\<in>\<D>. content k *\<^sub>R g x) - integral (cbox a b) g) < e/2"
by meson
ultimately have "gauge (\<lambda>x. \<gamma> x \<inter> \<delta> x)"
using gauge_Int by blast
with fine_division_exists obtain \<D>
where p: "\<D> tagged_division_of cbox a b" "(\<lambda>x. \<gamma> x \<inter> \<delta> x) fine \<D>"
by metis
have "\<gamma> fine \<D>" "\<delta> fine \<D>"
using fine_Int p(2) by blast+
show "norm (integral (cbox a b) f) \<le> integral (cbox a b) g + e"
proof (rule norm)
have "norm (content K *\<^sub>R f x) \<le> content K *\<^sub>R g x" if "(x, K) \<in> \<D>" for x K
proof-
have K: "x \<in> K" "K \<subseteq> cbox a b"
using \<open>(x, K) \<in> \<D>\<close> p(1) by blast+
obtain u v where "K = cbox u v"
using \<open>(x, K) \<in> \<D>\<close> p(1) by blast
moreover have "content K * norm (f x) \<le> content K * g x"
by (metis K subsetD dual_order.antisym measure_nonneg mult_zero_left nle not_le real_mult_le_cancel_iff2)
then show ?thesis
by simp
qed
then show "norm (\<Sum>(x, k)\<in>\<D>. content k *\<^sub>R f x) \<le> (\<Sum>(x, k)\<in>\<D>. content k *\<^sub>R g x)"
by (simp add: sum_norm_le split_def)
show "norm ((\<Sum>(x, k)\<in>\<D>. content k *\<^sub>R f x) - integral (cbox a b) f) < e/2"
using \<open>\<gamma> fine \<D>\<close> \<gamma> p(1) by simp
show "\<bar>(\<Sum>(x, k)\<in>\<D>. content k *\<^sub>R g x) - integral (cbox a b) g\<bar> < e/2"
using \<open>\<delta> fine \<D>\<close> \<delta> p(1) by simp
qed
qed
show ?thesis
proof (rule field_le_epsilon)
fix e :: real
assume "e > 0"
then have e: "e/2 > 0"
by auto
let ?f = "(\<lambda>x. if x \<in> S then f x else 0)"
let ?g = "(\<lambda>x. if x \<in> S then g x else 0)"
have f: "?f integrable_on cbox a b" and g: "?g integrable_on cbox a b" for a b
using int_f int_g integrable_altD by auto
obtain Bf where "0 < Bf"
and Bf: "\<And>a b. ball 0 Bf \<subseteq> cbox a b \<Longrightarrow>
\<exists>z. (?f has_integral z) (cbox a b) \<and> norm (z - integral S f) < e/2"
using integrable_integral [OF int_f,unfolded has_integral'[of f]] e that by blast
obtain Bg where "0 < Bg"
and Bg: "\<And>a b. ball 0 Bg \<subseteq> cbox a b \<Longrightarrow>
\<exists>z. (?g has_integral z) (cbox a b) \<and> norm (z - integral S g) < e/2"
using integrable_integral [OF int_g,unfolded has_integral'[of g]] e that by blast
obtain a b::'n where ab: "ball 0 Bf \<union> ball 0 Bg \<subseteq> cbox a b"
using ball_max_Un by (metis bounded_ball bounded_subset_cbox_symmetric)
have "ball 0 Bf \<subseteq> cbox a b"
using ab by auto
with Bf obtain z where int_fz: "(?f has_integral z) (cbox a b)" and z: "norm (z - integral S f) < e/2"
by meson
have "ball 0 Bg \<subseteq> cbox a b"
using ab by auto
with Bg obtain w where int_gw: "(?g has_integral w) (cbox a b)" and w: "norm (w - integral S g) < e/2"
by meson
show "norm (integral S f) \<le> integral S g + e"
proof (rule norm)
show "norm (integral (cbox a b) ?f) \<le> integral (cbox a b) ?g"
by (simp add: le_g lem[OF f g, of a b])
show "\<bar>integral (cbox a b) ?g - integral S g\<bar> < e/2"
using int_gw integral_unique w by auto
show "norm (integral (cbox a b) ?f - integral S f) < e/2"
using int_fz integral_unique z by blast
qed
qed
qed
lemma continuous_on_imp_absolutely_integrable_on:
fixes f::"real \<Rightarrow> 'a::banach"
shows "continuous_on {a..b} f \<Longrightarrow>
norm (integral {a..b} f) \<le> integral {a..b} (\<lambda>x. norm (f x))"
by (intro integral_norm_bound_integral integrable_continuous_real continuous_on_norm) auto
lemma integral_bound:
fixes f::"real \<Rightarrow> 'a::banach"
assumes "a \<le> b"
assumes "continuous_on {a .. b} f"
assumes "\<And>t. t \<in> {a .. b} \<Longrightarrow> norm (f t) \<le> B"
shows "norm (integral {a .. b} f) \<le> B * (b - a)"
proof -
note continuous_on_imp_absolutely_integrable_on[OF assms(2)]
also have "integral {a..b} (\<lambda>x. norm (f x)) \<le> integral {a..b} (\<lambda>_. B)"
by (rule integral_le)
(auto intro!: integrable_continuous_real continuous_intros assms)
also have "\<dots> = B * (b - a)" using assms by simp
finally show ?thesis .
qed
lemma integral_norm_bound_integral_component:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
fixes g :: "'n \<Rightarrow> 'b::euclidean_space"
assumes f: "f integrable_on S" and g: "g integrable_on S"
and fg: "\<And>x. x \<in> S \<Longrightarrow> norm(f x) \<le> (g x)\<bullet>k"
shows "norm (integral S f) \<le> (integral S g)\<bullet>k"
proof -
have "norm (integral S f) \<le> integral S ((\<lambda>x. x \<bullet> k) \<circ> g)"
apply (rule integral_norm_bound_integral[OF f integrable_linear[OF g]])
apply (simp add: bounded_linear_inner_left)
apply (metis fg o_def)
done
then show ?thesis
unfolding o_def integral_component_eq[OF g] .
qed
lemma has_integral_norm_bound_integral_component:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
fixes g :: "'n \<Rightarrow> 'b::euclidean_space"
assumes f: "(f has_integral i) S"
and g: "(g has_integral j) S"
and "\<And>x. x \<in> S \<Longrightarrow> norm (f x) \<le> (g x)\<bullet>k"
shows "norm i \<le> j\<bullet>k"
using integral_norm_bound_integral_component[of f S g k]
unfolding integral_unique[OF f] integral_unique[OF g]
using assms
by auto
lemma uniformly_convergent_improper_integral:
fixes f :: "'b \<Rightarrow> real \<Rightarrow> 'a :: {banach}"
assumes deriv: "\<And>x. x \<ge> a \<Longrightarrow> (G has_field_derivative g x) (at x within {a..})"
assumes integrable: "\<And>a' b x. x \<in> A \<Longrightarrow> a' \<ge> a \<Longrightarrow> b \<ge> a' \<Longrightarrow> f x integrable_on {a'..b}"
assumes G: "convergent G"
assumes le: "\<And>y x. y \<in> A \<Longrightarrow> x \<ge> a \<Longrightarrow> norm (f y x) \<le> g x"
shows "uniformly_convergent_on A (\<lambda>b x. integral {a..b} (f x))"
proof (intro Cauchy_uniformly_convergent uniformly_Cauchy_onI', goal_cases)
case (1 \<epsilon>)
from G have "Cauchy G"
by (auto intro!: convergent_Cauchy)
with 1 obtain M where M: "dist (G (real m)) (G (real n)) < \<epsilon>" if "m \<ge> M" "n \<ge> M" for m n
by (force simp: Cauchy_def)
define M' where "M' = max (nat \<lceil>a\<rceil>) M"
show ?case
proof (rule exI[of _ M'], safe, goal_cases)
case (1 x m n)
have M': "M' \<ge> a" "M' \<ge> M" unfolding M'_def by linarith+
have int_g: "(g has_integral (G (real n) - G (real m))) {real m..real n}"
using 1 M' by (intro fundamental_theorem_of_calculus)
(auto simp: has_field_derivative_iff_has_vector_derivative [symmetric]
intro!: DERIV_subset[OF deriv])
have int_f: "f x integrable_on {a'..real n}" if "a' \<ge> a" for a'
using that 1 by (cases "a' \<le> real n") (auto intro: integrable)
have "dist (integral {a..real m} (f x)) (integral {a..real n} (f x)) =
norm (integral {a..real n} (f x) - integral {a..real m} (f x))"
by (simp add: dist_norm norm_minus_commute)
also have "integral {a..real m} (f x) + integral {real m..real n} (f x) =
integral {a..real n} (f x)"
using M' and 1 by (intro integral_combine int_f) auto
hence "integral {a..real n} (f x) - integral {a..real m} (f x) =
integral {real m..real n} (f x)"
by (simp add: algebra_simps)
also have "norm \<dots> \<le> integral {real m..real n} g"
using le 1 M' int_f int_g by (intro integral_norm_bound_integral) auto
also from int_g have "integral {real m..real n} g = G (real n) - G (real m)"
by (simp add: has_integral_iff)
also have "\<dots> \<le> dist (G m) (G n)"
by (simp add: dist_norm)
also from 1 and M' have "\<dots> < \<epsilon>"
by (intro M) auto
finally show ?case .
qed
qed
lemma uniformly_convergent_improper_integral':
fixes f :: "'b \<Rightarrow> real \<Rightarrow> 'a :: {banach, real_normed_algebra}"
assumes deriv: "\<And>x. x \<ge> a \<Longrightarrow> (G has_field_derivative g x) (at x within {a..})"
assumes integrable: "\<And>a' b x. x \<in> A \<Longrightarrow> a' \<ge> a \<Longrightarrow> b \<ge> a' \<Longrightarrow> f x integrable_on {a'..b}"
assumes G: "convergent G"
assumes le: "eventually (\<lambda>x. \<forall>y\<in>A. norm (f y x) \<le> g x) at_top"
shows "uniformly_convergent_on A (\<lambda>b x. integral {a..b} (f x))"
proof -
from le obtain a'' where le: "\<And>y x. y \<in> A \<Longrightarrow> x \<ge> a'' \<Longrightarrow> norm (f y x) \<le> g x"
by (auto simp: eventually_at_top_linorder)
define a' where "a' = max a a''"
have "uniformly_convergent_on A (\<lambda>b x. integral {a'..real b} (f x))"
proof (rule uniformly_convergent_improper_integral)
fix t assume t: "t \<ge> a'"
hence "(G has_field_derivative g t) (at t within {a..})"
by (intro deriv) (auto simp: a'_def)
moreover have "{a'..} \<subseteq> {a..}" unfolding a'_def by auto
ultimately show "(G has_field_derivative g t) (at t within {a'..})"
by (rule DERIV_subset)
qed (insert le, auto simp: a'_def intro: integrable G)
hence "uniformly_convergent_on A (\<lambda>b x. integral {a..a'} (f x) + integral {a'..real b} (f x))"
(is ?P) by (intro uniformly_convergent_add) auto
also have "eventually (\<lambda>x. \<forall>y\<in>A. integral {a..a'} (f y) + integral {a'..x} (f y) =
integral {a..x} (f y)) sequentially"
by (intro eventually_mono [OF eventually_ge_at_top[of "nat \<lceil>a'\<rceil>"]] ballI integral_combine)
(auto simp: a'_def intro: integrable)
hence "?P \<longleftrightarrow> ?thesis"
by (intro uniformly_convergent_cong) simp_all
finally show ?thesis .
qed
subsection \<open>differentiation under the integral sign\<close>
lemma integral_continuous_on_param:
fixes f::"'a::topological_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::banach"
assumes cont_fx: "continuous_on (U \<times> cbox a b) (\<lambda>(x, t). f x t)"
shows "continuous_on U (\<lambda>x. integral (cbox a b) (f x))"
proof cases
assume "content (cbox a b) \<noteq> 0"
then have ne: "cbox a b \<noteq> {}" by auto
note [continuous_intros] =
continuous_on_compose2[OF cont_fx, where f="\<lambda>y. Pair x y" for x,
unfolded split_beta fst_conv snd_conv]
show ?thesis
unfolding continuous_on_def
proof (safe intro!: tendstoI)
fix e'::real and x
assume "e' > 0"
define e where "e = e' / (content (cbox a b) + 1)"
have "e > 0" using \<open>e' > 0\<close> by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos)
assume "x \<in> U"
from continuous_on_prod_compactE[OF cont_fx compact_cbox \<open>x \<in> U\<close> \<open>0 < e\<close>]
obtain X0 where X0: "x \<in> X0" "open X0"
and fx_bound: "\<And>y t. y \<in> X0 \<inter> U \<Longrightarrow> t \<in> cbox a b \<Longrightarrow> norm (f y t - f x t) \<le> e"
unfolding split_beta fst_conv snd_conv dist_norm
by metis
have "\<forall>\<^sub>F y in at x within U. y \<in> X0 \<inter> U"
using X0(1) X0(2) eventually_at_topological by auto
then show "\<forall>\<^sub>F y in at x within U. dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'"
proof eventually_elim
case (elim y)
have "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) =
norm (integral (cbox a b) (\<lambda>t. f y t - f x t))"
using elim \<open>x \<in> U\<close>
unfolding dist_norm
by (subst integral_diff)
(auto intro!: integrable_continuous continuous_intros)
also have "\<dots> \<le> e * content (cbox a b)"
using elim \<open>x \<in> U\<close>
by (intro integrable_bound)
(auto intro!: fx_bound \<open>x \<in> U \<close> less_imp_le[OF \<open>0 < e\<close>]
integrable_continuous continuous_intros)
also have "\<dots> < e'"
using \<open>0 < e'\<close> \<open>e > 0\<close>
by (auto simp: e_def field_split_simps)
finally show "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'" .
qed
qed
qed (auto intro!: continuous_on_const)
lemma leibniz_rule:
fixes f::"'a::banach \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::banach"
assumes fx: "\<And>x t. x \<in> U \<Longrightarrow> t \<in> cbox a b \<Longrightarrow>
((\<lambda>x. f x t) has_derivative blinfun_apply (fx x t)) (at x within U)"
assumes integrable_f2: "\<And>x. x \<in> U \<Longrightarrow> f x integrable_on cbox a b"
assumes cont_fx: "continuous_on (U \<times> (cbox a b)) (\<lambda>(x, t). fx x t)"
assumes [intro]: "x0 \<in> U"
assumes "convex U"
shows
"((\<lambda>x. integral (cbox a b) (f x)) has_derivative integral (cbox a b) (fx x0)) (at x0 within U)"
(is "(?F has_derivative ?dF) _")
proof cases
assume "content (cbox a b) \<noteq> 0"
then have ne: "cbox a b \<noteq> {}" by auto
note [continuous_intros] =
continuous_on_compose2[OF cont_fx, where f="\<lambda>y. Pair x y" for x,
unfolded split_beta fst_conv snd_conv]
show ?thesis
proof (intro has_derivativeI bounded_linear_scaleR_left tendstoI, fold norm_conv_dist)
have cont_f1: "\<And>t. t \<in> cbox a b \<Longrightarrow> continuous_on U (\<lambda>x. f x t)"
by (auto simp: continuous_on_eq_continuous_within intro!: has_derivative_continuous fx)
note [continuous_intros] = continuous_on_compose2[OF cont_f1]
fix e'::real
assume "e' > 0"
define e where "e = e' / (content (cbox a b) + 1)"
have "e > 0" using \<open>e' > 0\<close> by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos)
from continuous_on_prod_compactE[OF cont_fx compact_cbox \<open>x0 \<in> U\<close> \<open>e > 0\<close>]
obtain X0 where X0: "x0 \<in> X0" "open X0"
and fx_bound: "\<And>x t. x \<in> X0 \<inter> U \<Longrightarrow> t \<in> cbox a b \<Longrightarrow> norm (fx x t - fx x0 t) \<le> e"
unfolding split_beta fst_conv snd_conv
by (metis dist_norm)
note eventually_closed_segment[OF \<open>open X0\<close> \<open>x0 \<in> X0\<close>, of U]
moreover
have "\<forall>\<^sub>F x in at x0 within U. x \<in> X0"
using \<open>open X0\<close> \<open>x0 \<in> X0\<close> eventually_at_topological by blast
moreover have "\<forall>\<^sub>F x in at x0 within U. x \<noteq> x0"
by (auto simp: eventually_at_filter)
moreover have "\<forall>\<^sub>F x in at x0 within U. x \<in> U"
by (auto simp: eventually_at_filter)
ultimately
show "\<forall>\<^sub>F x in at x0 within U. norm ((?F x - ?F x0 - ?dF (x - x0)) /\<^sub>R norm (x - x0)) < e'"
proof eventually_elim
case (elim x)
from elim have "0 < norm (x - x0)" by simp
have "closed_segment x0 x \<subseteq> U"
by (rule \<open>convex U\<close>[unfolded convex_contains_segment, rule_format, OF \<open>x0 \<in> U\<close> \<open>x \<in> U\<close>])
from elim have [intro]: "x \<in> U" by auto
have "?F x - ?F x0 - ?dF (x - x0) =
integral (cbox a b) (\<lambda>y. f x y - f x0 y - fx x0 y (x - x0))"
(is "_ = ?id")
using \<open>x \<noteq> x0\<close>
by (subst blinfun_apply_integral integral_diff,
auto intro!: integrable_diff integrable_f2 continuous_intros
intro: integrable_continuous)+
also
{
fix t assume t: "t \<in> (cbox a b)"
have seg: "\<And>t. t \<in> {0..1} \<Longrightarrow> x0 + t *\<^sub>R (x - x0) \<in> X0 \<inter> U"
using \<open>closed_segment x0 x \<subseteq> U\<close>
\<open>closed_segment x0 x \<subseteq> X0\<close>
by (force simp: closed_segment_def algebra_simps)
from t have deriv:
"((\<lambda>x. f x t) has_derivative (fx y t)) (at y within X0 \<inter> U)"
if "y \<in> X0 \<inter> U" for y
unfolding has_vector_derivative_def[symmetric]
using that \<open>x \<in> X0\<close>
by (intro has_derivative_within_subset[OF fx]) auto
have "\<And>x. x \<in> X0 \<inter> U \<Longrightarrow> onorm (blinfun_apply (fx x t) - (fx x0 t)) \<le> e"
using fx_bound t
by (auto simp add: norm_blinfun_def fun_diff_def blinfun.bilinear_simps[symmetric])
from differentiable_bound_linearization[OF seg deriv this] X0
have "norm (f x t - f x0 t - fx x0 t (x - x0)) \<le> e * norm (x - x0)"
by (auto simp add: ac_simps)
}
then have "norm ?id \<le> integral (cbox a b) (\<lambda>_. e * norm (x - x0))"
by (intro integral_norm_bound_integral)
(auto intro!: continuous_intros integrable_diff integrable_f2
intro: integrable_continuous)
also have "\<dots> = content (cbox a b) * e * norm (x - x0)"
by simp
also have "\<dots> < e' * norm (x - x0)"
using \<open>e' > 0\<close>
apply (intro mult_strict_right_mono[OF _ \<open>0 < norm (x - x0)\<close>])
apply (simp add: divide_simps e_def not_less)
done
finally have "norm (?F x - ?F x0 - ?dF (x - x0)) < e' * norm (x - x0)" .
then show ?case
by (auto simp: divide_simps)
qed
qed (rule blinfun.bounded_linear_right)
qed (auto intro!: derivative_eq_intros simp: blinfun.bilinear_simps)
lemma has_vector_derivative_eq_has_derivative_blinfun:
"(f has_vector_derivative f') (at x within U) \<longleftrightarrow>
(f has_derivative blinfun_scaleR_left f') (at x within U)"
by (simp add: has_vector_derivative_def)
lemma leibniz_rule_vector_derivative:
fixes f::"real \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::banach"
assumes fx: "\<And>x t. x \<in> U \<Longrightarrow> t \<in> cbox a b \<Longrightarrow>
((\<lambda>x. f x t) has_vector_derivative (fx x t)) (at x within U)"
assumes integrable_f2: "\<And>x. x \<in> U \<Longrightarrow> (f x) integrable_on cbox a b"
assumes cont_fx: "continuous_on (U \<times> cbox a b) (\<lambda>(x, t). fx x t)"
assumes U: "x0 \<in> U" "convex U"
shows "((\<lambda>x. integral (cbox a b) (f x)) has_vector_derivative integral (cbox a b) (fx x0))
(at x0 within U)"
proof -
note [continuous_intros] =
continuous_on_compose2[OF cont_fx, where f="\<lambda>y. Pair x y" for x,
unfolded split_beta fst_conv snd_conv]
have *: "blinfun_scaleR_left (integral (cbox a b) (fx x0)) =
integral (cbox a b) (\<lambda>t. blinfun_scaleR_left (fx x0 t))"
by (subst integral_linear[symmetric])
(auto simp: has_vector_derivative_def o_def
intro!: integrable_continuous U continuous_intros bounded_linear_intros)
show ?thesis
unfolding has_vector_derivative_eq_has_derivative_blinfun
apply (rule has_derivative_eq_rhs)
apply (rule leibniz_rule[OF _ integrable_f2 _ U, where fx="\<lambda>x t. blinfun_scaleR_left (fx x t)"])
using fx cont_fx
apply (auto simp: has_vector_derivative_def * split_beta intro!: continuous_intros)
done
qed
lemma has_field_derivative_eq_has_derivative_blinfun:
"(f has_field_derivative f') (at x within U) \<longleftrightarrow> (f has_derivative blinfun_mult_right f') (at x within U)"
by (simp add: has_field_derivative_def)
lemma leibniz_rule_field_derivative:
fixes f::"'a::{real_normed_field, banach} \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'a"
assumes fx: "\<And>x t. x \<in> U \<Longrightarrow> t \<in> cbox a b \<Longrightarrow> ((\<lambda>x. f x t) has_field_derivative fx x t) (at x within U)"
assumes integrable_f2: "\<And>x. x \<in> U \<Longrightarrow> (f x) integrable_on cbox a b"
assumes cont_fx: "continuous_on (U \<times> (cbox a b)) (\<lambda>(x, t). fx x t)"
assumes U: "x0 \<in> U" "convex U"
shows "((\<lambda>x. integral (cbox a b) (f x)) has_field_derivative integral (cbox a b) (fx x0)) (at x0 within U)"
proof -
note [continuous_intros] =
continuous_on_compose2[OF cont_fx, where f="\<lambda>y. Pair x y" for x,
unfolded split_beta fst_conv snd_conv]
have *: "blinfun_mult_right (integral (cbox a b) (fx x0)) =
integral (cbox a b) (\<lambda>t. blinfun_mult_right (fx x0 t))"
by (subst integral_linear[symmetric])
(auto simp: has_vector_derivative_def o_def
intro!: integrable_continuous U continuous_intros bounded_linear_intros)
show ?thesis
unfolding has_field_derivative_eq_has_derivative_blinfun
apply (rule has_derivative_eq_rhs)
apply (rule leibniz_rule[OF _ integrable_f2 _ U, where fx="\<lambda>x t. blinfun_mult_right (fx x t)"])
using fx cont_fx
apply (auto simp: has_field_derivative_def * split_beta intro!: continuous_intros)
done
qed
lemma leibniz_rule_field_differentiable:
fixes f::"'a::{real_normed_field, banach} \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'a"
assumes "\<And>x t. x \<in> U \<Longrightarrow> t \<in> cbox a b \<Longrightarrow> ((\<lambda>x. f x t) has_field_derivative fx x t) (at x within U)"
assumes "\<And>x. x \<in> U \<Longrightarrow> (f x) integrable_on cbox a b"
assumes "continuous_on (U \<times> (cbox a b)) (\<lambda>(x, t). fx x t)"
assumes "x0 \<in> U" "convex U"
shows "(\<lambda>x. integral (cbox a b) (f x)) field_differentiable at x0 within U"
using leibniz_rule_field_derivative[OF assms] by (auto simp: field_differentiable_def)
subsection \<open>Exchange uniform limit and integral\<close>
lemma uniform_limit_integral_cbox:
fixes f::"'a \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::banach"
assumes u: "uniform_limit (cbox a b) f g F"
assumes c: "\<And>n. continuous_on (cbox a b) (f n)"
assumes [simp]: "F \<noteq> bot"
obtains I J where
"\<And>n. (f n has_integral I n) (cbox a b)"
"(g has_integral J) (cbox a b)"
"(I \<longlongrightarrow> J) F"
proof -
have fi[simp]: "f n integrable_on (cbox a b)" for n
by (auto intro!: integrable_continuous assms)
then obtain I where I: "\<And>n. (f n has_integral I n) (cbox a b)"
by atomize_elim (auto simp: integrable_on_def intro!: choice)
moreover
have gi[simp]: "g integrable_on (cbox a b)"
by (auto intro!: integrable_continuous uniform_limit_theorem[OF _ u] eventuallyI c)
then obtain J where J: "(g has_integral J) (cbox a b)"
by blast
moreover
have "(I \<longlongrightarrow> J) F"
proof cases
assume "content (cbox a b) = 0"
hence "I = (\<lambda>_. 0)" "J = 0"
by (auto intro!: has_integral_unique I J)
thus ?thesis by simp
next
assume content_nonzero: "content (cbox a b) \<noteq> 0"
show ?thesis
proof (rule tendstoI)
fix e::real
assume "e > 0"
define e' where "e' = e/2"
with \<open>e > 0\<close> have "e' > 0" by simp
then have "\<forall>\<^sub>F n in F. \<forall>x\<in>cbox a b. norm (f n x - g x) < e' / content (cbox a b)"
using u content_nonzero by (auto simp: uniform_limit_iff dist_norm zero_less_measure_iff)
then show "\<forall>\<^sub>F n in F. dist (I n) J < e"
proof eventually_elim
case (elim n)
have "I n = integral (cbox a b) (f n)"
"J = integral (cbox a b) g"
using I[of n] J by (simp_all add: integral_unique)
then have "dist (I n) J = norm (integral (cbox a b) (\<lambda>x. f n x - g x))"
by (simp add: integral_diff dist_norm)
also have "\<dots> \<le> integral (cbox a b) (\<lambda>x. (e' / content (cbox a b)))"
using elim
by (intro integral_norm_bound_integral) (auto intro!: integrable_diff)
also have "\<dots> < e"
using \<open>0 < e\<close>
by (simp add: e'_def)
finally show ?case .
qed
qed
qed
ultimately show ?thesis ..
qed
lemma uniform_limit_integral:
fixes f::"'a \<Rightarrow> 'b::ordered_euclidean_space \<Rightarrow> 'c::banach"
assumes u: "uniform_limit {a..b} f g F"
assumes c: "\<And>n. continuous_on {a..b} (f n)"
assumes [simp]: "F \<noteq> bot"
obtains I J where
"\<And>n. (f n has_integral I n) {a..b}"
"(g has_integral J) {a..b}"
"(I \<longlongrightarrow> J) F"
by (metis interval_cbox assms uniform_limit_integral_cbox)
subsection \<open>Integration by parts\<close>
lemma integration_by_parts_interior_strong:
fixes prod :: "_ \<Rightarrow> _ \<Rightarrow> 'b :: banach"
assumes bilinear: "bounded_bilinear (prod)"
assumes s: "finite s" and le: "a \<le> b"
assumes cont [continuous_intros]: "continuous_on {a..b} f" "continuous_on {a..b} g"
assumes deriv: "\<And>x. x\<in>{a<..<b} - s \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
"\<And>x. x\<in>{a<..<b} - s \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
assumes int: "((\<lambda>x. prod (f x) (g' x)) has_integral
(prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}"
shows "((\<lambda>x. prod (f' x) (g x)) has_integral y) {a..b}"
proof -
interpret bounded_bilinear prod by fact
have "((\<lambda>x. prod (f x) (g' x) + prod (f' x) (g x)) has_integral
(prod (f b) (g b) - prod (f a) (g a))) {a..b}"
using deriv by (intro fundamental_theorem_of_calculus_interior_strong[OF s le])
(auto intro!: continuous_intros continuous_on has_vector_derivative)
from has_integral_diff[OF this int] show ?thesis by (simp add: algebra_simps)
qed
lemma integration_by_parts_interior:
fixes prod :: "_ \<Rightarrow> _ \<Rightarrow> 'b :: banach"
assumes "bounded_bilinear (prod)" "a \<le> b"
"continuous_on {a..b} f" "continuous_on {a..b} g"
assumes "\<And>x. x\<in>{a<..<b} \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
"\<And>x. x\<in>{a<..<b} \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
assumes "((\<lambda>x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}"
shows "((\<lambda>x. prod (f' x) (g x)) has_integral y) {a..b}"
by (rule integration_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all)
lemma integration_by_parts:
fixes prod :: "_ \<Rightarrow> _ \<Rightarrow> 'b :: banach"
assumes "bounded_bilinear (prod)" "a \<le> b"
"continuous_on {a..b} f" "continuous_on {a..b} g"
assumes "\<And>x. x\<in>{a..b} \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
"\<And>x. x\<in>{a..b} \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
assumes "((\<lambda>x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}"
shows "((\<lambda>x. prod (f' x) (g x)) has_integral y) {a..b}"
by (rule integration_by_parts_interior[of _ _ _ f g f' g']) (insert assms, simp_all)
lemma integrable_by_parts_interior_strong:
fixes prod :: "_ \<Rightarrow> _ \<Rightarrow> 'b :: banach"
assumes bilinear: "bounded_bilinear (prod)"
assumes s: "finite s" and le: "a \<le> b"
assumes cont [continuous_intros]: "continuous_on {a..b} f" "continuous_on {a..b} g"
assumes deriv: "\<And>x. x\<in>{a<..<b} - s \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
"\<And>x. x\<in>{a<..<b} - s \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
assumes int: "(\<lambda>x. prod (f x) (g' x)) integrable_on {a..b}"
shows "(\<lambda>x. prod (f' x) (g x)) integrable_on {a..b}"
proof -
from int obtain I where "((\<lambda>x. prod (f x) (g' x)) has_integral I) {a..b}"
unfolding integrable_on_def by blast
hence "((\<lambda>x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) -
(prod (f b) (g b) - prod (f a) (g a) - I))) {a..b}" by simp
from integration_by_parts_interior_strong[OF assms(1-7) this]
show ?thesis unfolding integrable_on_def by blast
qed
lemma integrable_by_parts_interior:
fixes prod :: "_ \<Rightarrow> _ \<Rightarrow> 'b :: banach"
assumes "bounded_bilinear (prod)" "a \<le> b"
"continuous_on {a..b} f" "continuous_on {a..b} g"
assumes "\<And>x. x\<in>{a<..<b} \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
"\<And>x. x\<in>{a<..<b} \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
assumes "(\<lambda>x. prod (f x) (g' x)) integrable_on {a..b}"
shows "(\<lambda>x. prod (f' x) (g x)) integrable_on {a..b}"
by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all)
lemma integrable_by_parts:
fixes prod :: "_ \<Rightarrow> _ \<Rightarrow> 'b :: banach"
assumes "bounded_bilinear (prod)" "a \<le> b"
"continuous_on {a..b} f" "continuous_on {a..b} g"
assumes "\<And>x. x\<in>{a..b} \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
"\<And>x. x\<in>{a..b} \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
assumes "(\<lambda>x. prod (f x) (g' x)) integrable_on {a..b}"
shows "(\<lambda>x. prod (f' x) (g x)) integrable_on {a..b}"
by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all)
subsection \<open>Integration by substitution\<close>
lemma has_integral_substitution_general:
fixes f :: "real \<Rightarrow> 'a::euclidean_space" and g :: "real \<Rightarrow> real"
assumes s: "finite s" and le: "a \<le> b"
and subset: "g ` {a..b} \<subseteq> {c..d}"
and f [continuous_intros]: "continuous_on {c..d} f"
and g [continuous_intros]: "continuous_on {a..b} g"
and deriv [derivative_intros]:
"\<And>x. x \<in> {a..b} - s \<Longrightarrow> (g has_field_derivative g' x) (at x within {a..b})"
shows "((\<lambda>x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f - integral {g b..g a} f)) {a..b}"
proof -
let ?F = "\<lambda>x. integral {c..g x} f"
have cont_int: "continuous_on {a..b} ?F"
by (rule continuous_on_compose2[OF _ g subset] indefinite_integral_continuous_1
f integrable_continuous_real)+
have deriv: "(((\<lambda>x. integral {c..x} f) \<circ> g) has_vector_derivative g' x *\<^sub>R f (g x))
(at x within {a..b})" if "x \<in> {a..b} - s" for x
proof (rule has_vector_derivative_eq_rhs [OF vector_diff_chain_within refl])
show "(g has_vector_derivative g' x) (at x within {a..b})"
using deriv has_field_derivative_iff_has_vector_derivative that by blast
show "((\<lambda>x. integral {c..x} f) has_vector_derivative f (g x))
(at (g x) within g ` {a..b})"
using that le subset
by (blast intro: has_vector_derivative_within_subset integral_has_vector_derivative f)
qed
have deriv: "(?F has_vector_derivative g' x *\<^sub>R f (g x))
(at x)" if "x \<in> {a..b} - (s \<union> {a,b})" for x
using deriv[of x] that by (simp add: at_within_Icc_at o_def)
have "((\<lambda>x. g' x *\<^sub>R f (g x)) has_integral (?F b - ?F a)) {a..b}"
using le cont_int s deriv cont_int
by (intro fundamental_theorem_of_calculus_interior_strong[of "s \<union> {a,b}"]) simp_all
also
from subset have "g x \<in> {c..d}" if "x \<in> {a..b}" for x using that by blast
from this[of a] this[of b] le have cd: "c \<le> g a" "g b \<le> d" "c \<le> g b" "g a \<le> d" by auto
have "integral {c..g b} f - integral {c..g a} f = integral {g a..g b} f - integral {g b..g a} f"
proof cases
assume "g a \<le> g b"
note le = le this
from cd have "integral {c..g a} f + integral {g a..g b} f = integral {c..g b} f"
by (intro integral_combine integrable_continuous_real continuous_on_subset[OF f] le) simp_all
with le show ?thesis
by (cases "g a = g b") (simp_all add: algebra_simps)
next
assume less: "\<not>g a \<le> g b"
then have "g a \<ge> g b" by simp
note le = le this
from cd have "integral {c..g b} f + integral {g b..g a} f = integral {c..g a} f"
by (intro integral_combine integrable_continuous_real continuous_on_subset[OF f] le) simp_all
with less show ?thesis
by (simp_all add: algebra_simps)
qed
finally show ?thesis .
qed
lemma has_integral_substitution_strong:
fixes f :: "real \<Rightarrow> 'a::euclidean_space" and g :: "real \<Rightarrow> real"
assumes s: "finite s" and le: "a \<le> b" "g a \<le> g b"
and subset: "g ` {a..b} \<subseteq> {c..d}"
and f [continuous_intros]: "continuous_on {c..d} f"
and g [continuous_intros]: "continuous_on {a..b} g"
and deriv [derivative_intros]:
"\<And>x. x \<in> {a..b} - s \<Longrightarrow> (g has_field_derivative g' x) (at x within {a..b})"
shows "((\<lambda>x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f)) {a..b}"
using has_integral_substitution_general[OF s le(1) subset f g deriv] le(2)
by (cases "g a = g b") auto
lemma has_integral_substitution:
fixes f :: "real \<Rightarrow> 'a::euclidean_space" and g :: "real \<Rightarrow> real"
assumes "a \<le> b" "g a \<le> g b" "g ` {a..b} \<subseteq> {c..d}"
and "continuous_on {c..d} f"
and "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_field_derivative g' x) (at x within {a..b})"
shows "((\<lambda>x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f)) {a..b}"
by (intro has_integral_substitution_strong[of "{}" a b g c d] assms)
(auto intro: DERIV_continuous_on assms)
lemma integral_shift:
fixes f :: "real \<Rightarrow> 'a::euclidean_space"
assumes cont: "continuous_on {a + c..b + c} f"
shows "integral {a..b} (f \<circ> (\<lambda>x. x + c)) = integral {a + c..b + c} f"
proof (cases "a \<le> b")
case True
have "((\<lambda>x. 1 *\<^sub>R f (x + c)) has_integral integral {a+c..b+c} f) {a..b}"
using True cont
by (intro has_integral_substitution[where c = "a + c" and d = "b + c"])
(auto intro!: derivative_eq_intros)
thus ?thesis by (simp add: has_integral_iff o_def)
qed auto
subsection \<open>Compute a double integral using iterated integrals and switching the order of integration\<close>
lemma continuous_on_imp_integrable_on_Pair1:
fixes f :: "_ \<Rightarrow> 'b::banach"
assumes con: "continuous_on (cbox (a,c) (b,d)) f" and x: "x \<in> cbox a b"
shows "(\<lambda>y. f (x, y)) integrable_on (cbox c d)"
proof -
have "f \<circ> (\<lambda>y. (x, y)) integrable_on (cbox c d)"
apply (rule integrable_continuous)
apply (rule continuous_on_compose [OF _ continuous_on_subset [OF con]])
using x
apply (auto intro: continuous_on_Pair continuous_on_const continuous_on_id continuous_on_subset con)
done
then show ?thesis
by (simp add: o_def)
qed
lemma integral_integrable_2dim:
fixes f :: "('a::euclidean_space * 'b::euclidean_space) \<Rightarrow> 'c::banach"
assumes "continuous_on (cbox (a,c) (b,d)) f"
shows "(\<lambda>x. integral (cbox c d) (\<lambda>y. f (x,y))) integrable_on cbox a b"
proof (cases "content(cbox c d) = 0")
case True
then show ?thesis
by (simp add: True integrable_const)
next
case False
have uc: "uniformly_continuous_on (cbox (a,c) (b,d)) f"
by (simp add: assms compact_cbox compact_uniformly_continuous)
{ fix x::'a and e::real
assume x: "x \<in> cbox a b" and e: "0 < e"
then have e2_gt: "0 < e/2 / content (cbox c d)" and e2_less: "e/2 / content (cbox c d) * content (cbox c d) < e"
by (auto simp: False content_lt_nz e)
then obtain dd
where dd: "\<And>x x'. \<lbrakk>x\<in>cbox (a, c) (b, d); x'\<in>cbox (a, c) (b, d); norm (x' - x) < dd\<rbrakk>
\<Longrightarrow> norm (f x' - f x) \<le> e/(2 * content (cbox c d))" "dd>0"
using uc [unfolded uniformly_continuous_on_def, THEN spec, of "e/(2 * content (cbox c d))"]
by (auto simp: dist_norm intro: less_imp_le)
have "\<exists>delta>0. \<forall>x'\<in>cbox a b. norm (x' - x) < delta \<longrightarrow> norm (integral (cbox c d) (\<lambda>u. f (x', u) - f (x, u))) < e"
apply (rule_tac x=dd in exI)
using dd e2_gt assms x
apply clarify
apply (rule le_less_trans [OF _ e2_less])
apply (rule integrable_bound)
apply (auto intro: integrable_diff continuous_on_imp_integrable_on_Pair1)
done
} note * = this
show ?thesis
apply (rule integrable_continuous)
apply (simp add: * continuous_on_iff dist_norm integral_diff [symmetric] continuous_on_imp_integrable_on_Pair1 [OF assms])
done
qed
lemma integral_split:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::{real_normed_vector,complete_space}"
assumes f: "f integrable_on (cbox a b)"
and k: "k \<in> Basis"
shows "integral (cbox a b) f =
integral (cbox a b \<inter> {x. x\<bullet>k \<le> c}) f +
integral (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) f"
apply (rule integral_unique [OF has_integral_split [where c=c]])
using k f
apply (auto simp: has_integral_integral [symmetric])
done
lemma integral_swap_operativeI:
fixes f :: "('a::euclidean_space * 'b::euclidean_space) \<Rightarrow> 'c::banach"
assumes f: "continuous_on s f" and e: "0 < e"
shows "operative conj True
(\<lambda>k. \<forall>a b c d.
cbox (a,c) (b,d) \<subseteq> k \<and> cbox (a,c) (b,d) \<subseteq> s
\<longrightarrow> norm(integral (cbox (a,c) (b,d)) f -
integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f((x,y)))))
\<le> e * content (cbox (a,c) (b,d)))"
proof (standard, auto)
fix a::'a and c::'b and b::'a and d::'b and u::'a and v::'a and w::'b and z::'b
assume *: "box (a, c) (b, d) = {}"
and cb1: "cbox (u, w) (v, z) \<subseteq> cbox (a, c) (b, d)"
and cb2: "cbox (u, w) (v, z) \<subseteq> s"
then have c0: "content (cbox (a, c) (b, d)) = 0"
using * unfolding content_eq_0_interior by simp
have c0': "content (cbox (u, w) (v, z)) = 0"
by (fact content_0_subset [OF c0 cb1])
show "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y))))
\<le> e * content (cbox (u,w) (v,z))"
using content_cbox_pair_eq0_D [OF c0']
by (force simp add: c0')
next
fix a::'a and c::'b and b::'a and d::'b
and M::real and i::'a and j::'b
and u::'a and v::'a and w::'b and z::'b
assume ij: "(i,j) \<in> Basis"
and n1: "\<forall>a' b' c' d'.
cbox (a',c') (b',d') \<subseteq> cbox (a,c) (b,d) \<and>
cbox (a',c') (b',d') \<subseteq> {x. x \<bullet> (i,j) \<le> M} \<and> cbox (a',c') (b',d') \<subseteq> s \<longrightarrow>
norm (integral (cbox (a',c') (b',d')) f - integral (cbox a' b') (\<lambda>x. integral (cbox c' d') (\<lambda>y. f (x,y))))
\<le> e * content (cbox (a',c') (b',d'))"
and n2: "\<forall>a' b' c' d'.
cbox (a',c') (b',d') \<subseteq> cbox (a,c) (b,d) \<and>
cbox (a',c') (b',d') \<subseteq> {x. M \<le> x \<bullet> (i,j)} \<and> cbox (a',c') (b',d') \<subseteq> s \<longrightarrow>
norm (integral (cbox (a',c') (b',d')) f - integral (cbox a' b') (\<lambda>x. integral (cbox c' d') (\<lambda>y. f (x,y))))
\<le> e * content (cbox (a',c') (b',d'))"
and subs: "cbox (u,w) (v,z) \<subseteq> cbox (a,c) (b,d)" "cbox (u,w) (v,z) \<subseteq> s"
have fcont: "continuous_on (cbox (u, w) (v, z)) f"
using assms(1) continuous_on_subset subs(2) by blast
then have fint: "f integrable_on cbox (u, w) (v, z)"
by (metis integrable_continuous)
consider "i \<in> Basis" "j=0" | "j \<in> Basis" "i=0" using ij
by (auto simp: Euclidean_Space.Basis_prod_def)
then show "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x,y))))
\<le> e * content (cbox (u,w) (v,z))" (is ?normle)
proof cases
case 1
then have e: "e * content (cbox (u, w) (v, z)) =
e * (content (cbox u v \<inter> {x. x \<bullet> i \<le> M}) * content (cbox w z)) +
e * (content (cbox u v \<inter> {x. M \<le> x \<bullet> i}) * content (cbox w z))"
by (simp add: content_split [where c=M] content_Pair algebra_simps)
have *: "integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y))) =
integral (cbox u v \<inter> {x. x \<bullet> i \<le> M}) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y))) +
integral (cbox u v \<inter> {x. M \<le> x \<bullet> i}) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y)))"
using 1 f subs integral_integrable_2dim continuous_on_subset
by (blast intro: integral_split)
show ?normle
apply (rule norm_diff2 [OF integral_split [where c=M, OF fint ij] * e])
using 1 subs
apply (simp_all add: cbox_Pair_eq setcomp_dot1 [of "\<lambda>u. M\<le>u"] setcomp_dot1 [of "\<lambda>u. u\<le>M"] Sigma_Int_Paircomp1)
apply (simp_all add: interval_split ij)
apply (simp_all add: cbox_Pair_eq [symmetric] content_Pair [symmetric])
apply (force simp add: interval_split [symmetric] intro!: n1 [rule_format])
apply (force simp add: interval_split [symmetric] intro!: n2 [rule_format])
done
next
case 2
then have e: "e * content (cbox (u, w) (v, z)) =
e * (content (cbox u v) * content (cbox w z \<inter> {x. x \<bullet> j \<le> M})) +
e * (content (cbox u v) * content (cbox w z \<inter> {x. M \<le> x \<bullet> j}))"
by (simp add: content_split [where c=M] content_Pair algebra_simps)
have "(\<lambda>x. integral (cbox w z \<inter> {x. x \<bullet> j \<le> M}) (\<lambda>y. f (x, y))) integrable_on cbox u v"
"(\<lambda>x. integral (cbox w z \<inter> {x. M \<le> x \<bullet> j}) (\<lambda>y. f (x, y))) integrable_on cbox u v"
using 2 subs
apply (simp_all add: interval_split)
apply (rule_tac [!] integral_integrable_2dim [OF continuous_on_subset [OF f]])
apply (auto simp: cbox_Pair_eq interval_split [symmetric])
done
with 2 have *: "integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y))) =
integral (cbox u v) (\<lambda>x. integral (cbox w z \<inter> {x. x \<bullet> j \<le> M}) (\<lambda>y. f (x, y))) +
integral (cbox u v) (\<lambda>x. integral (cbox w z \<inter> {x. M \<le> x \<bullet> j}) (\<lambda>y. f (x, y)))"
by (simp add: integral_add [symmetric] integral_split [symmetric]
continuous_on_imp_integrable_on_Pair1 [OF fcont] cong: integral_cong)
show ?normle
apply (rule norm_diff2 [OF integral_split [where c=M, OF fint ij] * e])
using 2 subs
apply (simp_all add: cbox_Pair_eq setcomp_dot2 [of "\<lambda>u. M\<le>u"] setcomp_dot2 [of "\<lambda>u. u\<le>M"] Sigma_Int_Paircomp2)
apply (simp_all add: interval_split ij)
apply (simp_all add: cbox_Pair_eq [symmetric] content_Pair [symmetric])
apply (force simp add: interval_split [symmetric] intro!: n1 [rule_format])
apply (force simp add: interval_split [symmetric] intro!: n2 [rule_format])
done
qed
qed
lemma integral_Pair_const:
"integral (cbox (a,c) (b,d)) (\<lambda>x. k) =
integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. k))"
by (simp add: content_Pair)
lemma integral_prod_continuous:
fixes f :: "('a::euclidean_space * 'b::euclidean_space) \<Rightarrow> 'c::banach"
assumes "continuous_on (cbox (a, c) (b, d)) f" (is "continuous_on ?CBOX f")
shows "integral (cbox (a, c) (b, d)) f = integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f (x, y)))"
proof (cases "content ?CBOX = 0")
case True
then show ?thesis
by (auto simp: content_Pair)
next
case False
then have cbp: "content ?CBOX > 0"
using content_lt_nz by blast
have "norm (integral ?CBOX f - integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f (x,y)))) = 0"
proof (rule dense_eq0_I, simp)
fix e :: real
assume "0 < e"
with \<open>content ?CBOX > 0\<close> have "0 < e/content ?CBOX"
by simp
have f_int_acbd: "f integrable_on ?CBOX"
by (rule integrable_continuous [OF assms])
{ fix p
assume p: "p division_of ?CBOX"
then have "finite p"
by blast
define e' where "e' = e/content ?CBOX"
with \<open>0 < e\<close> \<open>0 < e/content ?CBOX\<close>
have "0 < e'"
by simp
interpret operative conj True
"\<lambda>k. \<forall>a' b' c' d'.
cbox (a', c') (b', d') \<subseteq> k \<and> cbox (a', c') (b', d') \<subseteq> ?CBOX
\<longrightarrow> norm (integral (cbox (a', c') (b', d')) f -
integral (cbox a' b') (\<lambda>x. integral (cbox c' d') (\<lambda>y. f ((x, y)))))
\<le> e' * content (cbox (a', c') (b', d'))"
using assms \<open>0 < e'\<close> by (rule integral_swap_operativeI)
have "norm (integral ?CBOX f - integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f (x, y))))
\<le> e' * content ?CBOX"
if "\<And>t u v w z. t \<in> p \<Longrightarrow> cbox (u, w) (v, z) \<subseteq> t \<Longrightarrow> cbox (u, w) (v, z) \<subseteq> ?CBOX
\<Longrightarrow> norm (integral (cbox (u, w) (v, z)) f -
integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y))))
\<le> e' * content (cbox (u, w) (v, z))"
using that division [of p "(a, c)" "(b, d)"] p \<open>finite p\<close> by (auto simp add: comm_monoid_set_F_and)
with False have "norm (integral ?CBOX f - integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f (x, y))))
\<le> e"
if "\<And>t u v w z. t \<in> p \<Longrightarrow> cbox (u, w) (v, z) \<subseteq> t \<Longrightarrow> cbox (u, w) (v, z) \<subseteq> ?CBOX
\<Longrightarrow> norm (integral (cbox (u, w) (v, z)) f -
integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y))))
\<le> e * content (cbox (u, w) (v, z)) / content ?CBOX"
using that by (simp add: e'_def)
} note op_acbd = this
{ fix k::real and \<D> and u::'a and v w and z::'b and t1 t2 l
assume k: "0 < k"
and nf: "\<And>x y u v.
\<lbrakk>x \<in> cbox a b; y \<in> cbox c d; u \<in> cbox a b; v\<in>cbox c d; norm (u-x, v-y) < k\<rbrakk>
\<Longrightarrow> norm (f(u,v) - f(x,y)) < e/(2 * (content ?CBOX))"
and p_acbd: "\<D> tagged_division_of cbox (a,c) (b,d)"
and fine: "(\<lambda>x. ball x k) fine \<D>" "((t1,t2), l) \<in> \<D>"
and uwvz_sub: "cbox (u,w) (v,z) \<subseteq> l" "cbox (u,w) (v,z) \<subseteq> cbox (a,c) (b,d)"
have t: "t1 \<in> cbox a b" "t2 \<in> cbox c d"
by (meson fine p_acbd cbox_Pair_iff tag_in_interval)+
have ls: "l \<subseteq> ball (t1,t2) k"
using fine by (simp add: fine_def Ball_def)
{ fix x1 x2
assume xuvwz: "x1 \<in> cbox u v" "x2 \<in> cbox w z"
then have x: "x1 \<in> cbox a b" "x2 \<in> cbox c d"
using uwvz_sub by auto
have "norm (x1 - t1, x2 - t2) = norm (t1 - x1, t2 - x2)"
by (simp add: norm_Pair norm_minus_commute)
also have "norm (t1 - x1, t2 - x2) < k"
using xuvwz ls uwvz_sub unfolding ball_def
by (force simp add: cbox_Pair_eq dist_norm )
finally have "norm (f (x1,x2) - f (t1,t2)) \<le> e/content ?CBOX/2"
using nf [OF t x] by simp
} note nf' = this
have f_int_uwvz: "f integrable_on cbox (u,w) (v,z)"
using f_int_acbd uwvz_sub integrable_on_subcbox by blast
have f_int_uv: "\<And>x. x \<in> cbox u v \<Longrightarrow> (\<lambda>y. f (x,y)) integrable_on cbox w z"
using assms continuous_on_subset uwvz_sub
by (blast intro: continuous_on_imp_integrable_on_Pair1)
have 1: "norm (integral (cbox (u,w) (v,z)) f - integral (cbox (u,w) (v,z)) (\<lambda>x. f (t1,t2)))
\<le> e * content (cbox (u,w) (v,z)) / content ?CBOX/2"
apply (simp only: integral_diff [symmetric] f_int_uwvz integrable_const)
apply (rule order_trans [OF integrable_bound [of "e/content ?CBOX/2"]])
using cbp \<open>0 < e/content ?CBOX\<close> nf'
apply (auto simp: integrable_diff f_int_uwvz integrable_const)
done
have int_integrable: "(\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y))) integrable_on cbox u v"
using assms integral_integrable_2dim continuous_on_subset uwvz_sub(2) by blast
have normint_wz:
"\<And>x. x \<in> cbox u v \<Longrightarrow>
norm (integral (cbox w z) (\<lambda>y. f (x, y)) - integral (cbox w z) (\<lambda>y. f (t1, t2)))
\<le> e * content (cbox w z) / content (cbox (a, c) (b, d))/2"
apply (simp only: integral_diff [symmetric] f_int_uv integrable_const)
apply (rule order_trans [OF integrable_bound [of "e/content ?CBOX/2"]])
using cbp \<open>0 < e/content ?CBOX\<close> nf'
apply (auto simp: integrable_diff f_int_uv integrable_const)
done
have "norm (integral (cbox u v)
(\<lambda>x. integral (cbox w z) (\<lambda>y. f (x,y)) - integral (cbox w z) (\<lambda>y. f (t1,t2))))
\<le> e * content (cbox w z) / content ?CBOX/2 * content (cbox u v)"
apply (rule integrable_bound [OF _ _ normint_wz])
using cbp \<open>0 < e/content ?CBOX\<close>
apply (auto simp: field_split_simps integrable_diff int_integrable integrable_const)
done
also have "... \<le> e * content (cbox (u,w) (v,z)) / content ?CBOX/2"
by (simp add: content_Pair field_split_simps)
finally have 2: "norm (integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x,y))) -
integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (t1,t2))))
\<le> e * content (cbox (u,w) (v,z)) / content ?CBOX/2"
by (simp only: integral_diff [symmetric] int_integrable integrable_const)
have norm_xx: "\<lbrakk>x' = y'; norm(x - x') \<le> e/2; norm(y - y') \<le> e/2\<rbrakk> \<Longrightarrow> norm(x - y) \<le> e" for x::'c and y x' y' e
using norm_triangle_mono [of "x-y'" "e/2" "y'-y" "e/2"] field_sum_of_halves
by (simp add: norm_minus_commute)
have "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x,y))))
\<le> e * content (cbox (u,w) (v,z)) / content ?CBOX"
by (rule norm_xx [OF integral_Pair_const 1 2])
} note * = this
have "norm (integral ?CBOX f - integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f (x,y)))) \<le> e"
if "\<forall>x\<in>?CBOX. \<forall>x'\<in>?CBOX. norm (x' - x) < k \<longrightarrow> norm (f x' - f x) < e /(2 * content (?CBOX))" "0 < k" for k
proof -
obtain p where ptag: "p tagged_division_of cbox (a, c) (b, d)"
and fine: "(\<lambda>x. ball x k) fine p"
using fine_division_exists \<open>0 < k\<close> by blast
show ?thesis
apply (rule op_acbd [OF division_of_tagged_division [OF ptag]])
using that fine ptag \<open>0 < k\<close> by (auto simp: *)
qed
then show "norm (integral ?CBOX f - integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f (x,y)))) \<le> e"
using compact_uniformly_continuous [OF assms compact_cbox]
apply (simp add: uniformly_continuous_on_def dist_norm)
apply (drule_tac x="e/2 / content?CBOX" in spec)
using cbp \<open>0 < e\<close> by (auto simp: zero_less_mult_iff)
qed
then show ?thesis
by simp
qed
lemma integral_swap_2dim:
fixes f :: "['a::euclidean_space, 'b::euclidean_space] \<Rightarrow> 'c::banach"
assumes "continuous_on (cbox (a,c) (b,d)) (\<lambda>(x,y). f x y)"
shows "integral (cbox (a, c) (b, d)) (\<lambda>(x, y). f x y) = integral (cbox (c, a) (d, b)) (\<lambda>(x, y). f y x)"
proof -
have "((\<lambda>(x, y). f x y) has_integral integral (cbox (c, a) (d, b)) (\<lambda>(x, y). f y x)) (prod.swap ` (cbox (c, a) (d, b)))"
apply (rule has_integral_twiddle [of 1 prod.swap prod.swap "\<lambda>(x,y). f y x" "integral (cbox (c, a) (d, b)) (\<lambda>(x, y). f y x)", simplified])
apply (force simp: isCont_swap content_Pair has_integral_integral [symmetric] integrable_continuous swap_continuous assms)+
done
then show ?thesis
by force
qed
theorem integral_swap_continuous:
fixes f :: "['a::euclidean_space, 'b::euclidean_space] \<Rightarrow> 'c::banach"
assumes "continuous_on (cbox (a,c) (b,d)) (\<lambda>(x,y). f x y)"
shows "integral (cbox a b) (\<lambda>x. integral (cbox c d) (f x)) =
integral (cbox c d) (\<lambda>y. integral (cbox a b) (\<lambda>x. f x y))"
proof -
have "integral (cbox a b) (\<lambda>x. integral (cbox c d) (f x)) = integral (cbox (a, c) (b, d)) (\<lambda>(x, y). f x y)"
using integral_prod_continuous [OF assms] by auto
also have "... = integral (cbox (c, a) (d, b)) (\<lambda>(x, y). f y x)"
by (rule integral_swap_2dim [OF assms])
also have "... = integral (cbox c d) (\<lambda>y. integral (cbox a b) (\<lambda>x. f x y))"
using integral_prod_continuous [OF swap_continuous] assms
by auto
finally show ?thesis .
qed
subsection \<open>Definite integrals for exponential and power function\<close>
lemma has_integral_exp_minus_to_infinity:
assumes a: "a > 0"
shows "((\<lambda>x::real. exp (-a*x)) has_integral exp (-a*c)/a) {c..}"
proof -
define f where "f = (\<lambda>k x. if x \<in> {c..real k} then exp (-a*x) else 0)"
{
fix k :: nat assume k: "of_nat k \<ge> c"
from k a
have "((\<lambda>x. exp (-a*x)) has_integral (-exp (-a*real k)/a - (-exp (-a*c)/a))) {c..real k}"
by (intro fundamental_theorem_of_calculus)
(auto intro!: derivative_eq_intros
simp: has_field_derivative_iff_has_vector_derivative [symmetric])
hence "(f k has_integral (exp (-a*c)/a - exp (-a*real k)/a)) {c..}" unfolding f_def
by (subst has_integral_restrict) simp_all
} note has_integral_f = this
have [simp]: "f k = (\<lambda>_. 0)" if "of_nat k < c" for k using that by (auto simp: fun_eq_iff f_def)
have integral_f: "integral {c..} (f k) =
(if real k \<ge> c then exp (-a*c)/a - exp (-a*real k)/a else 0)"
for k using integral_unique[OF has_integral_f[of k]] by simp
have A: "(\<lambda>x. exp (-a*x)) integrable_on {c..} \<and>
(\<lambda>k. integral {c..} (f k)) \<longlonglongrightarrow> integral {c..} (\<lambda>x. exp (-a*x))"
proof (intro monotone_convergence_increasing allI ballI)
fix k ::nat
have "(\<lambda>x. exp (-a*x)) integrable_on {c..of_real k}" (is ?P)
unfolding f_def by (auto intro!: continuous_intros integrable_continuous_real)
hence "(f k) integrable_on {c..of_real k}"
by (rule integrable_eq) (simp add: f_def)
then show "f k integrable_on {c..}"
by (rule integrable_on_superset) (auto simp: f_def)
next
fix x assume x: "x \<in> {c..}"
have "sequentially \<le> principal {nat \<lceil>x\<rceil>..}" unfolding at_top_def by (simp add: Inf_lower)
also have "{nat \<lceil>x\<rceil>..} \<subseteq> {k. x \<le> real k}" by auto
also have "inf (principal \<dots>) (principal {k. \<not>x \<le> real k}) =
principal ({k. x \<le> real k} \<inter> {k. \<not>x \<le> real k})" by simp
also have "{k. x \<le> real k} \<inter> {k. \<not>x \<le> real k} = {}" by blast
finally have "inf sequentially (principal {k. \<not>x \<le> real k}) = bot"
by (simp add: inf.coboundedI1 bot_unique)
with x show "(\<lambda>k. f k x) \<longlonglongrightarrow> exp (-a*x)" unfolding f_def
by (intro filterlim_If) auto
next
have "\<bar>integral {c..} (f k)\<bar> \<le> exp (-a*c)/a" for k
proof (cases "c > of_nat k")
case False
hence "abs (integral {c..} (f k)) = abs (exp (- (a * c)) / a - exp (- (a * real k)) / a)"
by (simp add: integral_f)
also have "abs (exp (- (a * c)) / a - exp (- (a * real k)) / a) =
exp (- (a * c)) / a - exp (- (a * real k)) / a"
using False a by (intro abs_of_nonneg) (simp_all add: field_simps)
also have "\<dots> \<le> exp (- a * c) / a" using a by simp
finally show ?thesis .
qed (insert a, simp_all add: integral_f)
thus "bounded (range(\<lambda>k. integral {c..} (f k)))"
by (intro boundedI[of _ "exp (-a*c)/a"]) auto
qed (auto simp: f_def)
have "(\<lambda>k. exp (-a*c)/a - exp (-a * of_nat k)/a) \<longlonglongrightarrow> exp (-a*c)/a - 0/a"
by (intro tendsto_intros filterlim_compose[OF exp_at_bot]
filterlim_tendsto_neg_mult_at_bot[OF tendsto_const] filterlim_real_sequentially)+
(insert a, simp_all)
moreover
from eventually_gt_at_top[of "nat \<lceil>c\<rceil>"] have "eventually (\<lambda>k. of_nat k > c) sequentially"
by eventually_elim linarith
hence "eventually (\<lambda>k. exp (-a*c)/a - exp (-a * of_nat k)/a = integral {c..} (f k)) sequentially"
by eventually_elim (simp add: integral_f)
ultimately have "(\<lambda>k. integral {c..} (f k)) \<longlonglongrightarrow> exp (-a*c)/a - 0/a"
by (rule Lim_transform_eventually)
from LIMSEQ_unique[OF conjunct2[OF A] this]
have "integral {c..} (\<lambda>x. exp (-a*x)) = exp (-a*c)/a" by simp
with conjunct1[OF A] show ?thesis
by (simp add: has_integral_integral)
qed
lemma integrable_on_exp_minus_to_infinity: "a > 0 \<Longrightarrow> (\<lambda>x. exp (-a*x) :: real) integrable_on {c..}"
using has_integral_exp_minus_to_infinity[of a c] unfolding integrable_on_def by blast
lemma has_integral_powr_from_0:
assumes a: "a > (-1::real)" and c: "c \<ge> 0"
shows "((\<lambda>x. x powr a) has_integral (c powr (a+1) / (a+1))) {0..c}"
proof (cases "c = 0")
case False
define f where "f = (\<lambda>k x. if x \<in> {inverse (of_nat (Suc k))..c} then x powr a else 0)"
define F where "F = (\<lambda>k. if inverse (of_nat (Suc k)) \<le> c then
c powr (a+1)/(a+1) - inverse (real (Suc k)) powr (a+1)/(a+1) else 0)"
{
fix k :: nat
have "(f k has_integral F k) {0..c}"
proof (cases "inverse (of_nat (Suc k)) \<le> c")
case True
{
fix x assume x: "x \<ge> inverse (1 + real k)"
have "0 < inverse (1 + real k)" by simp
also note x
finally have "x > 0" .
} note x = this
hence "((\<lambda>x. x powr a) has_integral c powr (a + 1) / (a + 1) -
inverse (real (Suc k)) powr (a + 1) / (a + 1)) {inverse (real (Suc k))..c}"
using True a by (intro fundamental_theorem_of_calculus)
(auto intro!: derivative_eq_intros continuous_on_powr' continuous_on_const
simp: has_field_derivative_iff_has_vector_derivative [symmetric])
with True show ?thesis unfolding f_def F_def by (subst has_integral_restrict) simp_all
next
case False
thus ?thesis unfolding f_def F_def by (subst has_integral_restrict) auto
qed
} note has_integral_f = this
have integral_f: "integral {0..c} (f k) = F k" for k
using has_integral_f[of k] by (rule integral_unique)
have A: "(\<lambda>x. x powr a) integrable_on {0..c} \<and>
(\<lambda>k. integral {0..c} (f k)) \<longlonglongrightarrow> integral {0..c} (\<lambda>x. x powr a)"
proof (intro monotone_convergence_increasing ballI allI)
fix k from has_integral_f[of k] show "f k integrable_on {0..c}"
by (auto simp: integrable_on_def)
next
fix k :: nat and x :: real
{
assume x: "inverse (real (Suc k)) \<le> x"
have "inverse (real (Suc (Suc k))) \<le> inverse (real (Suc k))" by (simp add: field_simps)
also note x
finally have "inverse (real (Suc (Suc k))) \<le> x" .
}
thus "f k x \<le> f (Suc k) x" by (auto simp: f_def simp del: of_nat_Suc)
next
fix x assume x: "x \<in> {0..c}"
show "(\<lambda>k. f k x) \<longlonglongrightarrow> x powr a"
proof (cases "x = 0")
case False
with x have "x > 0" by simp
from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this]
have "eventually (\<lambda>k. x powr a = f k x) sequentially"
by eventually_elim (insert x, simp add: f_def)
moreover have "(\<lambda>_. x powr a) \<longlonglongrightarrow> x powr a" by simp
ultimately show ?thesis by (blast intro: Lim_transform_eventually)
qed (simp_all add: f_def)
next
{
fix k
from a have "F k \<le> c powr (a + 1) / (a + 1)"
by (auto simp add: F_def divide_simps)
also from a have "F k \<ge> 0"
by (auto simp: F_def divide_simps simp del: of_nat_Suc intro!: powr_mono2)
hence "F k = abs (F k)" by simp
finally have "abs (F k) \<le> c powr (a + 1) / (a + 1)" .
}
thus "bounded (range(\<lambda>k. integral {0..c} (f k)))"
by (intro boundedI[of _ "c powr (a+1) / (a+1)"]) (auto simp: integral_f)
qed
from False c have "c > 0" by simp
from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this]
have "eventually (\<lambda>k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a+1) / (a+1) =
integral {0..c} (f k)) sequentially"
by eventually_elim (simp add: integral_f F_def)
moreover have "(\<lambda>k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1))
\<longlonglongrightarrow> c powr (a + 1) / (a + 1) - 0 powr (a + 1) / (a + 1)"
using a by (intro tendsto_intros LIMSEQ_inverse_real_of_nat) auto
hence "(\<lambda>k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1))
\<longlonglongrightarrow> c powr (a + 1) / (a + 1)" by simp
ultimately have "(\<lambda>k. integral {0..c} (f k)) \<longlonglongrightarrow> c powr (a+1) / (a+1)"
by (blast intro: Lim_transform_eventually)
with A have "integral {0..c} (\<lambda>x. x powr a) = c powr (a+1) / (a+1)"
by (blast intro: LIMSEQ_unique)
with A show ?thesis by (simp add: has_integral_integral)
qed (simp_all add: has_integral_refl)
lemma integrable_on_powr_from_0:
assumes a: "a > (-1::real)" and c: "c \<ge> 0"
shows "(\<lambda>x. x powr a) integrable_on {0..c}"
using has_integral_powr_from_0[OF assms] unfolding integrable_on_def by blast
lemma has_integral_powr_to_inf:
fixes a e :: real
assumes "e < -1" "a > 0"
shows "((\<lambda>x. x powr e) has_integral -(a powr (e + 1)) / (e + 1)) {a..}"
proof -
define f :: "nat \<Rightarrow> real \<Rightarrow> real" where "f = (\<lambda>n x. if x \<in> {a..n} then x powr e else 0)"
define F where "F = (\<lambda>x. x powr (e + 1) / (e + 1))"
have has_integral_f: "(f n has_integral (F n - F a)) {a..}"
if n: "n \<ge> a" for n :: nat
proof -
from n assms have "((\<lambda>x. x powr e) has_integral (F n - F a)) {a..n}"
by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros
simp: has_field_derivative_iff_has_vector_derivative [symmetric] F_def)
hence "(f n has_integral (F n - F a)) {a..n}"
by (rule has_integral_eq [rotated]) (simp add: f_def)
thus "(f n has_integral (F n - F a)) {a..}"
by (rule has_integral_on_superset) (auto simp: f_def)
qed
have integral_f: "integral {a..} (f n) = (if n \<ge> a then F n - F a else 0)" for n :: nat
proof (cases "n \<ge> a")
case True
with has_integral_f[OF this] show ?thesis by (simp add: integral_unique)
next
case False
have "(f n has_integral 0) {a}" by (rule has_integral_refl)
hence "(f n has_integral 0) {a..}"
by (rule has_integral_on_superset) (insert False, simp_all add: f_def)
with False show ?thesis by (simp add: integral_unique)
qed
have *: "(\<lambda>x. x powr e) integrable_on {a..} \<and>
(\<lambda>n. integral {a..} (f n)) \<longlonglongrightarrow> integral {a..} (\<lambda>x. x powr e)"
proof (intro monotone_convergence_increasing allI ballI)
fix n :: nat
from assms have "(\<lambda>x. x powr e) integrable_on {a..n}"
by (auto intro!: integrable_continuous_real continuous_intros)
hence "f n integrable_on {a..n}"
by (rule integrable_eq) (auto simp: f_def)
thus "f n integrable_on {a..}"
by (rule integrable_on_superset) (auto simp: f_def)
next
fix n :: nat and x :: real
show "f n x \<le> f (Suc n) x" by (auto simp: f_def)
next
fix x :: real assume x: "x \<in> {a..}"
from filterlim_real_sequentially
have "eventually (\<lambda>n. real n \<ge> x) at_top"
by (simp add: filterlim_at_top)
with x have "eventually (\<lambda>n. f n x = x powr e) at_top"
by (auto elim!: eventually_mono simp: f_def)
thus "(\<lambda>n. f n x) \<longlonglongrightarrow> x powr e" by (simp add: tendsto_eventually)
next
have "norm (integral {a..} (f n)) \<le> -F a" for n :: nat
proof (cases "n \<ge> a")
case True
with assms have "a powr (e + 1) \<ge> n powr (e + 1)"
by (intro powr_mono2') simp_all
with assms show ?thesis by (auto simp: divide_simps F_def integral_f)
qed (insert assms, simp add: integral_f F_def field_split_simps)
thus "bounded (range(\<lambda>k. integral {a..} (f k)))"
unfolding bounded_iff by (intro exI[of _ "-F a"]) auto
qed
from filterlim_real_sequentially
have "eventually (\<lambda>n. real n \<ge> a) at_top"
by (simp add: filterlim_at_top)
hence "eventually (\<lambda>n. F n - F a = integral {a..} (f n)) at_top"
by eventually_elim (simp add: integral_f)
moreover have "(\<lambda>n. F n - F a) \<longlonglongrightarrow> 0 / (e + 1) - F a" unfolding F_def
by (insert assms, (rule tendsto_intros filterlim_compose[OF tendsto_neg_powr]
filterlim_ident filterlim_real_sequentially | simp)+)
hence "(\<lambda>n. F n - F a) \<longlonglongrightarrow> -F a" by simp
ultimately have "(\<lambda>n. integral {a..} (f n)) \<longlonglongrightarrow> -F a" by (blast intro: Lim_transform_eventually)
from conjunct2[OF *] and this
have "integral {a..} (\<lambda>x. x powr e) = -F a" by (rule LIMSEQ_unique)
with conjunct1[OF *] show ?thesis
by (simp add: has_integral_integral F_def)
qed
lemma has_integral_inverse_power_to_inf:
fixes a :: real and n :: nat
assumes "n > 1" "a > 0"
shows "((\<lambda>x. 1 / x ^ n) has_integral 1 / (real (n - 1) * a ^ (n - 1))) {a..}"
proof -
from assms have "real_of_int (-int n) < -1" by simp
note has_integral_powr_to_inf[OF this \<open>a > 0\<close>]
also have "- (a powr (real_of_int (- int n) + 1)) / (real_of_int (- int n) + 1) =
1 / (real (n - 1) * a powr (real (n - 1)))" using assms
by (simp add: field_split_simps powr_add [symmetric] of_nat_diff)
also from assms have "a powr (real (n - 1)) = a ^ (n - 1)"
by (intro powr_realpow)
finally show ?thesis
by (rule has_integral_eq [rotated])
(insert assms, simp_all add: powr_minus powr_realpow field_split_simps)
qed
subsubsection \<open>Adaption to ordered Euclidean spaces and the Cartesian Euclidean space\<close>
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_vec_nth,unfolded o_def] .
lemma content_closed_interval:
fixes a :: "'a::ordered_euclidean_space"
assumes "a \<le> b"
shows "content {a..b} = (\<Prod>i\<in>Basis. b\<bullet>i - a\<bullet>i)"
using content_cbox[of a b] assms by (simp add: cbox_interval eucl_le[where 'a='a])
lemma integrable_const_ivl[intro]:
fixes a::"'a::ordered_euclidean_space"
shows "(\<lambda>x. c) integrable_on {a..b}"
unfolding cbox_interval[symmetric] by (rule integrable_const)
lemma integrable_on_subinterval:
fixes f :: "'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
assumes "f integrable_on S" "{a..b} \<subseteq> S"
shows "f integrable_on {a..b}"
using integrable_on_subcbox[of f S a b] assms by (simp add: cbox_interval)
end
diff --git a/src/HOL/Analysis/Homeomorphism.thy b/src/HOL/Analysis/Homeomorphism.thy
--- a/src/HOL/Analysis/Homeomorphism.thy
+++ b/src/HOL/Analysis/Homeomorphism.thy
@@ -1,2286 +1,2286 @@
(* Title: HOL/Analysis/Homeomorphism.thy
Author: LC Paulson (ported from HOL Light)
*)
section \<open>Homeomorphism Theorems\<close>
theory Homeomorphism
imports Homotopy
begin
lemma homeomorphic_spheres':
fixes a ::"'a::euclidean_space" and b ::"'b::euclidean_space"
assumes "0 < \<delta>" and dimeq: "DIM('a) = DIM('b)"
shows "(sphere a \<delta>) homeomorphic (sphere b \<delta>)"
proof -
obtain f :: "'a\<Rightarrow>'b" and g where "linear f" "linear g"
and fg: "\<And>x. norm(f x) = norm x" "\<And>y. norm(g y) = norm y" "\<And>x. g(f x) = x" "\<And>y. f(g y) = y"
by (blast intro: isomorphisms_UNIV_UNIV [OF dimeq])
then have "continuous_on UNIV f" "continuous_on UNIV g"
using linear_continuous_on linear_linear by blast+
then show ?thesis
unfolding homeomorphic_minimal
apply(rule_tac x="\<lambda>x. b + f(x - a)" in exI)
apply(rule_tac x="\<lambda>x. a + g(x - b)" in exI)
using assms
apply (force intro: continuous_intros
continuous_on_compose2 [of _ f] continuous_on_compose2 [of _ g] simp: dist_commute dist_norm fg)
done
qed
lemma homeomorphic_spheres_gen:
fixes a :: "'a::euclidean_space" and b :: "'b::euclidean_space"
assumes "0 < r" "0 < s" "DIM('a::euclidean_space) = DIM('b::euclidean_space)"
shows "(sphere a r homeomorphic sphere b s)"
apply (rule homeomorphic_trans [OF homeomorphic_spheres homeomorphic_spheres'])
using assms apply auto
done
subsection \<open>Homeomorphism of all convex compact sets with nonempty interior\<close>
proposition
fixes S :: "'a::euclidean_space set"
assumes "compact S" and 0: "0 \<in> rel_interior S"
and star: "\<And>x. x \<in> S \<Longrightarrow> open_segment 0 x \<subseteq> rel_interior S"
shows starlike_compact_projective1_0:
"S - rel_interior S homeomorphic sphere 0 1 \<inter> affine hull S"
(is "?SMINUS homeomorphic ?SPHER")
and starlike_compact_projective2_0:
"S homeomorphic cball 0 1 \<inter> affine hull S"
(is "S homeomorphic ?CBALL")
proof -
have starI: "(u *\<^sub>R x) \<in> rel_interior S" if "x \<in> S" "0 \<le> u" "u < 1" for x u
proof (cases "x=0 \<or> u=0")
case True with 0 show ?thesis by force
next
case False with that show ?thesis
by (auto simp: in_segment intro: star [THEN subsetD])
qed
have "0 \<in> S" using assms rel_interior_subset by auto
define proj where "proj \<equiv> \<lambda>x::'a. x /\<^sub>R norm x"
have eqI: "x = y" if "proj x = proj y" "norm x = norm y" for x y
using that by (force simp: proj_def)
then have iff_eq: "\<And>x y. (proj x = proj y \<and> norm x = norm y) \<longleftrightarrow> x = y"
by blast
have projI: "x \<in> affine hull S \<Longrightarrow> proj x \<in> affine hull S" for x
by (metis \<open>0 \<in> S\<close> affine_hull_span_0 hull_inc span_mul proj_def)
have nproj1 [simp]: "x \<noteq> 0 \<Longrightarrow> norm(proj x) = 1" for x
by (simp add: proj_def)
have proj0_iff [simp]: "proj x = 0 \<longleftrightarrow> x = 0" for x
by (simp add: proj_def)
have cont_proj: "continuous_on (UNIV - {0}) proj"
unfolding proj_def by (rule continuous_intros | force)+
have proj_spherI: "\<And>x. \<lbrakk>x \<in> affine hull S; x \<noteq> 0\<rbrakk> \<Longrightarrow> proj x \<in> ?SPHER"
by (simp add: projI)
have "bounded S" "closed S"
using \<open>compact S\<close> compact_eq_bounded_closed by blast+
have inj_on_proj: "inj_on proj (S - rel_interior S)"
proof
fix x y
assume x: "x \<in> S - rel_interior S"
and y: "y \<in> S - rel_interior S" and eq: "proj x = proj y"
then have xynot: "x \<noteq> 0" "y \<noteq> 0" "x \<in> S" "y \<in> S" "x \<notin> rel_interior S" "y \<notin> rel_interior S"
using 0 by auto
consider "norm x = norm y" | "norm x < norm y" | "norm x > norm y" by linarith
then show "x = y"
proof cases
assume "norm x = norm y"
with iff_eq eq show "x = y" by blast
next
assume *: "norm x < norm y"
have "x /\<^sub>R norm x = norm x *\<^sub>R (x /\<^sub>R norm x) /\<^sub>R norm (norm x *\<^sub>R (x /\<^sub>R norm x))"
by force
then have "proj ((norm x / norm y) *\<^sub>R y) = proj x"
by (metis (no_types) divide_inverse local.proj_def eq scaleR_scaleR)
then have [simp]: "(norm x / norm y) *\<^sub>R y = x"
by (rule eqI) (simp add: \<open>y \<noteq> 0\<close>)
have no: "0 \<le> norm x / norm y" "norm x / norm y < 1"
using * by (auto simp: field_split_simps)
then show "x = y"
using starI [OF \<open>y \<in> S\<close> no] xynot by auto
next
assume *: "norm x > norm y"
have "y /\<^sub>R norm y = norm y *\<^sub>R (y /\<^sub>R norm y) /\<^sub>R norm (norm y *\<^sub>R (y /\<^sub>R norm y))"
by force
then have "proj ((norm y / norm x) *\<^sub>R x) = proj y"
by (metis (no_types) divide_inverse local.proj_def eq scaleR_scaleR)
then have [simp]: "(norm y / norm x) *\<^sub>R x = y"
by (rule eqI) (simp add: \<open>x \<noteq> 0\<close>)
have no: "0 \<le> norm y / norm x" "norm y / norm x < 1"
using * by (auto simp: field_split_simps)
then show "x = y"
using starI [OF \<open>x \<in> S\<close> no] xynot by auto
qed
qed
have "\<exists>surf. homeomorphism (S - rel_interior S) ?SPHER proj surf"
proof (rule homeomorphism_compact)
show "compact (S - rel_interior S)"
using \<open>compact S\<close> compact_rel_boundary by blast
show "continuous_on (S - rel_interior S) proj"
using 0 by (blast intro: continuous_on_subset [OF cont_proj])
show "proj ` (S - rel_interior S) = ?SPHER"
proof
show "proj ` (S - rel_interior S) \<subseteq> ?SPHER"
using 0 by (force simp: hull_inc projI intro: nproj1)
show "?SPHER \<subseteq> proj ` (S - rel_interior S)"
proof (clarsimp simp: proj_def)
fix x
assume "x \<in> affine hull S" and nox: "norm x = 1"
then have "x \<noteq> 0" by auto
obtain d where "0 < d" and dx: "(d *\<^sub>R x) \<in> rel_frontier S"
and ri: "\<And>e. \<lbrakk>0 \<le> e; e < d\<rbrakk> \<Longrightarrow> (e *\<^sub>R x) \<in> rel_interior S"
using ray_to_rel_frontier [OF \<open>bounded S\<close> 0] \<open>x \<in> affine hull S\<close> \<open>x \<noteq> 0\<close> by auto
show "x \<in> (\<lambda>x. x /\<^sub>R norm x) ` (S - rel_interior S)"
apply (rule_tac x="d *\<^sub>R x" in image_eqI)
using \<open>0 < d\<close>
using dx \<open>closed S\<close> apply (auto simp: rel_frontier_def field_split_simps nox)
done
qed
qed
qed (rule inj_on_proj)
then obtain surf where surf: "homeomorphism (S - rel_interior S) ?SPHER proj surf"
by blast
then have cont_surf: "continuous_on (proj ` (S - rel_interior S)) surf"
by (auto simp: homeomorphism_def)
have surf_nz: "\<And>x. x \<in> ?SPHER \<Longrightarrow> surf x \<noteq> 0"
by (metis "0" DiffE homeomorphism_def imageI surf)
have cont_nosp: "continuous_on (?SPHER) (\<lambda>x. norm x *\<^sub>R ((surf o proj) x))"
apply (rule continuous_intros)+
apply (rule continuous_on_subset [OF cont_proj], force)
apply (rule continuous_on_subset [OF cont_surf])
apply (force simp: homeomorphism_image1 [OF surf] dest: proj_spherI)
done
have surfpS: "\<And>x. \<lbrakk>norm x = 1; x \<in> affine hull S\<rbrakk> \<Longrightarrow> surf (proj x) \<in> S"
by (metis (full_types) DiffE \<open>0 \<in> S\<close> homeomorphism_def image_eqI norm_zero proj_spherI real_vector.scale_zero_left scaleR_one surf)
have *: "\<exists>y. norm y = 1 \<and> y \<in> affine hull S \<and> x = surf (proj y)"
if "x \<in> S" "x \<notin> rel_interior S" for x
proof -
have "proj x \<in> ?SPHER"
by (metis (full_types) "0" hull_inc proj_spherI that)
moreover have "surf (proj x) = x"
by (metis Diff_iff homeomorphism_def surf that)
ultimately show ?thesis
by (metis \<open>\<And>x. x \<in> ?SPHER \<Longrightarrow> surf x \<noteq> 0\<close> hull_inc inverse_1 local.proj_def norm_sgn projI scaleR_one sgn_div_norm that(1))
qed
have surfp_notin: "\<And>x. \<lbrakk>norm x = 1; x \<in> affine hull S\<rbrakk> \<Longrightarrow> surf (proj x) \<notin> rel_interior S"
by (metis (full_types) DiffE one_neq_zero homeomorphism_def image_eqI norm_zero proj_spherI surf)
have no_sp_im: "(\<lambda>x. norm x *\<^sub>R surf (proj x)) ` (?SPHER) = S - rel_interior S"
by (auto simp: surfpS image_def Bex_def surfp_notin *)
have inj_spher: "inj_on (\<lambda>x. norm x *\<^sub>R surf (proj x)) ?SPHER"
proof
fix x y
assume xy: "x \<in> ?SPHER" "y \<in> ?SPHER"
and eq: " norm x *\<^sub>R surf (proj x) = norm y *\<^sub>R surf (proj y)"
then have "norm x = 1" "norm y = 1" "x \<in> affine hull S" "y \<in> affine hull S"
using 0 by auto
with eq show "x = y"
by (simp add: proj_def) (metis surf xy homeomorphism_def)
qed
have co01: "compact ?SPHER"
by (simp add: compact_Int_closed)
show "?SMINUS homeomorphic ?SPHER"
apply (subst homeomorphic_sym)
apply (rule homeomorphic_compact [OF co01 cont_nosp [unfolded o_def] no_sp_im inj_spher])
done
have proj_scaleR: "\<And>a x. 0 < a \<Longrightarrow> proj (a *\<^sub>R x) = proj x"
by (simp add: proj_def)
have cont_sp0: "continuous_on (affine hull S - {0}) (surf o proj)"
apply (rule continuous_on_compose [OF continuous_on_subset [OF cont_proj]], force)
apply (rule continuous_on_subset [OF cont_surf])
using homeomorphism_image1 proj_spherI surf by fastforce
obtain B where "B>0" and B: "\<And>x. x \<in> S \<Longrightarrow> norm x \<le> B"
by (metis compact_imp_bounded \<open>compact S\<close> bounded_pos_less less_eq_real_def)
have cont_nosp: "continuous (at x within ?CBALL) (\<lambda>x. norm x *\<^sub>R surf (proj x))"
if "norm x \<le> 1" "x \<in> affine hull S" for x
proof (cases "x=0")
case True
show ?thesis using True
apply (simp add: continuous_within)
apply (rule lim_null_scaleR_bounded [where B=B])
apply (simp_all add: tendsto_norm_zero eventually_at)
apply (rule_tac x=B in exI)
using B surfpS proj_def projI apply (auto simp: \<open>B > 0\<close>)
done
next
case False
then have "\<forall>\<^sub>F x in at x. (x \<in> affine hull S - {0}) = (x \<in> affine hull S)"
apply (simp add: eventually_at)
apply (rule_tac x="norm x" in exI)
apply (auto simp: False)
done
with cont_sp0 have *: "continuous (at x within affine hull S) (\<lambda>x. surf (proj x))"
apply (simp add: continuous_on_eq_continuous_within)
apply (drule_tac x=x in bspec, force simp: False that)
apply (simp add: continuous_within Lim_transform_within_set)
done
show ?thesis
apply (rule continuous_within_subset [where s = "affine hull S", OF _ Int_lower2])
apply (rule continuous_intros *)+
done
qed
have cont_nosp2: "continuous_on ?CBALL (\<lambda>x. norm x *\<^sub>R ((surf o proj) x))"
by (simp add: continuous_on_eq_continuous_within cont_nosp)
have "norm y *\<^sub>R surf (proj y) \<in> S" if "y \<in> cball 0 1" and yaff: "y \<in> affine hull S" for y
proof (cases "y=0")
case True then show ?thesis
by (simp add: \<open>0 \<in> S\<close>)
next
case False
then have "norm y *\<^sub>R surf (proj y) = norm y *\<^sub>R surf (proj (y /\<^sub>R norm y))"
by (simp add: proj_def)
have "norm y \<le> 1" using that by simp
have "surf (proj (y /\<^sub>R norm y)) \<in> S"
apply (rule surfpS)
using proj_def projI yaff
by (auto simp: False)
then have "surf (proj y) \<in> S"
by (simp add: False proj_def)
then show "norm y *\<^sub>R surf (proj y) \<in> S"
by (metis dual_order.antisym le_less_linear norm_ge_zero rel_interior_subset scaleR_one
starI subset_eq \<open>norm y \<le> 1\<close>)
qed
moreover have "x \<in> (\<lambda>x. norm x *\<^sub>R surf (proj x)) ` (?CBALL)" if "x \<in> S" for x
proof (cases "x=0")
case True with that hull_inc show ?thesis by fastforce
next
case False
then have psp: "proj (surf (proj x)) = proj x"
by (metis homeomorphism_def hull_inc proj_spherI surf that)
have nxx: "norm x *\<^sub>R proj x = x"
by (simp add: False local.proj_def)
have affineI: "(1 / norm (surf (proj x))) *\<^sub>R x \<in> affine hull S"
by (metis \<open>0 \<in> S\<close> affine_hull_span_0 hull_inc span_clauses(4) that)
have sproj_nz: "surf (proj x) \<noteq> 0"
by (metis False proj0_iff psp)
then have "proj x = proj (proj x)"
by (metis False nxx proj_scaleR zero_less_norm_iff)
moreover have scaleproj: "\<And>a r. r *\<^sub>R proj a = (r / norm a) *\<^sub>R a"
by (simp add: divide_inverse local.proj_def)
ultimately have "(norm (surf (proj x)) / norm x) *\<^sub>R x \<notin> rel_interior S"
by (metis (no_types) sproj_nz divide_self_if hull_inc norm_eq_zero nproj1 projI psp scaleR_one surfp_notin that)
then have "(norm (surf (proj x)) / norm x) \<ge> 1"
using starI [OF that] by (meson starI [OF that] le_less_linear norm_ge_zero zero_le_divide_iff)
then have nole: "norm x \<le> norm (surf (proj x))"
by (simp add: le_divide_eq_1)
show ?thesis
apply (rule_tac x="inverse(norm(surf (proj x))) *\<^sub>R x" in image_eqI)
apply (metis (no_types, hide_lams) mult.commute scaleproj abs_inverse abs_norm_cancel divide_inverse norm_scaleR nxx positive_imp_inverse_positive proj_scaleR psp sproj_nz zero_less_norm_iff)
apply (auto simp: field_split_simps nole affineI)
done
qed
ultimately have im_cball: "(\<lambda>x. norm x *\<^sub>R surf (proj x)) ` ?CBALL = S"
by blast
have inj_cball: "inj_on (\<lambda>x. norm x *\<^sub>R surf (proj x)) ?CBALL"
proof
fix x y
assume "x \<in> ?CBALL" "y \<in> ?CBALL"
and eq: "norm x *\<^sub>R surf (proj x) = norm y *\<^sub>R surf (proj y)"
then have x: "x \<in> affine hull S" and y: "y \<in> affine hull S"
using 0 by auto
show "x = y"
proof (cases "x=0 \<or> y=0")
case True then show "x = y" using eq proj_spherI surf_nz x y by force
next
case False
with x y have speq: "surf (proj x) = surf (proj y)"
by (metis eq homeomorphism_apply2 proj_scaleR proj_spherI surf zero_less_norm_iff)
then have "norm x = norm y"
by (metis \<open>x \<in> affine hull S\<close> \<open>y \<in> affine hull S\<close> eq proj_spherI real_vector.scale_cancel_right surf_nz)
moreover have "proj x = proj y"
by (metis (no_types) False speq homeomorphism_apply2 proj_spherI surf x y)
ultimately show "x = y"
using eq eqI by blast
qed
qed
have co01: "compact ?CBALL"
by (simp add: compact_Int_closed)
show "S homeomorphic ?CBALL"
apply (subst homeomorphic_sym)
apply (rule homeomorphic_compact [OF co01 cont_nosp2 [unfolded o_def] im_cball inj_cball])
done
qed
corollary
fixes S :: "'a::euclidean_space set"
assumes "compact S" and a: "a \<in> rel_interior S"
and star: "\<And>x. x \<in> S \<Longrightarrow> open_segment a x \<subseteq> rel_interior S"
shows starlike_compact_projective1:
"S - rel_interior S homeomorphic sphere a 1 \<inter> affine hull S"
and starlike_compact_projective2:
"S homeomorphic cball a 1 \<inter> affine hull S"
proof -
have 1: "compact ((+) (-a) ` S)" by (meson assms compact_translation)
have 2: "0 \<in> rel_interior ((+) (-a) ` S)"
using a rel_interior_translation [of "- a" S] by (simp cong: image_cong_simp)
have 3: "open_segment 0 x \<subseteq> rel_interior ((+) (-a) ` S)" if "x \<in> ((+) (-a) ` S)" for x
proof -
have "x+a \<in> S" using that by auto
then have "open_segment a (x+a) \<subseteq> rel_interior S" by (metis star)
then show ?thesis using open_segment_translation [of a 0 x]
using rel_interior_translation [of "- a" S] by (fastforce simp add: ac_simps image_iff cong: image_cong_simp)
qed
have "S - rel_interior S homeomorphic ((+) (-a) ` S) - rel_interior ((+) (-a) ` S)"
by (metis rel_interior_translation translation_diff homeomorphic_translation)
also have "... homeomorphic sphere 0 1 \<inter> affine hull ((+) (-a) ` S)"
by (rule starlike_compact_projective1_0 [OF 1 2 3])
also have "... = (+) (-a) ` (sphere a 1 \<inter> affine hull S)"
by (metis affine_hull_translation left_minus sphere_translation translation_Int)
also have "... homeomorphic sphere a 1 \<inter> affine hull S"
using homeomorphic_translation homeomorphic_sym by blast
finally show "S - rel_interior S homeomorphic sphere a 1 \<inter> affine hull S" .
have "S homeomorphic ((+) (-a) ` S)"
by (metis homeomorphic_translation)
also have "... homeomorphic cball 0 1 \<inter> affine hull ((+) (-a) ` S)"
by (rule starlike_compact_projective2_0 [OF 1 2 3])
also have "... = (+) (-a) ` (cball a 1 \<inter> affine hull S)"
by (metis affine_hull_translation left_minus cball_translation translation_Int)
also have "... homeomorphic cball a 1 \<inter> affine hull S"
using homeomorphic_translation homeomorphic_sym by blast
finally show "S homeomorphic cball a 1 \<inter> affine hull S" .
qed
corollary starlike_compact_projective_special:
assumes "compact S"
and cb01: "cball (0::'a::euclidean_space) 1 \<subseteq> S"
and scale: "\<And>x u. \<lbrakk>x \<in> S; 0 \<le> u; u < 1\<rbrakk> \<Longrightarrow> u *\<^sub>R x \<in> S - frontier S"
shows "S homeomorphic (cball (0::'a::euclidean_space) 1)"
proof -
have "ball 0 1 \<subseteq> interior S"
using cb01 interior_cball interior_mono by blast
then have 0: "0 \<in> rel_interior S"
by (meson centre_in_ball subsetD interior_subset_rel_interior le_numeral_extra(2) not_le)
have [simp]: "affine hull S = UNIV"
using \<open>ball 0 1 \<subseteq> interior S\<close> by (auto intro!: affine_hull_nonempty_interior)
have star: "open_segment 0 x \<subseteq> rel_interior S" if "x \<in> S" for x
proof
fix p assume "p \<in> open_segment 0 x"
then obtain u where "x \<noteq> 0" and u: "0 \<le> u" "u < 1" and p: "u *\<^sub>R x = p"
by (auto simp: in_segment)
then show "p \<in> rel_interior S"
using scale [OF that u] closure_subset frontier_def interior_subset_rel_interior by fastforce
qed
show ?thesis
using starlike_compact_projective2_0 [OF \<open>compact S\<close> 0 star] by simp
qed
lemma homeomorphic_convex_lemma:
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
assumes "convex S" "compact S" "convex T" "compact T"
and affeq: "aff_dim S = aff_dim T"
shows "(S - rel_interior S) homeomorphic (T - rel_interior T) \<and>
S homeomorphic T"
proof (cases "rel_interior S = {} \<or> rel_interior T = {}")
case True
then show ?thesis
by (metis Diff_empty affeq \<open>convex S\<close> \<open>convex T\<close> aff_dim_empty homeomorphic_empty rel_interior_eq_empty aff_dim_empty)
next
case False
then obtain a b where a: "a \<in> rel_interior S" and b: "b \<in> rel_interior T" by auto
have starS: "\<And>x. x \<in> S \<Longrightarrow> open_segment a x \<subseteq> rel_interior S"
using rel_interior_closure_convex_segment
a \<open>convex S\<close> closure_subset subsetCE by blast
have starT: "\<And>x. x \<in> T \<Longrightarrow> open_segment b x \<subseteq> rel_interior T"
using rel_interior_closure_convex_segment
b \<open>convex T\<close> closure_subset subsetCE by blast
let ?aS = "(+) (-a) ` S" and ?bT = "(+) (-b) ` T"
have 0: "0 \<in> affine hull ?aS" "0 \<in> affine hull ?bT"
by (metis a b subsetD hull_inc image_eqI left_minus rel_interior_subset)+
have subs: "subspace (span ?aS)" "subspace (span ?bT)"
by (rule subspace_span)+
moreover
have "dim (span ((+) (- a) ` S)) = dim (span ((+) (- b) ` T))"
by (metis 0 aff_dim_translation_eq aff_dim_zero affeq dim_span nat_int)
ultimately obtain f g where "linear f" "linear g"
and fim: "f ` span ?aS = span ?bT"
and gim: "g ` span ?bT = span ?aS"
and fno: "\<And>x. x \<in> span ?aS \<Longrightarrow> norm(f x) = norm x"
and gno: "\<And>x. x \<in> span ?bT \<Longrightarrow> norm(g x) = norm x"
and gf: "\<And>x. x \<in> span ?aS \<Longrightarrow> g(f x) = x"
and fg: "\<And>x. x \<in> span ?bT \<Longrightarrow> f(g x) = x"
by (rule isometries_subspaces) blast
have [simp]: "continuous_on A f" for A
using \<open>linear f\<close> linear_conv_bounded_linear linear_continuous_on by blast
have [simp]: "continuous_on B g" for B
using \<open>linear g\<close> linear_conv_bounded_linear linear_continuous_on by blast
have eqspanS: "affine hull ?aS = span ?aS"
by (metis a affine_hull_span_0 subsetD hull_inc image_eqI left_minus rel_interior_subset)
have eqspanT: "affine hull ?bT = span ?bT"
by (metis b affine_hull_span_0 subsetD hull_inc image_eqI left_minus rel_interior_subset)
have "S homeomorphic cball a 1 \<inter> affine hull S"
by (rule starlike_compact_projective2 [OF \<open>compact S\<close> a starS])
also have "... homeomorphic (+) (-a) ` (cball a 1 \<inter> affine hull S)"
by (metis homeomorphic_translation)
also have "... = cball 0 1 \<inter> (+) (-a) ` (affine hull S)"
by (auto simp: dist_norm)
also have "... = cball 0 1 \<inter> span ?aS"
using eqspanS affine_hull_translation by blast
also have "... homeomorphic cball 0 1 \<inter> span ?bT"
proof (rule homeomorphicI [where f=f and g=g])
show fim1: "f ` (cball 0 1 \<inter> span ?aS) = cball 0 1 \<inter> span ?bT"
apply (rule subset_antisym)
using fim fno apply (force simp:, clarify)
by (metis IntI fg gim gno image_eqI mem_cball_0)
show "g ` (cball 0 1 \<inter> span ?bT) = cball 0 1 \<inter> span ?aS"
apply (rule subset_antisym)
using gim gno apply (force simp:, clarify)
by (metis IntI fim1 gf image_eqI)
qed (auto simp: fg gf)
also have "... = cball 0 1 \<inter> (+) (-b) ` (affine hull T)"
using eqspanT affine_hull_translation by blast
also have "... = (+) (-b) ` (cball b 1 \<inter> affine hull T)"
by (auto simp: dist_norm)
also have "... homeomorphic (cball b 1 \<inter> affine hull T)"
by (metis homeomorphic_translation homeomorphic_sym)
also have "... homeomorphic T"
by (metis starlike_compact_projective2 [OF \<open>compact T\<close> b starT] homeomorphic_sym)
finally have 1: "S homeomorphic T" .
have "S - rel_interior S homeomorphic sphere a 1 \<inter> affine hull S"
by (rule starlike_compact_projective1 [OF \<open>compact S\<close> a starS])
also have "... homeomorphic (+) (-a) ` (sphere a 1 \<inter> affine hull S)"
by (metis homeomorphic_translation)
also have "... = sphere 0 1 \<inter> (+) (-a) ` (affine hull S)"
by (auto simp: dist_norm)
also have "... = sphere 0 1 \<inter> span ?aS"
using eqspanS affine_hull_translation by blast
also have "... homeomorphic sphere 0 1 \<inter> span ?bT"
proof (rule homeomorphicI [where f=f and g=g])
show fim1: "f ` (sphere 0 1 \<inter> span ?aS) = sphere 0 1 \<inter> span ?bT"
apply (rule subset_antisym)
using fim fno apply (force simp:, clarify)
by (metis IntI fg gim gno image_eqI mem_sphere_0)
show "g ` (sphere 0 1 \<inter> span ?bT) = sphere 0 1 \<inter> span ?aS"
apply (rule subset_antisym)
using gim gno apply (force simp:, clarify)
by (metis IntI fim1 gf image_eqI)
qed (auto simp: fg gf)
also have "... = sphere 0 1 \<inter> (+) (-b) ` (affine hull T)"
using eqspanT affine_hull_translation by blast
also have "... = (+) (-b) ` (sphere b 1 \<inter> affine hull T)"
by (auto simp: dist_norm)
also have "... homeomorphic (sphere b 1 \<inter> affine hull T)"
by (metis homeomorphic_translation homeomorphic_sym)
also have "... homeomorphic T - rel_interior T"
by (metis starlike_compact_projective1 [OF \<open>compact T\<close> b starT] homeomorphic_sym)
finally have 2: "S - rel_interior S homeomorphic T - rel_interior T" .
show ?thesis
using 1 2 by blast
qed
lemma homeomorphic_convex_compact_sets:
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
assumes "convex S" "compact S" "convex T" "compact T"
and affeq: "aff_dim S = aff_dim T"
shows "S homeomorphic T"
using homeomorphic_convex_lemma [OF assms] assms
by (auto simp: rel_frontier_def)
lemma homeomorphic_rel_frontiers_convex_bounded_sets:
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
assumes "convex S" "bounded S" "convex T" "bounded T"
and affeq: "aff_dim S = aff_dim T"
shows "rel_frontier S homeomorphic rel_frontier T"
using assms homeomorphic_convex_lemma [of "closure S" "closure T"]
by (simp add: rel_frontier_def convex_rel_interior_closure)
subsection\<open>Homeomorphisms between punctured spheres and affine sets\<close>
text\<open>Including the famous stereoscopic projection of the 3-D sphere to the complex plane\<close>
text\<open>The special case with centre 0 and radius 1\<close>
lemma homeomorphic_punctured_affine_sphere_affine_01:
assumes "b \<in> sphere 0 1" "affine T" "0 \<in> T" "b \<in> T" "affine p"
and affT: "aff_dim T = aff_dim p + 1"
shows "(sphere 0 1 \<inter> T) - {b} homeomorphic p"
proof -
have [simp]: "norm b = 1" "b\<bullet>b = 1"
using assms by (auto simp: norm_eq_1)
have [simp]: "T \<inter> {v. b\<bullet>v = 0} \<noteq> {}"
using \<open>0 \<in> T\<close> by auto
have [simp]: "\<not> T \<subseteq> {v. b\<bullet>v = 0}"
using \<open>norm b = 1\<close> \<open>b \<in> T\<close> by auto
define f where "f \<equiv> \<lambda>x. 2 *\<^sub>R b + (2 / (1 - b\<bullet>x)) *\<^sub>R (x - b)"
define g where "g \<equiv> \<lambda>y. b + (4 / (norm y ^ 2 + 4)) *\<^sub>R (y - 2 *\<^sub>R b)"
have [simp]: "\<And>x. \<lbrakk>x \<in> T; b\<bullet>x = 0\<rbrakk> \<Longrightarrow> f (g x) = x"
unfolding f_def g_def by (simp add: algebra_simps field_split_simps add_nonneg_eq_0_iff)
have no: "(norm (f x))\<^sup>2 = 4 * (1 + b \<bullet> x) / (1 - b \<bullet> x)"
if "norm x = 1" and "b \<bullet> x \<noteq> 1" for x
using that
apply (simp flip: dot_square_norm add: norm_eq_1 nonzero_eq_divide_eq)
apply (simp add: f_def vector_add_divide_simps inner_simps)
apply (use sum_sqs_eq [of 1 "b \<bullet> x"]
in \<open>auto simp add: field_split_simps inner_commute\<close>)
done
have [simp]: "\<And>u::real. 8 + u * (u * 8) = u * 16 \<longleftrightarrow> u=1"
by algebra
have [simp]: "\<And>x. \<lbrakk>norm x = 1; b \<bullet> x \<noteq> 1\<rbrakk> \<Longrightarrow> g (f x) = x"
unfolding g_def no by (auto simp: f_def field_split_simps)
have [simp]: "norm (g x) = 1" if "x \<in> T" and "b \<bullet> x = 0" for x
using that
apply (simp only: g_def)
apply (rule power2_eq_imp_eq)
apply (simp_all add: dot_square_norm [symmetric] divide_simps vector_add_divide_simps)
apply (simp add: algebra_simps inner_commute)
done
have [simp]: "b \<bullet> g x \<noteq> 1" if "x \<in> T" and "b \<bullet> x = 0" for x
using that unfolding g_def
apply (simp_all add: dot_square_norm [symmetric] divide_simps vector_add_divide_simps add_nonneg_eq_0_iff)
apply (auto simp: algebra_simps)
done
have "subspace T"
by (simp add: assms subspace_affine)
have [simp]: "\<And>x. \<lbrakk>x \<in> T; b \<bullet> x = 0\<rbrakk> \<Longrightarrow> g x \<in> T"
unfolding g_def
by (blast intro: \<open>subspace T\<close> \<open>b \<in> T\<close> subspace_add subspace_mul subspace_diff)
have "f ` {x. norm x = 1 \<and> b\<bullet>x \<noteq> 1} \<subseteq> {x. b\<bullet>x = 0}"
unfolding f_def using \<open>norm b = 1\<close> norm_eq_1
by (force simp: field_simps inner_add_right inner_diff_right)
moreover have "f ` T \<subseteq> T"
unfolding f_def using assms \<open>subspace T\<close>
by (auto simp add: inner_add_right inner_diff_right mem_affine_3_minus subspace_mul)
moreover have "{x. b\<bullet>x = 0} \<inter> T \<subseteq> f ` ({x. norm x = 1 \<and> b\<bullet>x \<noteq> 1} \<inter> T)"
apply clarify
apply (rule_tac x = "g x" in image_eqI, auto)
done
ultimately have imf: "f ` ({x. norm x = 1 \<and> b\<bullet>x \<noteq> 1} \<inter> T) = {x. b\<bullet>x = 0} \<inter> T"
by blast
have no4: "\<And>y. b\<bullet>y = 0 \<Longrightarrow> norm ((y\<bullet>y + 4) *\<^sub>R b + 4 *\<^sub>R (y - 2 *\<^sub>R b)) = y\<bullet>y + 4"
apply (rule power2_eq_imp_eq)
apply (simp_all add: dot_square_norm [symmetric])
apply (auto simp: power2_eq_square algebra_simps inner_commute)
done
have [simp]: "\<And>x. \<lbrakk>norm x = 1; b \<bullet> x \<noteq> 1\<rbrakk> \<Longrightarrow> b \<bullet> f x = 0"
by (simp add: f_def algebra_simps field_split_simps)
have [simp]: "\<And>x. \<lbrakk>x \<in> T; norm x = 1; b \<bullet> x \<noteq> 1\<rbrakk> \<Longrightarrow> f x \<in> T"
unfolding f_def
by (blast intro: \<open>subspace T\<close> \<open>b \<in> T\<close> subspace_add subspace_mul subspace_diff)
have "g ` {x. b\<bullet>x = 0} \<subseteq> {x. norm x = 1 \<and> b\<bullet>x \<noteq> 1}"
unfolding g_def
apply (clarsimp simp: no4 vector_add_divide_simps divide_simps add_nonneg_eq_0_iff dot_square_norm [symmetric])
apply (auto simp: algebra_simps)
done
moreover have "g ` T \<subseteq> T"
unfolding g_def
by (blast intro: \<open>subspace T\<close> \<open>b \<in> T\<close> subspace_add subspace_mul subspace_diff)
moreover have "{x. norm x = 1 \<and> b\<bullet>x \<noteq> 1} \<inter> T \<subseteq> g ` ({x. b\<bullet>x = 0} \<inter> T)"
apply clarify
apply (rule_tac x = "f x" in image_eqI, auto)
done
ultimately have img: "g ` ({x. b\<bullet>x = 0} \<inter> T) = {x. norm x = 1 \<and> b\<bullet>x \<noteq> 1} \<inter> T"
by blast
have aff: "affine ({x. b\<bullet>x = 0} \<inter> T)"
by (blast intro: affine_hyperplane assms)
have contf: "continuous_on ({x. norm x = 1 \<and> b\<bullet>x \<noteq> 1} \<inter> T) f"
unfolding f_def by (rule continuous_intros | force)+
have contg: "continuous_on ({x. b\<bullet>x = 0} \<inter> T) g"
unfolding g_def by (rule continuous_intros | force simp: add_nonneg_eq_0_iff)+
have "(sphere 0 1 \<inter> T) - {b} = {x. norm x = 1 \<and> (b\<bullet>x \<noteq> 1)} \<inter> T"
using \<open>norm b = 1\<close> by (auto simp: norm_eq_1) (metis vector_eq \<open>b\<bullet>b = 1\<close>)
also have "... homeomorphic {x. b\<bullet>x = 0} \<inter> T"
by (rule homeomorphicI [OF imf img contf contg]) auto
also have "... homeomorphic p"
apply (rule homeomorphic_affine_sets [OF aff \<open>affine p\<close>])
apply (simp add: Int_commute aff_dim_affine_Int_hyperplane [OF \<open>affine T\<close>] affT)
done
finally show ?thesis .
qed
theorem homeomorphic_punctured_affine_sphere_affine:
fixes a :: "'a :: euclidean_space"
assumes "0 < r" "b \<in> sphere a r" "affine T" "a \<in> T" "b \<in> T" "affine p"
and aff: "aff_dim T = aff_dim p + 1"
shows "(sphere a r \<inter> T) - {b} homeomorphic p"
proof -
have "a \<noteq> b" using assms by auto
then have inj: "inj (\<lambda>x::'a. x /\<^sub>R norm (a - b))"
by (simp add: inj_on_def)
have "((sphere a r \<inter> T) - {b}) homeomorphic
(+) (-a) ` ((sphere a r \<inter> T) - {b})"
by (rule homeomorphic_translation)
also have "... homeomorphic (*\<^sub>R) (inverse r) ` (+) (- a) ` (sphere a r \<inter> T - {b})"
by (metis \<open>0 < r\<close> homeomorphic_scaling inverse_inverse_eq inverse_zero less_irrefl)
also have "... = sphere 0 1 \<inter> ((*\<^sub>R) (inverse r) ` (+) (- a) ` T) - {(b - a) /\<^sub>R r}"
using assms by (auto simp: dist_norm norm_minus_commute divide_simps)
also have "... homeomorphic p"
apply (rule homeomorphic_punctured_affine_sphere_affine_01)
using assms affine_translation [symmetric, of "- a"] aff_dim_translation_eq [of "- a"]
apply (auto simp: dist_norm norm_minus_commute affine_scaling inj)
done
finally show ?thesis .
qed
corollary homeomorphic_punctured_sphere_affine:
fixes a :: "'a :: euclidean_space"
assumes "0 < r" and b: "b \<in> sphere a r"
and "affine T" and affS: "aff_dim T + 1 = DIM('a)"
shows "(sphere a r - {b}) homeomorphic T"
using homeomorphic_punctured_affine_sphere_affine [of r b a UNIV T] assms by auto
corollary homeomorphic_punctured_sphere_hyperplane:
fixes a :: "'a :: euclidean_space"
assumes "0 < r" and b: "b \<in> sphere a r"
and "c \<noteq> 0"
shows "(sphere a r - {b}) homeomorphic {x::'a. c \<bullet> x = d}"
apply (rule homeomorphic_punctured_sphere_affine)
using assms
apply (auto simp: affine_hyperplane of_nat_diff)
done
proposition homeomorphic_punctured_sphere_affine_gen:
fixes a :: "'a :: euclidean_space"
assumes "convex S" "bounded S" and a: "a \<in> rel_frontier S"
and "affine T" and affS: "aff_dim S = aff_dim T + 1"
shows "rel_frontier S - {a} homeomorphic T"
proof -
obtain U :: "'a set" where "affine U" "convex U" and affdS: "aff_dim U = aff_dim S"
using choose_affine_subset [OF affine_UNIV aff_dim_geq]
by (meson aff_dim_affine_hull affine_affine_hull affine_imp_convex)
have "S \<noteq> {}" using assms by auto
then obtain z where "z \<in> U"
by (metis aff_dim_negative_iff equals0I affdS)
then have bne: "ball z 1 \<inter> U \<noteq> {}" by force
then have [simp]: "aff_dim(ball z 1 \<inter> U) = aff_dim U"
using aff_dim_convex_Int_open [OF \<open>convex U\<close> open_ball]
by (fastforce simp add: Int_commute)
have "rel_frontier S homeomorphic rel_frontier (ball z 1 \<inter> U)"
by (rule homeomorphic_rel_frontiers_convex_bounded_sets)
(auto simp: \<open>affine U\<close> affine_imp_convex convex_Int affdS assms)
also have "... = sphere z 1 \<inter> U"
using convex_affine_rel_frontier_Int [of "ball z 1" U]
by (simp add: \<open>affine U\<close> bne)
finally have "rel_frontier S homeomorphic sphere z 1 \<inter> U" .
then obtain h k where him: "h ` rel_frontier S = sphere z 1 \<inter> U"
and kim: "k ` (sphere z 1 \<inter> U) = rel_frontier S"
and hcon: "continuous_on (rel_frontier S) h"
and kcon: "continuous_on (sphere z 1 \<inter> U) k"
and kh: "\<And>x. x \<in> rel_frontier S \<Longrightarrow> k(h(x)) = x"
and hk: "\<And>y. y \<in> sphere z 1 \<inter> U \<Longrightarrow> h(k(y)) = y"
unfolding homeomorphic_def homeomorphism_def by auto
have "rel_frontier S - {a} homeomorphic (sphere z 1 \<inter> U) - {h a}"
proof (rule homeomorphicI)
show h: "h ` (rel_frontier S - {a}) = sphere z 1 \<inter> U - {h a}"
using him a kh by auto metis
show "k ` (sphere z 1 \<inter> U - {h a}) = rel_frontier S - {a}"
by (force simp: h [symmetric] image_comp o_def kh)
qed (auto intro: continuous_on_subset hcon kcon simp: kh hk)
also have "... homeomorphic T"
by (rule homeomorphic_punctured_affine_sphere_affine)
(use a him in \<open>auto simp: affS affdS \<open>affine T\<close> \<open>affine U\<close> \<open>z \<in> U\<close>\<close>)
finally show ?thesis .
qed
text\<open> When dealing with AR, ANR and ANR later, it's useful to know that every set
is homeomorphic to a closed subset of a convex set, and
if the set is locally compact we can take the convex set to be the universe.\<close>
proposition homeomorphic_closedin_convex:
fixes S :: "'m::euclidean_space set"
assumes "aff_dim S < DIM('n)"
obtains U and T :: "'n::euclidean_space set"
where "convex U" "U \<noteq> {}" "closedin (top_of_set U) T"
"S homeomorphic T"
proof (cases "S = {}")
case True then show ?thesis
by (rule_tac U=UNIV and T="{}" in that) auto
next
case False
then obtain a where "a \<in> S" by auto
obtain i::'n where i: "i \<in> Basis" "i \<noteq> 0"
using SOME_Basis Basis_zero by force
have "0 \<in> affine hull ((+) (- a) ` S)"
by (simp add: \<open>a \<in> S\<close> hull_inc)
then have "dim ((+) (- a) ` S) = aff_dim ((+) (- a) ` S)"
by (simp add: aff_dim_zero)
also have "... < DIM('n)"
by (simp add: aff_dim_translation_eq_subtract assms cong: image_cong_simp)
finally have dd: "dim ((+) (- a) ` S) < DIM('n)"
by linarith
have span: "span {x. i \<bullet> x = 0} = {x. i \<bullet> x = 0}"
using span_eq_iff [symmetric, of "{x. i \<bullet> x = 0}"] subspace_hyperplane [of i] by simp
have "dim ((+) (- a) ` S) \<le> dim {x. i \<bullet> x = 0}"
using dd by (simp add: dim_hyperplane [OF \<open>i \<noteq> 0\<close>])
then obtain T where "subspace T" and Tsub: "T \<subseteq> {x. i \<bullet> x = 0}"
and dimT: "dim T = dim ((+) (- a) ` S)"
by (rule choose_subspace_of_subspace) (simp add: span)
have "subspace (span ((+) (- a) ` S))"
using subspace_span by blast
then obtain h k where "linear h" "linear k"
and heq: "h ` span ((+) (- a) ` S) = T"
and keq:"k ` T = span ((+) (- a) ` S)"
and hinv [simp]: "\<And>x. x \<in> span ((+) (- a) ` S) \<Longrightarrow> k(h x) = x"
and kinv [simp]: "\<And>x. x \<in> T \<Longrightarrow> h(k x) = x"
apply (rule isometries_subspaces [OF _ \<open>subspace T\<close>])
apply (auto simp: dimT)
done
have hcont: "continuous_on A h" and kcont: "continuous_on B k" for A B
using \<open>linear h\<close> \<open>linear k\<close> linear_continuous_on linear_conv_bounded_linear by blast+
have ihhhh[simp]: "\<And>x. x \<in> S \<Longrightarrow> i \<bullet> h (x - a) = 0"
using Tsub [THEN subsetD] heq span_superset by fastforce
have "sphere 0 1 - {i} homeomorphic {x. i \<bullet> x = 0}"
apply (rule homeomorphic_punctured_sphere_affine)
using i
apply (auto simp: affine_hyperplane)
by (metis DIM_positive Suc_eq_plus1 add.left_neutral diff_add_cancel not_le not_less_eq_eq of_nat_1 of_nat_diff)
then obtain f g where fg: "homeomorphism (sphere 0 1 - {i}) {x. i \<bullet> x = 0} f g"
by (force simp: homeomorphic_def)
have "h ` (+) (- a) ` S \<subseteq> T"
using heq span_superset span_linear_image by blast
then have "g ` h ` (+) (- a) ` S \<subseteq> g ` {x. i \<bullet> x = 0}"
using Tsub by (simp add: image_mono)
also have "... \<subseteq> sphere 0 1 - {i}"
by (simp add: fg [unfolded homeomorphism_def])
finally have gh_sub_sph: "(g \<circ> h) ` (+) (- a) ` S \<subseteq> sphere 0 1 - {i}"
by (metis image_comp)
then have gh_sub_cb: "(g \<circ> h) ` (+) (- a) ` S \<subseteq> cball 0 1"
by (metis Diff_subset order_trans sphere_cball)
have [simp]: "\<And>u. u \<in> S \<Longrightarrow> norm (g (h (u - a))) = 1"
using gh_sub_sph [THEN subsetD] by (auto simp: o_def)
have ghcont: "continuous_on ((\<lambda>x. x - a) ` S) (\<lambda>x. g (h x))"
apply (rule continuous_on_compose2 [OF homeomorphism_cont2 [OF fg] hcont], force)
done
have kfcont: "continuous_on ((\<lambda>x. g (h (x - a))) ` S) (\<lambda>x. k (f x))"
apply (rule continuous_on_compose2 [OF kcont])
using homeomorphism_cont1 [OF fg] gh_sub_sph apply (force intro: continuous_on_subset, blast)
done
have "S homeomorphic (+) (- a) ` S"
by (fact homeomorphic_translation)
also have "\<dots> homeomorphic (g \<circ> h) ` (+) (- a) ` S"
apply (simp add: homeomorphic_def homeomorphism_def cong: image_cong_simp)
apply (rule_tac x="g \<circ> h" in exI)
apply (rule_tac x="k \<circ> f" in exI)
apply (auto simp: ghcont kfcont span_base homeomorphism_apply2 [OF fg] image_comp cong: image_cong_simp)
done
finally have Shom: "S homeomorphic (\<lambda>x. g (h x)) ` (\<lambda>x. x - a) ` S"
by (simp cong: image_cong_simp)
show ?thesis
apply (rule_tac U = "ball 0 1 \<union> image (g o h) ((+) (- a) ` S)"
and T = "image (g o h) ((+) (- a) ` S)"
in that)
apply (rule convex_intermediate_ball [of 0 1], force)
using gh_sub_cb apply force
apply force
apply (simp add: closedin_closed)
apply (rule_tac x="sphere 0 1" in exI)
apply (auto simp: Shom cong: image_cong_simp)
done
qed
subsection\<open>Locally compact sets in an open set\<close>
text\<open> Locally compact sets are closed in an open set and are homeomorphic
to an absolutely closed set if we have one more dimension to play with.\<close>
lemma locally_compact_open_Int_closure:
fixes S :: "'a :: metric_space set"
assumes "locally compact S"
obtains T where "open T" "S = T \<inter> closure S"
proof -
have "\<forall>x\<in>S. \<exists>T v u. u = S \<inter> T \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> S \<and> open T \<and> compact v"
by (metis assms locally_compact openin_open)
then obtain t v where
tv: "\<And>x. x \<in> S
\<Longrightarrow> v x \<subseteq> S \<and> open (t x) \<and> compact (v x) \<and> (\<exists>u. x \<in> u \<and> u \<subseteq> v x \<and> u = S \<inter> t x)"
by metis
then have o: "open (\<Union>(t ` S))"
by blast
have "S = \<Union> (v ` S)"
using tv by blast
also have "... = \<Union>(t ` S) \<inter> closure S"
proof
show "\<Union>(v ` S) \<subseteq> \<Union>(t ` S) \<inter> closure S"
apply safe
apply (metis Int_iff subsetD UN_iff tv)
apply (simp add: closure_def rev_subsetD tv)
done
have "t x \<inter> closure S \<subseteq> v x" if "x \<in> S" for x
proof -
have "t x \<inter> closure S \<subseteq> closure (t x \<inter> S)"
by (simp add: open_Int_closure_subset that tv)
also have "... \<subseteq> v x"
by (metis Int_commute closure_minimal compact_imp_closed that tv)
finally show ?thesis .
qed
then show "\<Union>(t ` S) \<inter> closure S \<subseteq> \<Union>(v ` S)"
by blast
qed
finally have e: "S = \<Union>(t ` S) \<inter> closure S" .
show ?thesis
by (rule that [OF o e])
qed
lemma locally_compact_closedin_open:
fixes S :: "'a :: metric_space set"
assumes "locally compact S"
obtains T where "open T" "closedin (top_of_set T) S"
by (metis locally_compact_open_Int_closure [OF assms] closed_closure closedin_closed_Int)
lemma locally_compact_homeomorphism_projection_closed:
assumes "locally compact S"
obtains T and f :: "'a \<Rightarrow> 'a :: euclidean_space \<times> 'b :: euclidean_space"
where "closed T" "homeomorphism S T f fst"
proof (cases "closed S")
case True
then show ?thesis
apply (rule_tac T = "S \<times> {0}" and f = "\<lambda>x. (x, 0)" in that)
apply (auto simp: closed_Times homeomorphism_def continuous_intros)
done
next
case False
obtain U where "open U" and US: "U \<inter> closure S = S"
by (metis locally_compact_open_Int_closure [OF assms])
with False have Ucomp: "-U \<noteq> {}"
using closure_eq by auto
have [simp]: "closure (- U) = -U"
by (simp add: \<open>open U\<close> closed_Compl)
define f :: "'a \<Rightarrow> 'a \<times> 'b" where "f \<equiv> \<lambda>x. (x, One /\<^sub>R setdist {x} (- U))"
have "continuous_on U (\<lambda>x. (x, One /\<^sub>R setdist {x} (- U)))"
apply (intro continuous_intros continuous_on_setdist)
by (simp add: Ucomp setdist_eq_0_sing_1)
then have homU: "homeomorphism U (f`U) f fst"
by (auto simp: f_def homeomorphism_def image_iff continuous_intros)
have cloS: "closedin (top_of_set U) S"
by (metis US closed_closure closedin_closed_Int)
have cont: "isCont ((\<lambda>x. setdist {x} (- U)) o fst) z" for z :: "'a \<times> 'b"
by (rule continuous_at_compose continuous_intros continuous_at_setdist)+
have setdist1D: "setdist {a} (- U) *\<^sub>R b = One \<Longrightarrow> setdist {a} (- U) \<noteq> 0" for a::'a and b::'b
by force
have *: "r *\<^sub>R b = One \<Longrightarrow> b = (1 / r) *\<^sub>R One" for r and b::'b
by (metis One_non_0 nonzero_divide_eq_eq real_vector.scale_eq_0_iff real_vector.scale_scale scaleR_one)
have "f ` U = (\<lambda>z. (setdist {fst z} (- U) *\<^sub>R snd z)) -` {One}"
apply (auto simp: f_def setdist_eq_0_sing_1 field_simps Ucomp)
apply (rule_tac x=a in image_eqI)
apply (auto simp: * setdist_eq_0_sing_1 dest: setdist1D)
done
then have clfU: "closed (f ` U)"
apply (rule ssubst)
apply (rule continuous_closed_vimage)
apply (auto intro: continuous_intros cont [unfolded o_def])
done
have "closed (f ` S)"
apply (rule closedin_closed_trans [OF _ clfU])
apply (rule homeomorphism_imp_closed_map [OF homU cloS])
done
then show ?thesis
apply (rule that)
apply (rule homeomorphism_of_subsets [OF homU])
using US apply auto
done
qed
lemma locally_compact_closed_Int_open:
fixes S :: "'a :: euclidean_space set"
shows
"locally compact S \<longleftrightarrow> (\<exists>U u. closed U \<and> open u \<and> S = U \<inter> u)"
by (metis closed_closure closed_imp_locally_compact inf_commute locally_compact_Int locally_compact_open_Int_closure open_imp_locally_compact)
lemma lowerdim_embeddings:
assumes "DIM('a) < DIM('b)"
obtains f :: "'a::euclidean_space*real \<Rightarrow> 'b::euclidean_space"
and g :: "'b \<Rightarrow> 'a*real"
and j :: 'b
where "linear f" "linear g" "\<And>z. g (f z) = z" "j \<in> Basis" "\<And>x. f(x,0) \<bullet> j = 0"
proof -
let ?B = "Basis :: ('a*real) set"
have b01: "(0,1) \<in> ?B"
by (simp add: Basis_prod_def)
have "DIM('a * real) \<le> DIM('b)"
by (simp add: Suc_leI assms)
then obtain basf :: "'a*real \<Rightarrow> 'b" where sbf: "basf ` ?B \<subseteq> Basis" and injbf: "inj_on basf Basis"
by (metis finite_Basis card_le_inj)
define basg:: "'b \<Rightarrow> 'a * real" where
"basg \<equiv> \<lambda>i. if i \<in> basf ` Basis then inv_into Basis basf i else (0,1)"
have bgf[simp]: "basg (basf i) = i" if "i \<in> Basis" for i
using inv_into_f_f injbf that by (force simp: basg_def)
have sbg: "basg ` Basis \<subseteq> ?B"
by (force simp: basg_def injbf b01)
define f :: "'a*real \<Rightarrow> 'b" where "f \<equiv> \<lambda>u. \<Sum>j\<in>Basis. (u \<bullet> basg j) *\<^sub>R j"
define g :: "'b \<Rightarrow> 'a*real" where "g \<equiv> \<lambda>z. (\<Sum>i\<in>Basis. (z \<bullet> basf i) *\<^sub>R i)"
show ?thesis
proof
show "linear f"
unfolding f_def
by (intro linear_compose_sum linearI ballI) (auto simp: algebra_simps)
show "linear g"
unfolding g_def
by (intro linear_compose_sum linearI ballI) (auto simp: algebra_simps)
have *: "(\<Sum>a \<in> Basis. a \<bullet> basf b * (x \<bullet> basg a)) = x \<bullet> b" if "b \<in> Basis" for x b
using sbf that by auto
show gf: "g (f x) = x" for x
apply (rule euclidean_eqI)
apply (simp add: f_def g_def inner_sum_left scaleR_sum_left algebra_simps)
apply (simp add: Groups_Big.sum_distrib_left [symmetric] *)
done
show "basf(0,1) \<in> Basis"
using b01 sbf by auto
then show "f(x,0) \<bullet> basf(0,1) = 0" for x
apply (simp add: f_def inner_sum_left)
apply (rule comm_monoid_add_class.sum.neutral)
using b01 inner_not_same_Basis by fastforce
qed
qed
proposition locally_compact_homeomorphic_closed:
fixes S :: "'a::euclidean_space set"
assumes "locally compact S" and dimlt: "DIM('a) < DIM('b)"
obtains T :: "'b::euclidean_space set" where "closed T" "S homeomorphic T"
proof -
obtain U:: "('a*real)set" and h
where "closed U" and homU: "homeomorphism S U h fst"
using locally_compact_homeomorphism_projection_closed assms by metis
obtain f :: "'a*real \<Rightarrow> 'b" and g :: "'b \<Rightarrow> 'a*real"
where "linear f" "linear g" and gf [simp]: "\<And>z. g (f z) = z"
using lowerdim_embeddings [OF dimlt] by metis
then have "inj f"
by (metis injI)
have gfU: "g ` f ` U = U"
by (simp add: image_comp o_def)
have "S homeomorphic U"
using homU homeomorphic_def by blast
also have "... homeomorphic f ` U"
apply (rule homeomorphicI [OF refl gfU])
apply (meson \<open>inj f\<close> \<open>linear f\<close> homeomorphism_cont2 linear_homeomorphism_image)
using \<open>linear g\<close> linear_continuous_on linear_conv_bounded_linear apply blast
apply (auto simp: o_def)
done
finally show ?thesis
apply (rule_tac T = "f ` U" in that)
apply (rule closed_injective_linear_image [OF \<open>closed U\<close> \<open>linear f\<close> \<open>inj f\<close>], assumption)
done
qed
lemma homeomorphic_convex_compact_lemma:
fixes S :: "'a::euclidean_space set"
assumes "convex S"
and "compact S"
and "cball 0 1 \<subseteq> S"
shows "S homeomorphic (cball (0::'a) 1)"
proof (rule starlike_compact_projective_special[OF assms(2-3)])
fix x u
assume "x \<in> S" and "0 \<le> u" and "u < (1::real)"
have "open (ball (u *\<^sub>R x) (1 - u))"
by (rule open_ball)
moreover have "u *\<^sub>R x \<in> ball (u *\<^sub>R x) (1 - u)"
unfolding centre_in_ball using \<open>u < 1\<close> by simp
moreover have "ball (u *\<^sub>R x) (1 - u) \<subseteq> S"
proof
fix y
assume "y \<in> ball (u *\<^sub>R x) (1 - u)"
then have "dist (u *\<^sub>R x) y < 1 - u"
unfolding mem_ball .
with \<open>u < 1\<close> have "inverse (1 - u) *\<^sub>R (y - u *\<^sub>R x) \<in> cball 0 1"
by (simp add: dist_norm inverse_eq_divide norm_minus_commute)
with assms(3) have "inverse (1 - u) *\<^sub>R (y - u *\<^sub>R x) \<in> S" ..
with assms(1) have "(1 - u) *\<^sub>R ((y - u *\<^sub>R x) /\<^sub>R (1 - u)) + u *\<^sub>R x \<in> S"
using \<open>x \<in> S\<close> \<open>0 \<le> u\<close> \<open>u < 1\<close> [THEN less_imp_le] by (rule convexD_alt)
then show "y \<in> S" using \<open>u < 1\<close>
by simp
qed
ultimately have "u *\<^sub>R x \<in> interior S" ..
then show "u *\<^sub>R x \<in> S - frontier S"
using frontier_def and interior_subset by auto
qed
proposition homeomorphic_convex_compact_cball:
fixes e :: real
and S :: "'a::euclidean_space set"
assumes "convex S"
and "compact S"
and "interior S \<noteq> {}"
and "e > 0"
shows "S homeomorphic (cball (b::'a) e)"
proof -
obtain a where "a \<in> interior S"
using assms(3) by auto
then obtain d where "d > 0" and d: "cball a d \<subseteq> S"
unfolding mem_interior_cball by auto
let ?d = "inverse d" and ?n = "0::'a"
have "cball ?n 1 \<subseteq> (\<lambda>x. inverse d *\<^sub>R (x - a)) ` S"
apply rule
apply (rule_tac x="d *\<^sub>R x + a" in image_eqI)
defer
apply (rule d[unfolded subset_eq, rule_format])
using \<open>d > 0\<close>
unfolding mem_cball dist_norm
apply (auto simp add: mult_right_le_one_le)
done
then have "(\<lambda>x. inverse d *\<^sub>R (x - a)) ` S homeomorphic cball ?n 1"
using homeomorphic_convex_compact_lemma[of "(\<lambda>x. ?d *\<^sub>R -a + ?d *\<^sub>R x) ` S",
OF convex_affinity compact_affinity]
using assms(1,2)
by (auto simp add: scaleR_right_diff_distrib)
then show ?thesis
apply (rule_tac homeomorphic_trans[OF _ homeomorphic_balls(2)[of 1 _ ?n]])
apply (rule homeomorphic_trans[OF homeomorphic_affinity[of "?d" S "?d *\<^sub>R -a"]])
using \<open>d>0\<close> \<open>e>0\<close>
apply (auto simp add: scaleR_right_diff_distrib)
done
qed
corollary homeomorphic_convex_compact:
fixes S :: "'a::euclidean_space set"
and T :: "'a set"
assumes "convex S" "compact S" "interior S \<noteq> {}"
and "convex T" "compact T" "interior T \<noteq> {}"
shows "S homeomorphic T"
using assms
by (meson zero_less_one homeomorphic_trans homeomorphic_convex_compact_cball homeomorphic_sym)
lemma homeomorphic_closed_intervals:
fixes a :: "'a::euclidean_space" and b and c :: "'a::euclidean_space" and d
assumes "box a b \<noteq> {}" and "box c d \<noteq> {}"
shows "(cbox a b) homeomorphic (cbox c d)"
apply (rule homeomorphic_convex_compact)
using assms apply auto
done
lemma homeomorphic_closed_intervals_real:
fixes a::real and b and c::real and d
assumes "a<b" and "c<d"
shows "{a..b} homeomorphic {c..d}"
using assms by (auto intro: homeomorphic_convex_compact)
subsection\<open>Covering spaces and lifting results for them\<close>
definition\<^marker>\<open>tag important\<close> covering_space
:: "'a::topological_space set \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b::topological_space set \<Rightarrow> bool"
where
"covering_space c p S \<equiv>
continuous_on c p \<and> p ` c = S \<and>
(\<forall>x \<in> S. \<exists>T. x \<in> T \<and> openin (top_of_set S) T \<and>
(\<exists>v. \<Union>v = c \<inter> p -` T \<and>
(\<forall>u \<in> v. openin (top_of_set c) u) \<and>
pairwise disjnt v \<and>
(\<forall>u \<in> v. \<exists>q. homeomorphism u T p q)))"
lemma covering_space_imp_continuous: "covering_space c p S \<Longrightarrow> continuous_on c p"
by (simp add: covering_space_def)
lemma covering_space_imp_surjective: "covering_space c p S \<Longrightarrow> p ` c = S"
by (simp add: covering_space_def)
lemma homeomorphism_imp_covering_space: "homeomorphism S T f g \<Longrightarrow> covering_space S f T"
apply (simp add: homeomorphism_def covering_space_def, clarify)
apply (rule_tac x=T in exI, simp)
apply (rule_tac x="{S}" in exI, auto)
done
lemma covering_space_local_homeomorphism:
assumes "covering_space c p S" "x \<in> c"
obtains T u q where "x \<in> T" "openin (top_of_set c) T"
"p x \<in> u" "openin (top_of_set S) u"
"homeomorphism T u p q"
using assms
apply (simp add: covering_space_def, clarify)
apply (drule_tac x="p x" in bspec, force)
by (metis IntI UnionE vimage_eq)
lemma covering_space_local_homeomorphism_alt:
assumes p: "covering_space c p S" and "y \<in> S"
obtains x T U q where "p x = y"
"x \<in> T" "openin (top_of_set c) T"
"y \<in> U" "openin (top_of_set S) U"
"homeomorphism T U p q"
proof -
obtain x where "p x = y" "x \<in> c"
using assms covering_space_imp_surjective by blast
show ?thesis
apply (rule covering_space_local_homeomorphism [OF p \<open>x \<in> c\<close>])
using that \<open>p x = y\<close> by blast
qed
proposition covering_space_open_map:
fixes S :: "'a :: metric_space set" and T :: "'b :: metric_space set"
assumes p: "covering_space c p S" and T: "openin (top_of_set c) T"
shows "openin (top_of_set S) (p ` T)"
proof -
have pce: "p ` c = S"
and covs:
"\<And>x. x \<in> S \<Longrightarrow>
\<exists>X VS. x \<in> X \<and> openin (top_of_set S) X \<and>
\<Union>VS = c \<inter> p -` X \<and>
(\<forall>u \<in> VS. openin (top_of_set c) u) \<and>
pairwise disjnt VS \<and>
(\<forall>u \<in> VS. \<exists>q. homeomorphism u X p q)"
using p by (auto simp: covering_space_def)
have "T \<subseteq> c" by (metis openin_euclidean_subtopology_iff T)
have "\<exists>X. openin (top_of_set S) X \<and> y \<in> X \<and> X \<subseteq> p ` T"
if "y \<in> p ` T" for y
proof -
have "y \<in> S" using \<open>T \<subseteq> c\<close> pce that by blast
obtain U VS where "y \<in> U" and U: "openin (top_of_set S) U"
and VS: "\<Union>VS = c \<inter> p -` U"
and openVS: "\<forall>V \<in> VS. openin (top_of_set c) V"
and homVS: "\<And>V. V \<in> VS \<Longrightarrow> \<exists>q. homeomorphism V U p q"
using covs [OF \<open>y \<in> S\<close>] by auto
obtain x where "x \<in> c" "p x \<in> U" "x \<in> T" "p x = y"
apply simp
using T [unfolded openin_euclidean_subtopology_iff] \<open>y \<in> U\<close> \<open>y \<in> p ` T\<close> by blast
with VS obtain V where "x \<in> V" "V \<in> VS" by auto
then obtain q where q: "homeomorphism V U p q" using homVS by blast
then have ptV: "p ` (T \<inter> V) = U \<inter> q -` (T \<inter> V)"
using VS \<open>V \<in> VS\<close> by (auto simp: homeomorphism_def)
have ocv: "openin (top_of_set c) V"
by (simp add: \<open>V \<in> VS\<close> openVS)
have "openin (top_of_set U) (U \<inter> q -` (T \<inter> V))"
apply (rule continuous_on_open [THEN iffD1, rule_format])
using homeomorphism_def q apply blast
using openin_subtopology_Int_subset [of c] q T unfolding homeomorphism_def
by (metis inf.absorb_iff2 Int_commute ocv openin_euclidean_subtopology_iff)
then have os: "openin (top_of_set S) (U \<inter> q -` (T \<inter> V))"
using openin_trans [of U] by (simp add: Collect_conj_eq U)
show ?thesis
apply (rule_tac x = "p ` (T \<inter> V)" in exI)
apply (rule conjI)
apply (simp only: ptV os)
using \<open>p x = y\<close> \<open>x \<in> V\<close> \<open>x \<in> T\<close> apply blast
done
qed
with openin_subopen show ?thesis by blast
qed
lemma covering_space_lift_unique_gen:
fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
fixes g1 :: "'a \<Rightarrow> 'c::real_normed_vector"
assumes cov: "covering_space c p S"
and eq: "g1 a = g2 a"
and f: "continuous_on T f" "f ` T \<subseteq> S"
and g1: "continuous_on T g1" "g1 ` T \<subseteq> c"
and fg1: "\<And>x. x \<in> T \<Longrightarrow> f x = p(g1 x)"
and g2: "continuous_on T g2" "g2 ` T \<subseteq> c"
and fg2: "\<And>x. x \<in> T \<Longrightarrow> f x = p(g2 x)"
and u_compt: "U \<in> components T" and "a \<in> U" "x \<in> U"
shows "g1 x = g2 x"
proof -
have "U \<subseteq> T" by (rule in_components_subset [OF u_compt])
define G12 where "G12 \<equiv> {x \<in> U. g1 x - g2 x = 0}"
have "connected U" by (rule in_components_connected [OF u_compt])
have contu: "continuous_on U g1" "continuous_on U g2"
using \<open>U \<subseteq> T\<close> continuous_on_subset g1 g2 by blast+
have o12: "openin (top_of_set U) G12"
unfolding G12_def
proof (subst openin_subopen, clarify)
fix z
assume z: "z \<in> U" "g1 z - g2 z = 0"
obtain v w q where "g1 z \<in> v" and ocv: "openin (top_of_set c) v"
and "p (g1 z) \<in> w" and osw: "openin (top_of_set S) w"
and hom: "homeomorphism v w p q"
apply (rule_tac x = "g1 z" in covering_space_local_homeomorphism [OF cov])
using \<open>U \<subseteq> T\<close> \<open>z \<in> U\<close> g1(2) apply blast+
done
have "g2 z \<in> v" using \<open>g1 z \<in> v\<close> z by auto
have gg: "U \<inter> g -` v = U \<inter> g -` (v \<inter> g ` U)" for g
by auto
have "openin (top_of_set (g1 ` U)) (v \<inter> g1 ` U)"
using ocv \<open>U \<subseteq> T\<close> g1 by (fastforce simp add: openin_open)
then have 1: "openin (top_of_set U) (U \<inter> g1 -` v)"
unfolding gg by (blast intro: contu continuous_on_open [THEN iffD1, rule_format])
have "openin (top_of_set (g2 ` U)) (v \<inter> g2 ` U)"
using ocv \<open>U \<subseteq> T\<close> g2 by (fastforce simp add: openin_open)
then have 2: "openin (top_of_set U) (U \<inter> g2 -` v)"
unfolding gg by (blast intro: contu continuous_on_open [THEN iffD1, rule_format])
show "\<exists>T. openin (top_of_set U) T \<and> z \<in> T \<and> T \<subseteq> {z \<in> U. g1 z - g2 z = 0}"
using z
apply (rule_tac x = "(U \<inter> g1 -` v) \<inter> (U \<inter> g2 -` v)" in exI)
apply (intro conjI)
apply (rule openin_Int [OF 1 2])
using \<open>g1 z \<in> v\<close> \<open>g2 z \<in> v\<close> apply (force simp:, clarify)
apply (metis \<open>U \<subseteq> T\<close> subsetD eq_iff_diff_eq_0 fg1 fg2 hom homeomorphism_def)
done
qed
have c12: "closedin (top_of_set U) G12"
unfolding G12_def
by (intro continuous_intros continuous_closedin_preimage_constant contu)
have "G12 = {} \<or> G12 = U"
by (intro connected_clopen [THEN iffD1, rule_format] \<open>connected U\<close> conjI o12 c12)
with eq \<open>a \<in> U\<close> have "\<And>x. x \<in> U \<Longrightarrow> g1 x - g2 x = 0" by (auto simp: G12_def)
then show ?thesis
using \<open>x \<in> U\<close> by force
qed
proposition covering_space_lift_unique:
fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
fixes g1 :: "'a \<Rightarrow> 'c::real_normed_vector"
assumes "covering_space c p S"
"g1 a = g2 a"
"continuous_on T f" "f ` T \<subseteq> S"
"continuous_on T g1" "g1 ` T \<subseteq> c" "\<And>x. x \<in> T \<Longrightarrow> f x = p(g1 x)"
"continuous_on T g2" "g2 ` T \<subseteq> c" "\<And>x. x \<in> T \<Longrightarrow> f x = p(g2 x)"
"connected T" "a \<in> T" "x \<in> T"
shows "g1 x = g2 x"
using covering_space_lift_unique_gen [of c p S] in_components_self assms ex_in_conv
by blast
lemma covering_space_locally:
fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
assumes loc: "locally \<phi> C" and cov: "covering_space C p S"
and pim: "\<And>T. \<lbrakk>T \<subseteq> C; \<phi> T\<rbrakk> \<Longrightarrow> \<psi>(p ` T)"
shows "locally \<psi> S"
proof -
have "locally \<psi> (p ` C)"
apply (rule locally_open_map_image [OF loc])
using cov covering_space_imp_continuous apply blast
using cov covering_space_imp_surjective covering_space_open_map apply blast
by (simp add: pim)
then show ?thesis
using covering_space_imp_surjective [OF cov] by metis
qed
proposition covering_space_locally_eq:
fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
assumes cov: "covering_space C p S"
and pim: "\<And>T. \<lbrakk>T \<subseteq> C; \<phi> T\<rbrakk> \<Longrightarrow> \<psi>(p ` T)"
and qim: "\<And>q U. \<lbrakk>U \<subseteq> S; continuous_on U q; \<psi> U\<rbrakk> \<Longrightarrow> \<phi>(q ` U)"
shows "locally \<psi> S \<longleftrightarrow> locally \<phi> C"
(is "?lhs = ?rhs")
proof
assume L: ?lhs
show ?rhs
proof (rule locallyI)
fix V x
assume V: "openin (top_of_set C) V" and "x \<in> V"
have "p x \<in> p ` C"
by (metis IntE V \<open>x \<in> V\<close> imageI openin_open)
then obtain T \<V> where "p x \<in> T"
and opeT: "openin (top_of_set S) T"
and veq: "\<Union>\<V> = C \<inter> p -` T"
and ope: "\<forall>U\<in>\<V>. openin (top_of_set C) U"
and hom: "\<forall>U\<in>\<V>. \<exists>q. homeomorphism U T p q"
using cov unfolding covering_space_def by (blast intro: that)
have "x \<in> \<Union>\<V>"
using V veq \<open>p x \<in> T\<close> \<open>x \<in> V\<close> openin_imp_subset by fastforce
then obtain U where "x \<in> U" "U \<in> \<V>"
by blast
then obtain q where opeU: "openin (top_of_set C) U" and q: "homeomorphism U T p q"
using ope hom by blast
with V have "openin (top_of_set C) (U \<inter> V)"
by blast
then have UV: "openin (top_of_set S) (p ` (U \<inter> V))"
using cov covering_space_open_map by blast
obtain W W' where opeW: "openin (top_of_set S) W" and "\<psi> W'" "p x \<in> W" "W \<subseteq> W'" and W'sub: "W' \<subseteq> p ` (U \<inter> V)"
using locallyE [OF L UV] \<open>x \<in> U\<close> \<open>x \<in> V\<close> by blast
then have "W \<subseteq> T"
by (metis Int_lower1 q homeomorphism_image1 image_Int_subset order_trans)
show "\<exists>U Z. openin (top_of_set C) U \<and>
\<phi> Z \<and> x \<in> U \<and> U \<subseteq> Z \<and> Z \<subseteq> V"
proof (intro exI conjI)
have "openin (top_of_set T) W"
by (meson opeW opeT openin_imp_subset openin_subset_trans \<open>W \<subseteq> T\<close>)
then have "openin (top_of_set U) (q ` W)"
by (meson homeomorphism_imp_open_map homeomorphism_symD q)
then show "openin (top_of_set C) (q ` W)"
using opeU openin_trans by blast
show "\<phi> (q ` W')"
by (metis (mono_tags, lifting) Int_subset_iff UV W'sub \<open>\<psi> W'\<close> continuous_on_subset dual_order.trans homeomorphism_def image_Int_subset openin_imp_subset q qim)
show "x \<in> q ` W"
by (metis \<open>p x \<in> W\<close> \<open>x \<in> U\<close> homeomorphism_def imageI q)
show "q ` W \<subseteq> q ` W'"
using \<open>W \<subseteq> W'\<close> by blast
have "W' \<subseteq> p ` V"
using W'sub by blast
then show "q ` W' \<subseteq> V"
using W'sub homeomorphism_apply1 [OF q] by auto
qed
qed
next
assume ?rhs
then show ?lhs
using cov covering_space_locally pim by blast
qed
lemma covering_space_locally_compact_eq:
fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
assumes "covering_space C p S"
shows "locally compact S \<longleftrightarrow> locally compact C"
apply (rule covering_space_locally_eq [OF assms])
apply (meson assms compact_continuous_image continuous_on_subset covering_space_imp_continuous)
using compact_continuous_image by blast
lemma covering_space_locally_connected_eq:
fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
assumes "covering_space C p S"
shows "locally connected S \<longleftrightarrow> locally connected C"
apply (rule covering_space_locally_eq [OF assms])
apply (meson connected_continuous_image assms continuous_on_subset covering_space_imp_continuous)
using connected_continuous_image by blast
lemma covering_space_locally_path_connected_eq:
fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
assumes "covering_space C p S"
shows "locally path_connected S \<longleftrightarrow> locally path_connected C"
apply (rule covering_space_locally_eq [OF assms])
apply (meson path_connected_continuous_image assms continuous_on_subset covering_space_imp_continuous)
using path_connected_continuous_image by blast
lemma covering_space_locally_compact:
fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
assumes "locally compact C" "covering_space C p S"
shows "locally compact S"
using assms covering_space_locally_compact_eq by blast
lemma covering_space_locally_connected:
fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
assumes "locally connected C" "covering_space C p S"
shows "locally connected S"
using assms covering_space_locally_connected_eq by blast
lemma covering_space_locally_path_connected:
fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
assumes "locally path_connected C" "covering_space C p S"
shows "locally path_connected S"
using assms covering_space_locally_path_connected_eq by blast
proposition covering_space_lift_homotopy:
fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
and h :: "real \<times> 'c::real_normed_vector \<Rightarrow> 'b"
assumes cov: "covering_space C p S"
and conth: "continuous_on ({0..1} \<times> U) h"
and him: "h ` ({0..1} \<times> U) \<subseteq> S"
and heq: "\<And>y. y \<in> U \<Longrightarrow> h (0,y) = p(f y)"
and contf: "continuous_on U f" and fim: "f ` U \<subseteq> C"
obtains k where "continuous_on ({0..1} \<times> U) k"
"k ` ({0..1} \<times> U) \<subseteq> C"
"\<And>y. y \<in> U \<Longrightarrow> k(0, y) = f y"
"\<And>z. z \<in> {0..1} \<times> U \<Longrightarrow> h z = p(k z)"
proof -
have "\<exists>V k. openin (top_of_set U) V \<and> y \<in> V \<and>
continuous_on ({0..1} \<times> V) k \<and> k ` ({0..1} \<times> V) \<subseteq> C \<and>
(\<forall>z \<in> V. k(0, z) = f z) \<and> (\<forall>z \<in> {0..1} \<times> V. h z = p(k z))"
if "y \<in> U" for y
proof -
obtain UU where UU: "\<And>s. s \<in> S \<Longrightarrow> s \<in> (UU s) \<and> openin (top_of_set S) (UU s) \<and>
(\<exists>\<V>. \<Union>\<V> = C \<inter> p -` UU s \<and>
(\<forall>U \<in> \<V>. openin (top_of_set C) U) \<and>
pairwise disjnt \<V> \<and>
(\<forall>U \<in> \<V>. \<exists>q. homeomorphism U (UU s) p q))"
using cov unfolding covering_space_def by (metis (mono_tags))
then have ope: "\<And>s. s \<in> S \<Longrightarrow> s \<in> (UU s) \<and> openin (top_of_set S) (UU s)"
by blast
have "\<exists>k n i. open k \<and> open n \<and>
t \<in> k \<and> y \<in> n \<and> i \<in> S \<and> h ` (({0..1} \<inter> k) \<times> (U \<inter> n)) \<subseteq> UU i" if "t \<in> {0..1}" for t
proof -
have hinS: "h (t, y) \<in> S"
using \<open>y \<in> U\<close> him that by blast
then have "(t,y) \<in> ({0..1} \<times> U) \<inter> h -` UU(h(t, y))"
using \<open>y \<in> U\<close> \<open>t \<in> {0..1}\<close> by (auto simp: ope)
moreover have ope_01U: "openin (top_of_set ({0..1} \<times> U)) (({0..1} \<times> U) \<inter> h -` UU(h(t, y)))"
using hinS ope continuous_on_open_gen [OF him] conth by blast
ultimately obtain V W where opeV: "open V" and "t \<in> {0..1} \<inter> V" "t \<in> {0..1} \<inter> V"
and opeW: "open W" and "y \<in> U" "y \<in> W"
and VW: "({0..1} \<inter> V) \<times> (U \<inter> W) \<subseteq> (({0..1} \<times> U) \<inter> h -` UU(h(t, y)))"
by (rule Times_in_interior_subtopology) (auto simp: openin_open)
then show ?thesis
using hinS by blast
qed
then obtain K NN X where
K: "\<And>t. t \<in> {0..1} \<Longrightarrow> open (K t)"
and NN: "\<And>t. t \<in> {0..1} \<Longrightarrow> open (NN t)"
and inUS: "\<And>t. t \<in> {0..1} \<Longrightarrow> t \<in> K t \<and> y \<in> NN t \<and> X t \<in> S"
and him: "\<And>t. t \<in> {0..1} \<Longrightarrow> h ` (({0..1} \<inter> K t) \<times> (U \<inter> NN t)) \<subseteq> UU (X t)"
by (metis (mono_tags))
obtain \<T> where "\<T> \<subseteq> ((\<lambda>i. K i \<times> NN i)) ` {0..1}" "finite \<T>" "{0::real..1} \<times> {y} \<subseteq> \<Union>\<T>"
proof (rule compactE)
show "compact ({0::real..1} \<times> {y})"
by (simp add: compact_Times)
show "{0..1} \<times> {y} \<subseteq> (\<Union>i\<in>{0..1}. K i \<times> NN i)"
using K inUS by auto
show "\<And>B. B \<in> (\<lambda>i. K i \<times> NN i) ` {0..1} \<Longrightarrow> open B"
using K NN by (auto simp: open_Times)
qed blast
then obtain tk where "tk \<subseteq> {0..1}" "finite tk"
and tk: "{0::real..1} \<times> {y} \<subseteq> (\<Union>i \<in> tk. K i \<times> NN i)"
by (metis (no_types, lifting) finite_subset_image)
then have "tk \<noteq> {}"
by auto
define n where "n = \<Inter>(NN ` tk)"
have "y \<in> n" "open n"
using inUS NN \<open>tk \<subseteq> {0..1}\<close> \<open>finite tk\<close>
by (auto simp: n_def open_INT subset_iff)
obtain \<delta> where "0 < \<delta>" and \<delta>: "\<And>T. \<lbrakk>T \<subseteq> {0..1}; diameter T < \<delta>\<rbrakk> \<Longrightarrow> \<exists>B\<in>K ` tk. T \<subseteq> B"
proof (rule Lebesgue_number_lemma [of "{0..1}" "K ` tk"])
show "K ` tk \<noteq> {}"
using \<open>tk \<noteq> {}\<close> by auto
show "{0..1} \<subseteq> \<Union>(K ` tk)"
using tk by auto
show "\<And>B. B \<in> K ` tk \<Longrightarrow> open B"
using \<open>tk \<subseteq> {0..1}\<close> K by auto
qed auto
obtain N::nat where N: "N > 1 / \<delta>"
using reals_Archimedean2 by blast
then have "N > 0"
using \<open>0 < \<delta>\<close> order.asym by force
have *: "\<exists>V k. openin (top_of_set U) V \<and> y \<in> V \<and>
continuous_on ({0..of_nat n / N} \<times> V) k \<and>
k ` ({0..of_nat n / N} \<times> V) \<subseteq> C \<and>
(\<forall>z\<in>V. k (0, z) = f z) \<and>
(\<forall>z\<in>{0..of_nat n / N} \<times> V. h z = p (k z))" if "n \<le> N" for n
using that
proof (induction n)
case 0
show ?case
apply (rule_tac x=U in exI)
apply (rule_tac x="f \<circ> snd" in exI)
apply (intro conjI \<open>y \<in> U\<close> continuous_intros continuous_on_subset [OF contf])
using fim apply (auto simp: heq)
done
next
case (Suc n)
then obtain V k where opeUV: "openin (top_of_set U) V"
and "y \<in> V"
and contk: "continuous_on ({0..real n / real N} \<times> V) k"
and kim: "k ` ({0..real n / real N} \<times> V) \<subseteq> C"
and keq: "\<And>z. z \<in> V \<Longrightarrow> k (0, z) = f z"
and heq: "\<And>z. z \<in> {0..real n / real N} \<times> V \<Longrightarrow> h z = p (k z)"
using Suc_leD by auto
have "n \<le> N"
using Suc.prems by auto
obtain t where "t \<in> tk" and t: "{real n / real N .. (1 + real n) / real N} \<subseteq> K t"
proof (rule bexE [OF \<delta>])
show "{real n / real N .. (1 + real n) / real N} \<subseteq> {0..1}"
using Suc.prems by (auto simp: field_split_simps)
show diameter_less: "diameter {real n / real N .. (1 + real n) / real N} < \<delta>"
using \<open>0 < \<delta>\<close> N by (auto simp: field_split_simps)
qed blast
have t01: "t \<in> {0..1}"
using \<open>t \<in> tk\<close> \<open>tk \<subseteq> {0..1}\<close> by blast
obtain \<V> where \<V>: "\<Union>\<V> = C \<inter> p -` UU (X t)"
and opeC: "\<And>U. U \<in> \<V> \<Longrightarrow> openin (top_of_set C) U"
and "pairwise disjnt \<V>"
and homuu: "\<And>U. U \<in> \<V> \<Longrightarrow> \<exists>q. homeomorphism U (UU (X t)) p q"
using inUS [OF t01] UU by meson
have n_div_N_in: "real n / real N \<in> {real n / real N .. (1 + real n) / real N}"
using N by (auto simp: field_split_simps)
with t have nN_in_kkt: "real n / real N \<in> K t"
by blast
have "k (real n / real N, y) \<in> C \<inter> p -` UU (X t)"
proof (simp, rule conjI)
show "k (real n / real N, y) \<in> C"
using \<open>y \<in> V\<close> kim keq by force
have "p (k (real n / real N, y)) = h (real n / real N, y)"
by (simp add: \<open>y \<in> V\<close> heq)
also have "... \<in> h ` (({0..1} \<inter> K t) \<times> (U \<inter> NN t))"
apply (rule imageI)
using \<open>y \<in> V\<close> t01 \<open>n \<le> N\<close>
apply (simp add: nN_in_kkt \<open>y \<in> U\<close> inUS field_split_simps)
done
also have "... \<subseteq> UU (X t)"
using him t01 by blast
finally show "p (k (real n / real N, y)) \<in> UU (X t)" .
qed
with \<V> have "k (real n / real N, y) \<in> \<Union>\<V>"
by blast
then obtain W where W: "k (real n / real N, y) \<in> W" and "W \<in> \<V>"
by blast
then obtain p' where opeC': "openin (top_of_set C) W"
and hom': "homeomorphism W (UU (X t)) p p'"
using homuu opeC by blast
then have "W \<subseteq> C"
using openin_imp_subset by blast
define W' where "W' = UU(X t)"
have opeVW: "openin (top_of_set V) (V \<inter> (k \<circ> Pair (n / N)) -` W)"
apply (rule continuous_openin_preimage [OF _ _ opeC'])
apply (intro continuous_intros continuous_on_subset [OF contk])
using kim apply (auto simp: \<open>y \<in> V\<close> W)
done
obtain N' where opeUN': "openin (top_of_set U) N'"
and "y \<in> N'" and kimw: "k ` ({(real n / real N)} \<times> N') \<subseteq> W"
apply (rule_tac N' = "(V \<inter> (k \<circ> Pair (n / N)) -` W)" in that)
apply (fastforce simp: \<open>y \<in> V\<close> W intro!: openin_trans [OF opeVW opeUV])+
done
obtain Q Q' where opeUQ: "openin (top_of_set U) Q"
and cloUQ': "closedin (top_of_set U) Q'"
and "y \<in> Q" "Q \<subseteq> Q'"
and Q': "Q' \<subseteq> (U \<inter> NN(t)) \<inter> N' \<inter> V"
proof -
obtain VO VX where "open VO" "open VX" and VO: "V = U \<inter> VO" and VX: "N' = U \<inter> VX"
using opeUV opeUN' by (auto simp: openin_open)
then have "open (NN(t) \<inter> VO \<inter> VX)"
using NN t01 by blast
then obtain e where "e > 0" and e: "cball y e \<subseteq> NN(t) \<inter> VO \<inter> VX"
by (metis Int_iff \<open>N' = U \<inter> VX\<close> \<open>V = U \<inter> VO\<close> \<open>y \<in> N'\<close> \<open>y \<in> V\<close> inUS open_contains_cball t01)
show ?thesis
proof
show "openin (top_of_set U) (U \<inter> ball y e)"
by blast
show "closedin (top_of_set U) (U \<inter> cball y e)"
using e by (auto simp: closedin_closed)
qed (use \<open>y \<in> U\<close> \<open>e > 0\<close> VO VX e in auto)
qed
then have "y \<in> Q'" "Q \<subseteq> (U \<inter> NN(t)) \<inter> N' \<inter> V"
by blast+
have neq: "{0..real n / real N} \<union> {real n / real N..(1 + real n) / real N} = {0..(1 + real n) / real N}"
apply (auto simp: field_split_simps)
by (metis mult_zero_left of_nat_0_le_iff of_nat_0_less_iff order_trans real_mult_le_cancel_iff1)
then have neqQ': "{0..real n / real N} \<times> Q' \<union> {real n / real N..(1 + real n) / real N} \<times> Q' = {0..(1 + real n) / real N} \<times> Q'"
by blast
have cont: "continuous_on ({0..(1 + real n) / real N} \<times> Q')
(\<lambda>x. if x \<in> {0..real n / real N} \<times> Q' then k x else (p' \<circ> h) x)"
unfolding neqQ' [symmetric]
proof (rule continuous_on_cases_local, simp_all add: neqQ' del: comp_apply)
show "closedin (top_of_set ({0..(1 + real n) / real N} \<times> Q'))
({0..real n / real N} \<times> Q')"
apply (simp add: closedin_closed)
apply (rule_tac x="{0 .. real n / real N} \<times> UNIV" in exI)
using n_div_N_in apply (auto simp: closed_Times)
done
show "closedin (top_of_set ({0..(1 + real n) / real N} \<times> Q'))
({real n / real N..(1 + real n) / real N} \<times> Q')"
apply (simp add: closedin_closed)
apply (rule_tac x="{real n / real N .. (1 + real n) / real N} \<times> UNIV" in exI)
apply (auto simp: closed_Times)
by (meson divide_nonneg_nonneg of_nat_0_le_iff order_trans)
show "continuous_on ({0..real n / real N} \<times> Q') k"
apply (rule continuous_on_subset [OF contk])
using Q' by auto
have "continuous_on ({real n / real N..(1 + real n) / real N} \<times> Q') h"
proof (rule continuous_on_subset [OF conth])
show "{real n / real N..(1 + real n) / real N} \<times> Q' \<subseteq> {0..1} \<times> U"
using \<open>N > 0\<close>
apply auto
apply (meson divide_nonneg_nonneg of_nat_0_le_iff order_trans)
using Suc.prems order_trans apply fastforce
apply (metis IntE cloUQ' closedin_closed)
done
qed
moreover have "continuous_on (h ` ({real n / real N..(1 + real n) / real N} \<times> Q')) p'"
proof (rule continuous_on_subset [OF homeomorphism_cont2 [OF hom']])
have "h ` ({real n / real N..(1 + real n) / real N} \<times> Q') \<subseteq> h ` (({0..1} \<inter> K t) \<times> (U \<inter> NN t))"
apply (rule image_mono)
using \<open>0 < \<delta>\<close> \<open>N > 0\<close> Suc.prems apply auto
apply (meson divide_nonneg_nonneg of_nat_0_le_iff order_trans)
using Suc.prems order_trans apply fastforce
using t Q' apply auto
done
with him show "h ` ({real n / real N..(1 + real n) / real N} \<times> Q') \<subseteq> UU (X t)"
using t01 by blast
qed
ultimately show "continuous_on ({real n / real N..(1 + real n) / real N} \<times> Q') (p' \<circ> h)"
by (rule continuous_on_compose)
have "k (real n / real N, b) = p' (h (real n / real N, b))" if "b \<in> Q'" for b
proof -
have "k (real n / real N, b) \<in> W"
using that Q' kimw by force
then have "k (real n / real N, b) = p' (p (k (real n / real N, b)))"
by (simp add: homeomorphism_apply1 [OF hom'])
then show ?thesis
using Q' that by (force simp: heq)
qed
then show "\<And>x. x \<in> {real n / real N..(1 + real n) / real N} \<times> Q' \<and>
x \<in> {0..real n / real N} \<times> Q' \<Longrightarrow> k x = (p' \<circ> h) x"
by auto
qed
have h_in_UU: "h (x, y) \<in> UU (X t)" if "y \<in> Q" "\<not> x \<le> real n / real N" "0 \<le> x" "x \<le> (1 + real n) / real N" for x y
proof -
have "x \<le> 1"
using Suc.prems that order_trans by force
moreover have "x \<in> K t"
by (meson atLeastAtMost_iff le_less not_le subset_eq t that)
moreover have "y \<in> U"
using \<open>y \<in> Q\<close> opeUQ openin_imp_subset by blast
moreover have "y \<in> NN t"
using Q' \<open>Q \<subseteq> Q'\<close> \<open>y \<in> Q\<close> by auto
ultimately have "(x, y) \<in> (({0..1} \<inter> K t) \<times> (U \<inter> NN t))"
using that by auto
then have "h (x, y) \<in> h ` (({0..1} \<inter> K t) \<times> (U \<inter> NN t))"
by blast
also have "... \<subseteq> UU (X t)"
by (metis him t01)
finally show ?thesis .
qed
let ?k = "(\<lambda>x. if x \<in> {0..real n / real N} \<times> Q' then k x else (p' \<circ> h) x)"
show ?case
proof (intro exI conjI)
show "continuous_on ({0..real (Suc n) / real N} \<times> Q) ?k"
apply (rule continuous_on_subset [OF cont])
using \<open>Q \<subseteq> Q'\<close> by auto
have "\<And>a b. \<lbrakk>a \<le> real n / real N; b \<in> Q'; 0 \<le> a\<rbrakk> \<Longrightarrow> k (a, b) \<in> C"
using kim Q' by force
moreover have "\<And>a b. \<lbrakk>b \<in> Q; 0 \<le> a; a \<le> (1 + real n) / real N; \<not> a \<le> real n / real N\<rbrakk> \<Longrightarrow> p' (h (a, b)) \<in> C"
apply (rule \<open>W \<subseteq> C\<close> [THEN subsetD])
using homeomorphism_image2 [OF hom', symmetric] h_in_UU Q' \<open>Q \<subseteq> Q'\<close> \<open>W \<subseteq> C\<close>
apply auto
done
ultimately show "?k ` ({0..real (Suc n) / real N} \<times> Q) \<subseteq> C"
using Q' \<open>Q \<subseteq> Q'\<close> by force
show "\<forall>z\<in>Q. ?k (0, z) = f z"
using Q' keq \<open>Q \<subseteq> Q'\<close> by auto
show "\<forall>z \<in> {0..real (Suc n) / real N} \<times> Q. h z = p(?k z)"
using \<open>Q \<subseteq> U \<inter> NN t \<inter> N' \<inter> V\<close> heq apply clarsimp
using h_in_UU Q' \<open>Q \<subseteq> Q'\<close> apply (auto simp: homeomorphism_apply2 [OF hom', symmetric])
done
qed (auto simp: \<open>y \<in> Q\<close> opeUQ)
qed
show ?thesis
using*[OF order_refl] N \<open>0 < \<delta>\<close> by (simp add: split: if_split_asm)
qed
then obtain V fs where opeV: "\<And>y. y \<in> U \<Longrightarrow> openin (top_of_set U) (V y)"
and V: "\<And>y. y \<in> U \<Longrightarrow> y \<in> V y"
and contfs: "\<And>y. y \<in> U \<Longrightarrow> continuous_on ({0..1} \<times> V y) (fs y)"
and *: "\<And>y. y \<in> U \<Longrightarrow> (fs y) ` ({0..1} \<times> V y) \<subseteq> C \<and>
(\<forall>z \<in> V y. fs y (0, z) = f z) \<and>
(\<forall>z \<in> {0..1} \<times> V y. h z = p(fs y z))"
by (metis (mono_tags))
then have VU: "\<And>y. y \<in> U \<Longrightarrow> V y \<subseteq> U"
by (meson openin_imp_subset)
obtain k where contk: "continuous_on ({0..1} \<times> U) k"
and k: "\<And>x i. \<lbrakk>i \<in> U; x \<in> {0..1} \<times> U \<inter> {0..1} \<times> V i\<rbrakk> \<Longrightarrow> k x = fs i x"
proof (rule pasting_lemma_exists)
let ?X = "top_of_set ({0..1::real} \<times> U)"
show "topspace ?X \<subseteq> (\<Union>i\<in>U. {0..1} \<times> V i)"
using V by force
show "\<And>i. i \<in> U \<Longrightarrow> openin (top_of_set ({0..1} \<times> U)) ({0..1} \<times> V i)"
by (simp add: Abstract_Topology.openin_Times opeV)
show "\<And>i. i \<in> U \<Longrightarrow> continuous_map
(subtopology (top_of_set ({0..1} \<times> U)) ({0..1} \<times> V i)) euclidean (fs i)"
using contfs
apply simp
by (metis continuous_map_iff_continuous continuous_on_subset openin_imp_subset openin_subtopology_self subtopology_subtopology)
show "fs i x = fs j x" if "i \<in> U" "j \<in> U" and x: "x \<in> topspace ?X \<inter> {0..1} \<times> V i \<inter> {0..1} \<times> V j"
for i j x
proof -
obtain u y where "x = (u, y)" "y \<in> V i" "y \<in> V j" "0 \<le> u" "u \<le> 1"
using x by auto
show ?thesis
proof (rule covering_space_lift_unique [OF cov, of _ "(0,y)" _ "{0..1} \<times> {y}" h])
show "fs i (0, y) = fs j (0, y)"
using*V by (simp add: \<open>y \<in> V i\<close> \<open>y \<in> V j\<close> that)
show conth_y: "continuous_on ({0..1} \<times> {y}) h"
apply (rule continuous_on_subset [OF conth])
using VU \<open>y \<in> V j\<close> that by auto
show "h ` ({0..1} \<times> {y}) \<subseteq> S"
using \<open>y \<in> V i\<close> assms(3) VU that by fastforce
show "continuous_on ({0..1} \<times> {y}) (fs i)"
using continuous_on_subset [OF contfs] \<open>i \<in> U\<close>
by (simp add: \<open>y \<in> V i\<close> subset_iff)
show "fs i ` ({0..1} \<times> {y}) \<subseteq> C"
using "*" \<open>y \<in> V i\<close> \<open>i \<in> U\<close> by fastforce
show "\<And>x. x \<in> {0..1} \<times> {y} \<Longrightarrow> h x = p (fs i x)"
using "*" \<open>y \<in> V i\<close> \<open>i \<in> U\<close> by blast
show "continuous_on ({0..1} \<times> {y}) (fs j)"
using continuous_on_subset [OF contfs] \<open>j \<in> U\<close>
by (simp add: \<open>y \<in> V j\<close> subset_iff)
show "fs j ` ({0..1} \<times> {y}) \<subseteq> C"
using "*" \<open>y \<in> V j\<close> \<open>j \<in> U\<close> by fastforce
show "\<And>x. x \<in> {0..1} \<times> {y} \<Longrightarrow> h x = p (fs j x)"
using "*" \<open>y \<in> V j\<close> \<open>j \<in> U\<close> by blast
show "connected ({0..1::real} \<times> {y})"
using connected_Icc connected_Times connected_sing by blast
show "(0, y) \<in> {0..1::real} \<times> {y}"
by force
show "x \<in> {0..1} \<times> {y}"
using \<open>x = (u, y)\<close> x by blast
qed
qed
qed force
show ?thesis
proof
show "k ` ({0..1} \<times> U) \<subseteq> C"
using V*k VU by fastforce
show "\<And>y. y \<in> U \<Longrightarrow> k (0, y) = f y"
by (simp add: V*k)
show "\<And>z. z \<in> {0..1} \<times> U \<Longrightarrow> h z = p (k z)"
using V*k by auto
qed (auto simp: contk)
qed
corollary covering_space_lift_homotopy_alt:
fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
and h :: "'c::real_normed_vector \<times> real \<Rightarrow> 'b"
assumes cov: "covering_space C p S"
and conth: "continuous_on (U \<times> {0..1}) h"
and him: "h ` (U \<times> {0..1}) \<subseteq> S"
and heq: "\<And>y. y \<in> U \<Longrightarrow> h (y,0) = p(f y)"
and contf: "continuous_on U f" and fim: "f ` U \<subseteq> C"
obtains k where "continuous_on (U \<times> {0..1}) k"
"k ` (U \<times> {0..1}) \<subseteq> C"
"\<And>y. y \<in> U \<Longrightarrow> k(y, 0) = f y"
"\<And>z. z \<in> U \<times> {0..1} \<Longrightarrow> h z = p(k z)"
proof -
have "continuous_on ({0..1} \<times> U) (h \<circ> (\<lambda>z. (snd z, fst z)))"
by (intro continuous_intros continuous_on_subset [OF conth]) auto
then obtain k where contk: "continuous_on ({0..1} \<times> U) k"
and kim: "k ` ({0..1} \<times> U) \<subseteq> C"
and k0: "\<And>y. y \<in> U \<Longrightarrow> k(0, y) = f y"
and heqp: "\<And>z. z \<in> {0..1} \<times> U \<Longrightarrow> (h \<circ> (\<lambda>z. Pair (snd z) (fst z))) z = p(k z)"
apply (rule covering_space_lift_homotopy [OF cov _ _ _ contf fim])
using him by (auto simp: contf heq)
show ?thesis
apply (rule_tac k="k \<circ> (\<lambda>z. Pair (snd z) (fst z))" in that)
apply (intro continuous_intros continuous_on_subset [OF contk])
using kim heqp apply (auto simp: k0)
done
qed
corollary covering_space_lift_homotopic_function:
fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" and g:: "'c::real_normed_vector \<Rightarrow> 'a"
assumes cov: "covering_space C p S"
and contg: "continuous_on U g"
and gim: "g ` U \<subseteq> C"
and pgeq: "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y"
and hom: "homotopic_with_canon (\<lambda>x. True) U S f f'"
obtains g' where "continuous_on U g'" "image g' U \<subseteq> C" "\<And>y. y \<in> U \<Longrightarrow> p(g' y) = f' y"
proof -
obtain h where conth: "continuous_on ({0..1::real} \<times> U) h"
and him: "h ` ({0..1} \<times> U) \<subseteq> S"
and h0: "\<And>x. h(0, x) = f x"
and h1: "\<And>x. h(1, x) = f' x"
using hom by (auto simp: homotopic_with_def)
have "\<And>y. y \<in> U \<Longrightarrow> h (0, y) = p (g y)"
by (simp add: h0 pgeq)
then obtain k where contk: "continuous_on ({0..1} \<times> U) k"
and kim: "k ` ({0..1} \<times> U) \<subseteq> C"
and k0: "\<And>y. y \<in> U \<Longrightarrow> k(0, y) = g y"
and heq: "\<And>z. z \<in> {0..1} \<times> U \<Longrightarrow> h z = p(k z)"
using covering_space_lift_homotopy [OF cov conth him _ contg gim] by metis
show ?thesis
proof
show "continuous_on U (k \<circ> Pair 1)"
by (meson contk atLeastAtMost_iff continuous_on_o_Pair order_refl zero_le_one)
show "(k \<circ> Pair 1) ` U \<subseteq> C"
using kim by auto
show "\<And>y. y \<in> U \<Longrightarrow> p ((k \<circ> Pair 1) y) = f' y"
by (auto simp: h1 heq [symmetric])
qed
qed
corollary covering_space_lift_inessential_function:
fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" and U :: "'c::real_normed_vector set"
assumes cov: "covering_space C p S"
and hom: "homotopic_with_canon (\<lambda>x. True) U S f (\<lambda>x. a)"
obtains g where "continuous_on U g" "g ` U \<subseteq> C" "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y"
proof (cases "U = {}")
case True
then show ?thesis
using that continuous_on_empty by blast
next
case False
then obtain b where b: "b \<in> C" "p b = a"
using covering_space_imp_surjective [OF cov] homotopic_with_imp_subset2 [OF hom]
by auto
then have gim: "(\<lambda>y. b) ` U \<subseteq> C"
by blast
show ?thesis
apply (rule covering_space_lift_homotopic_function
[OF cov continuous_on_const gim _ homotopic_with_symD [OF hom]])
using b that apply auto
done
qed
subsection\<open> Lifting of general functions to covering space\<close>
proposition covering_space_lift_path_strong:
fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
and f :: "'c::real_normed_vector \<Rightarrow> 'b"
assumes cov: "covering_space C p S" and "a \<in> C"
and "path g" and pag: "path_image g \<subseteq> S" and pas: "pathstart g = p a"
obtains h where "path h" "path_image h \<subseteq> C" "pathstart h = a"
and "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h t) = g t"
proof -
obtain k:: "real \<times> 'c \<Rightarrow> 'a"
where contk: "continuous_on ({0..1} \<times> {undefined}) k"
and kim: "k ` ({0..1} \<times> {undefined}) \<subseteq> C"
and k0: "k (0, undefined) = a"
and pk: "\<And>z. z \<in> {0..1} \<times> {undefined} \<Longrightarrow> p(k z) = (g \<circ> fst) z"
proof (rule covering_space_lift_homotopy [OF cov, of "{undefined}" "g \<circ> fst"])
show "continuous_on ({0..1::real} \<times> {undefined::'c}) (g \<circ> fst)"
apply (intro continuous_intros)
using \<open>path g\<close> by (simp add: path_def)
show "(g \<circ> fst) ` ({0..1} \<times> {undefined}) \<subseteq> S"
using pag by (auto simp: path_image_def)
show "(g \<circ> fst) (0, y) = p a" if "y \<in> {undefined}" for y::'c
by (metis comp_def fst_conv pas pathstart_def)
qed (use assms in auto)
show ?thesis
proof
show "path (k \<circ> (\<lambda>t. Pair t undefined))"
unfolding path_def
by (intro continuous_on_compose continuous_intros continuous_on_subset [OF contk]) auto
show "path_image (k \<circ> (\<lambda>t. (t, undefined))) \<subseteq> C"
using kim by (auto simp: path_image_def)
show "pathstart (k \<circ> (\<lambda>t. (t, undefined))) = a"
by (auto simp: pathstart_def k0)
show "\<And>t. t \<in> {0..1} \<Longrightarrow> p ((k \<circ> (\<lambda>t. (t, undefined))) t) = g t"
by (auto simp: pk)
qed
qed
corollary covering_space_lift_path:
fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
assumes cov: "covering_space C p S" and "path g" and pig: "path_image g \<subseteq> S"
obtains h where "path h" "path_image h \<subseteq> C" "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h t) = g t"
proof -
obtain a where "a \<in> C" "pathstart g = p a"
by (metis pig cov covering_space_imp_surjective imageE pathstart_in_path_image subsetCE)
show ?thesis
using covering_space_lift_path_strong [OF cov \<open>a \<in> C\<close> \<open>path g\<close> pig]
by (metis \<open>pathstart g = p a\<close> that)
qed
proposition covering_space_lift_homotopic_paths:
fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
assumes cov: "covering_space C p S"
and "path g1" and pig1: "path_image g1 \<subseteq> S"
and "path g2" and pig2: "path_image g2 \<subseteq> S"
and hom: "homotopic_paths S g1 g2"
and "path h1" and pih1: "path_image h1 \<subseteq> C" and ph1: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h1 t) = g1 t"
and "path h2" and pih2: "path_image h2 \<subseteq> C" and ph2: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h2 t) = g2 t"
and h1h2: "pathstart h1 = pathstart h2"
shows "homotopic_paths C h1 h2"
proof -
obtain h :: "real \<times> real \<Rightarrow> 'b"
where conth: "continuous_on ({0..1} \<times> {0..1}) h"
and him: "h ` ({0..1} \<times> {0..1}) \<subseteq> S"
and h0: "\<And>x. h (0, x) = g1 x" and h1: "\<And>x. h (1, x) = g2 x"
and heq0: "\<And>t. t \<in> {0..1} \<Longrightarrow> h (t, 0) = g1 0"
and heq1: "\<And>t. t \<in> {0..1} \<Longrightarrow> h (t, 1) = g1 1"
using hom by (auto simp: homotopic_paths_def homotopic_with_def pathstart_def pathfinish_def)
obtain k where contk: "continuous_on ({0..1} \<times> {0..1}) k"
and kim: "k ` ({0..1} \<times> {0..1}) \<subseteq> C"
and kh2: "\<And>y. y \<in> {0..1} \<Longrightarrow> k (y, 0) = h2 0"
and hpk: "\<And>z. z \<in> {0..1} \<times> {0..1} \<Longrightarrow> h z = p (k z)"
apply (rule covering_space_lift_homotopy_alt [OF cov conth him, of "\<lambda>x. h2 0"])
using h1h2 ph1 ph2 apply (force simp: heq0 pathstart_def pathfinish_def)
using path_image_def pih2 by fastforce+
have contg1: "continuous_on {0..1} g1" and contg2: "continuous_on {0..1} g2"
using \<open>path g1\<close> \<open>path g2\<close> path_def by blast+
have g1im: "g1 ` {0..1} \<subseteq> S" and g2im: "g2 ` {0..1} \<subseteq> S"
using path_image_def pig1 pig2 by auto
have conth1: "continuous_on {0..1} h1" and conth2: "continuous_on {0..1} h2"
using \<open>path h1\<close> \<open>path h2\<close> path_def by blast+
have h1im: "h1 ` {0..1} \<subseteq> C" and h2im: "h2 ` {0..1} \<subseteq> C"
using path_image_def pih1 pih2 by auto
show ?thesis
unfolding homotopic_paths pathstart_def pathfinish_def
proof (intro exI conjI ballI)
show keqh1: "k(0, x) = h1 x" if "x \<in> {0..1}" for x
proof (rule covering_space_lift_unique [OF cov _ contg1 g1im])
show "k (0,0) = h1 0"
by (metis atLeastAtMost_iff h1h2 kh2 order_refl pathstart_def zero_le_one)
show "continuous_on {0..1} (\<lambda>a. k (0, a))"
by (intro continuous_intros continuous_on_compose2 [OF contk]) auto
show "\<And>x. x \<in> {0..1} \<Longrightarrow> g1 x = p (k (0, x))"
by (metis atLeastAtMost_iff h0 hpk zero_le_one mem_Sigma_iff order_refl)
qed (use conth1 h1im kim that in \<open>auto simp: ph1\<close>)
show "k(1, x) = h2 x" if "x \<in> {0..1}" for x
proof (rule covering_space_lift_unique [OF cov _ contg2 g2im])
show "k (1,0) = h2 0"
by (metis atLeastAtMost_iff kh2 order_refl zero_le_one)
show "continuous_on {0..1} (\<lambda>a. k (1, a))"
by (intro continuous_intros continuous_on_compose2 [OF contk]) auto
show "\<And>x. x \<in> {0..1} \<Longrightarrow> g2 x = p (k (1, x))"
by (metis atLeastAtMost_iff h1 hpk mem_Sigma_iff order_refl zero_le_one)
qed (use conth2 h2im kim that in \<open>auto simp: ph2\<close>)
show "\<And>t. t \<in> {0..1} \<Longrightarrow> (k \<circ> Pair t) 0 = h1 0"
by (metis comp_apply h1h2 kh2 pathstart_def)
show "(k \<circ> Pair t) 1 = h1 1" if "t \<in> {0..1}" for t
proof (rule covering_space_lift_unique
[OF cov, of "\<lambda>a. (k \<circ> Pair a) 1" 0 "\<lambda>a. h1 1" "{0..1}" "\<lambda>x. g1 1"])
show "(k \<circ> Pair 0) 1 = h1 1"
using keqh1 by auto
show "continuous_on {0..1} (\<lambda>a. (k \<circ> Pair a) 1)"
apply simp
by (intro continuous_intros continuous_on_compose2 [OF contk]) auto
show "\<And>x. x \<in> {0..1} \<Longrightarrow> g1 1 = p ((k \<circ> Pair x) 1)"
using heq1 hpk by auto
qed (use contk kim g1im h1im that in \<open>auto simp: ph1\<close>)
qed (use contk kim in auto)
qed
corollary covering_space_monodromy:
fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
assumes cov: "covering_space C p S"
and "path g1" and pig1: "path_image g1 \<subseteq> S"
and "path g2" and pig2: "path_image g2 \<subseteq> S"
and hom: "homotopic_paths S g1 g2"
and "path h1" and pih1: "path_image h1 \<subseteq> C" and ph1: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h1 t) = g1 t"
and "path h2" and pih2: "path_image h2 \<subseteq> C" and ph2: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h2 t) = g2 t"
and h1h2: "pathstart h1 = pathstart h2"
shows "pathfinish h1 = pathfinish h2"
using covering_space_lift_homotopic_paths [OF assms] homotopic_paths_imp_pathfinish
by blast
corollary covering_space_lift_homotopic_path:
fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
assumes cov: "covering_space C p S"
and hom: "homotopic_paths S f f'"
and "path g" and pig: "path_image g \<subseteq> C"
and a: "pathstart g = a" and b: "pathfinish g = b"
and pgeq: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(g t) = f t"
obtains g' where "path g'" "path_image g' \<subseteq> C"
"pathstart g' = a" "pathfinish g' = b" "\<And>t. t \<in> {0..1} \<Longrightarrow> p(g' t) = f' t"
proof (rule covering_space_lift_path_strong [OF cov, of a f'])
show "a \<in> C"
using a pig by auto
show "path f'" "path_image f' \<subseteq> S"
using hom homotopic_paths_imp_path homotopic_paths_imp_subset by blast+
show "pathstart f' = p a"
by (metis a atLeastAtMost_iff hom homotopic_paths_imp_pathstart order_refl pathstart_def pgeq zero_le_one)
qed (metis (mono_tags, lifting) assms cov covering_space_monodromy hom homotopic_paths_imp_path homotopic_paths_imp_subset pgeq pig)
proposition covering_space_lift_general:
fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
and f :: "'c::real_normed_vector \<Rightarrow> 'b"
assumes cov: "covering_space C p S" and "a \<in> C" "z \<in> U"
and U: "path_connected U" "locally path_connected U"
and contf: "continuous_on U f" and fim: "f ` U \<subseteq> S"
and feq: "f z = p a"
and hom: "\<And>r. \<lbrakk>path r; path_image r \<subseteq> U; pathstart r = z; pathfinish r = z\<rbrakk>
\<Longrightarrow> \<exists>q. path q \<and> path_image q \<subseteq> C \<and>
pathstart q = a \<and> pathfinish q = a \<and>
homotopic_paths S (f \<circ> r) (p \<circ> q)"
obtains g where "continuous_on U g" "g ` U \<subseteq> C" "g z = a" "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y"
proof -
have *: "\<exists>g h. path g \<and> path_image g \<subseteq> U \<and>
pathstart g = z \<and> pathfinish g = y \<and>
path h \<and> path_image h \<subseteq> C \<and> pathstart h = a \<and>
(\<forall>t \<in> {0..1}. p(h t) = f(g t))"
if "y \<in> U" for y
proof -
obtain g where "path g" "path_image g \<subseteq> U" and pastg: "pathstart g = z"
and pafig: "pathfinish g = y"
using U \<open>z \<in> U\<close> \<open>y \<in> U\<close> by (force simp: path_connected_def)
obtain h where "path h" "path_image h \<subseteq> C" "pathstart h = a"
and "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h t) = (f \<circ> g) t"
proof (rule covering_space_lift_path_strong [OF cov \<open>a \<in> C\<close>])
show "path (f \<circ> g)"
using \<open>path g\<close> \<open>path_image g \<subseteq> U\<close> contf continuous_on_subset path_continuous_image by blast
show "path_image (f \<circ> g) \<subseteq> S"
by (metis \<open>path_image g \<subseteq> U\<close> fim image_mono path_image_compose subset_trans)
show "pathstart (f \<circ> g) = p a"
by (simp add: feq pastg pathstart_compose)
qed auto
then show ?thesis
by (metis \<open>path g\<close> \<open>path_image g \<subseteq> U\<close> comp_apply pafig pastg)
qed
have "\<exists>l. \<forall>g h. path g \<and> path_image g \<subseteq> U \<and> pathstart g = z \<and> pathfinish g = y \<and>
path h \<and> path_image h \<subseteq> C \<and> pathstart h = a \<and>
(\<forall>t \<in> {0..1}. p(h t) = f(g t)) \<longrightarrow> pathfinish h = l" for y
proof -
have "pathfinish h = pathfinish h'"
if g: "path g" "path_image g \<subseteq> U" "pathstart g = z" "pathfinish g = y"
and h: "path h" "path_image h \<subseteq> C" "pathstart h = a"
and phg: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h t) = f(g t)"
and g': "path g'" "path_image g' \<subseteq> U" "pathstart g' = z" "pathfinish g' = y"
and h': "path h'" "path_image h' \<subseteq> C" "pathstart h' = a"
and phg': "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h' t) = f(g' t)"
for g h g' h'
proof -
obtain q where "path q" and piq: "path_image q \<subseteq> C" and pastq: "pathstart q = a" and pafiq: "pathfinish q = a"
and homS: "homotopic_paths S (f \<circ> g +++ reversepath g') (p \<circ> q)"
using g g' hom [of "g +++ reversepath g'"] by (auto simp: subset_path_image_join)
have papq: "path (p \<circ> q)"
using homS homotopic_paths_imp_path by blast
have pipq: "path_image (p \<circ> q) \<subseteq> S"
using homS homotopic_paths_imp_subset by blast
obtain q' where "path q'" "path_image q' \<subseteq> C"
and "pathstart q' = pathstart q" "pathfinish q' = pathfinish q"
and pq'_eq: "\<And>t. t \<in> {0..1} \<Longrightarrow> p (q' t) = (f \<circ> g +++ reversepath g') t"
using covering_space_lift_homotopic_path [OF cov homotopic_paths_sym [OF homS] \<open>path q\<close> piq refl refl]
by auto
have "q' t = (h \<circ> (*\<^sub>R) 2) t" if "0 \<le> t" "t \<le> 1/2" for t
proof (rule covering_space_lift_unique [OF cov, of q' 0 "h \<circ> (*\<^sub>R) 2" "{0..1/2}" "f \<circ> g \<circ> (*\<^sub>R) 2" t])
show "q' 0 = (h \<circ> (*\<^sub>R) 2) 0"
- by (metis \<open>pathstart q' = pathstart q\<close> comp_def h pastq pathstart_def pth_4(2))
+ by (metis \<open>pathstart q' = pathstart q\<close> comp_def h(3) pastq pathstart_def pth_4(2))
show "continuous_on {0..1/2} (f \<circ> g \<circ> (*\<^sub>R) 2)"
apply (intro continuous_intros continuous_on_compose continuous_on_path [OF \<open>path g\<close>] continuous_on_subset [OF contf])
using g(2) path_image_def by fastforce+
show "(f \<circ> g \<circ> (*\<^sub>R) 2) ` {0..1/2} \<subseteq> S"
using g(2) path_image_def fim by fastforce
show "(h \<circ> (*\<^sub>R) 2) ` {0..1/2} \<subseteq> C"
using h path_image_def by fastforce
show "q' ` {0..1/2} \<subseteq> C"
using \<open>path_image q' \<subseteq> C\<close> path_image_def by fastforce
show "\<And>x. x \<in> {0..1/2} \<Longrightarrow> (f \<circ> g \<circ> (*\<^sub>R) 2) x = p (q' x)"
by (auto simp: joinpaths_def pq'_eq)
show "\<And>x. x \<in> {0..1/2} \<Longrightarrow> (f \<circ> g \<circ> (*\<^sub>R) 2) x = p ((h \<circ> (*\<^sub>R) 2) x)"
by (simp add: phg)
show "continuous_on {0..1/2} q'"
by (simp add: continuous_on_path \<open>path q'\<close>)
show "continuous_on {0..1/2} (h \<circ> (*\<^sub>R) 2)"
apply (intro continuous_intros continuous_on_compose continuous_on_path [OF \<open>path h\<close>], force)
done
qed (use that in auto)
moreover have "q' t = (reversepath h' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)) t" if "1/2 < t" "t \<le> 1" for t
proof (rule covering_space_lift_unique [OF cov, of q' 1 "reversepath h' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)" "{1/2<..1}" "f \<circ> reversepath g' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)" t])
show "q' 1 = (reversepath h' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)) 1"
using h' \<open>pathfinish q' = pathfinish q\<close> pafiq
by (simp add: pathstart_def pathfinish_def reversepath_def)
show "continuous_on {1/2<..1} (f \<circ> reversepath g' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1))"
apply (intro continuous_intros continuous_on_compose continuous_on_path \<open>path g'\<close> continuous_on_subset [OF contf])
using g' apply simp_all
by (auto simp: path_image_def reversepath_def)
show "(f \<circ> reversepath g' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)) ` {1/2<..1} \<subseteq> S"
using g'(2) path_image_def fim by (auto simp: image_subset_iff path_image_def reversepath_def)
show "q' ` {1/2<..1} \<subseteq> C"
using \<open>path_image q' \<subseteq> C\<close> path_image_def by fastforce
show "(reversepath h' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)) ` {1/2<..1} \<subseteq> C"
using h' by (simp add: path_image_def reversepath_def subset_eq)
show "\<And>x. x \<in> {1/2<..1} \<Longrightarrow> (f \<circ> reversepath g' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)) x = p (q' x)"
by (auto simp: joinpaths_def pq'_eq)
show "\<And>x. x \<in> {1/2<..1} \<Longrightarrow>
(f \<circ> reversepath g' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)) x = p ((reversepath h' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)) x)"
by (simp add: phg' reversepath_def)
show "continuous_on {1/2<..1} q'"
by (auto intro: continuous_on_path [OF \<open>path q'\<close>])
show "continuous_on {1/2<..1} (reversepath h' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1))"
apply (intro continuous_intros continuous_on_compose continuous_on_path \<open>path h'\<close>)
using h' apply auto
done
qed (use that in auto)
ultimately have "q' t = (h +++ reversepath h') t" if "0 \<le> t" "t \<le> 1" for t
using that by (simp add: joinpaths_def)
then have "path(h +++ reversepath h')"
by (auto intro: path_eq [OF \<open>path q'\<close>])
then show ?thesis
by (auto simp: \<open>path h\<close> \<open>path h'\<close>)
qed
then show ?thesis by metis
qed
then obtain l :: "'c \<Rightarrow> 'a"
where l: "\<And>y g h. \<lbrakk>path g; path_image g \<subseteq> U; pathstart g = z; pathfinish g = y;
path h; path_image h \<subseteq> C; pathstart h = a;
\<And>t. t \<in> {0..1} \<Longrightarrow> p(h t) = f(g t)\<rbrakk> \<Longrightarrow> pathfinish h = l y"
by metis
show ?thesis
proof
show pleq: "p (l y) = f y" if "y \<in> U" for y
using*[OF \<open>y \<in> U\<close>] by (metis l atLeastAtMost_iff order_refl pathfinish_def zero_le_one)
show "l z = a"
using l [of "linepath z z" z "linepath a a"] by (auto simp: assms)
show LC: "l ` U \<subseteq> C"
by (clarify dest!: *) (metis (full_types) l pathfinish_in_path_image subsetCE)
have "\<exists>T. openin (top_of_set U) T \<and> y \<in> T \<and> T \<subseteq> U \<inter> l -` X"
if X: "openin (top_of_set C) X" and "y \<in> U" "l y \<in> X" for X y
proof -
have "X \<subseteq> C"
using X openin_euclidean_subtopology_iff by blast
have "f y \<in> S"
using fim \<open>y \<in> U\<close> by blast
then obtain W \<V>
where WV: "f y \<in> W \<and> openin (top_of_set S) W \<and>
(\<Union>\<V> = C \<inter> p -` W \<and>
(\<forall>U \<in> \<V>. openin (top_of_set C) U) \<and>
pairwise disjnt \<V> \<and>
(\<forall>U \<in> \<V>. \<exists>q. homeomorphism U W p q))"
using cov by (force simp: covering_space_def)
then have "l y \<in> \<Union>\<V>"
using \<open>X \<subseteq> C\<close> pleq that by auto
then obtain W' where "l y \<in> W'" and "W' \<in> \<V>"
by blast
with WV obtain p' where opeCW': "openin (top_of_set C) W'"
and homUW': "homeomorphism W' W p p'"
by blast
then have contp': "continuous_on W p'" and p'im: "p' ` W \<subseteq> W'"
using homUW' homeomorphism_image2 homeomorphism_cont2 by fastforce+
obtain V where "y \<in> V" "y \<in> U" and fimW: "f ` V \<subseteq> W" "V \<subseteq> U"
and "path_connected V" and opeUV: "openin (top_of_set U) V"
proof -
have "openin (top_of_set U) (U \<inter> f -` W)"
using WV contf continuous_on_open_gen fim by auto
then show ?thesis
using U WV
apply (auto simp: locally_path_connected)
apply (drule_tac x="U \<inter> f -` W" in spec)
apply (drule_tac x=y in spec)
apply (auto simp: \<open>y \<in> U\<close> intro: that)
done
qed
have "W' \<subseteq> C" "W \<subseteq> S"
using opeCW' WV openin_imp_subset by auto
have p'im: "p' ` W \<subseteq> W'"
using homUW' homeomorphism_image2 by fastforce
show ?thesis
proof (intro exI conjI)
have "openin (top_of_set S) (W \<inter> p' -` (W' \<inter> X))"
proof (rule openin_trans)
show "openin (top_of_set W) (W \<inter> p' -` (W' \<inter> X))"
apply (rule continuous_openin_preimage [OF contp' p'im])
using X \<open>W' \<subseteq> C\<close> apply (auto simp: openin_open)
done
show "openin (top_of_set S) W"
using WV by blast
qed
then show "openin (top_of_set U) (V \<inter> (U \<inter> (f -` (W \<inter> (p' -` (W' \<inter> X))))))"
by (blast intro: opeUV openin_subtopology_self continuous_openin_preimage [OF contf fim])
have "p' (f y) \<in> X"
using \<open>l y \<in> W'\<close> homeomorphism_apply1 [OF homUW'] pleq \<open>y \<in> U\<close> \<open>l y \<in> X\<close> by fastforce
then show "y \<in> V \<inter> (U \<inter> f -` (W \<inter> p' -` (W' \<inter> X)))"
using \<open>y \<in> U\<close> \<open>y \<in> V\<close> WV p'im by auto
show "V \<inter> (U \<inter> f -` (W \<inter> p' -` (W' \<inter> X))) \<subseteq> U \<inter> l -` X"
proof (intro subsetI IntI; clarify)
fix y'
assume y': "y' \<in> V" "y' \<in> U" "f y' \<in> W" "p' (f y') \<in> W'" "p' (f y') \<in> X"
then obtain \<gamma> where "path \<gamma>" "path_image \<gamma> \<subseteq> V" "pathstart \<gamma> = y" "pathfinish \<gamma> = y'"
by (meson \<open>path_connected V\<close> \<open>y \<in> V\<close> path_connected_def)
obtain pp qq where "path pp" "path_image pp \<subseteq> U"
"pathstart pp = z" "pathfinish pp = y"
"path qq" "path_image qq \<subseteq> C" "pathstart qq = a"
and pqqeq: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(qq t) = f(pp t)"
using*[OF \<open>y \<in> U\<close>] by blast
have finW: "\<And>x. \<lbrakk>0 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> f (\<gamma> x) \<in> W"
using \<open>path_image \<gamma> \<subseteq> V\<close> by (auto simp: image_subset_iff path_image_def fimW [THEN subsetD])
have "pathfinish (qq +++ (p' \<circ> f \<circ> \<gamma>)) = l y'"
proof (rule l [of "pp +++ \<gamma>" y' "qq +++ (p' \<circ> f \<circ> \<gamma>)"])
show "path (pp +++ \<gamma>)"
by (simp add: \<open>path \<gamma>\<close> \<open>path pp\<close> \<open>pathfinish pp = y\<close> \<open>pathstart \<gamma> = y\<close>)
show "path_image (pp +++ \<gamma>) \<subseteq> U"
using \<open>V \<subseteq> U\<close> \<open>path_image \<gamma> \<subseteq> V\<close> \<open>path_image pp \<subseteq> U\<close> not_in_path_image_join by blast
show "pathstart (pp +++ \<gamma>) = z"
by (simp add: \<open>pathstart pp = z\<close>)
show "pathfinish (pp +++ \<gamma>) = y'"
by (simp add: \<open>pathfinish \<gamma> = y'\<close>)
have paqq: "pathfinish qq = pathstart (p' \<circ> f \<circ> \<gamma>)"
apply (simp add: \<open>pathstart \<gamma> = y\<close> pathstart_compose)
apply (metis (mono_tags, lifting) \<open>l y \<in> W'\<close> \<open>path pp\<close> \<open>path qq\<close> \<open>path_image pp \<subseteq> U\<close> \<open>path_image qq \<subseteq> C\<close>
\<open>pathfinish pp = y\<close> \<open>pathstart pp = z\<close> \<open>pathstart qq = a\<close>
homeomorphism_apply1 [OF homUW'] l pleq pqqeq \<open>y \<in> U\<close>)
done
have "continuous_on (path_image \<gamma>) (p' \<circ> f)"
proof (rule continuous_on_compose)
show "continuous_on (path_image \<gamma>) f"
using \<open>path_image \<gamma> \<subseteq> V\<close> \<open>V \<subseteq> U\<close> contf continuous_on_subset by blast
show "continuous_on (f ` path_image \<gamma>) p'"
apply (rule continuous_on_subset [OF contp'])
apply (auto simp: path_image_def pathfinish_def pathstart_def finW)
done
qed
then show "path (qq +++ (p' \<circ> f \<circ> \<gamma>))"
using \<open>path \<gamma>\<close> \<open>path qq\<close> paqq path_continuous_image path_join_imp by blast
show "path_image (qq +++ (p' \<circ> f \<circ> \<gamma>)) \<subseteq> C"
apply (rule subset_path_image_join)
apply (simp add: \<open>path_image qq \<subseteq> C\<close>)
by (metis \<open>W' \<subseteq> C\<close> \<open>path_image \<gamma> \<subseteq> V\<close> dual_order.trans fimW(1) image_comp image_mono p'im path_image_compose)
show "pathstart (qq +++ (p' \<circ> f \<circ> \<gamma>)) = a"
by (simp add: \<open>pathstart qq = a\<close>)
show "p ((qq +++ (p' \<circ> f \<circ> \<gamma>)) \<xi>) = f ((pp +++ \<gamma>) \<xi>)" if \<xi>: "\<xi> \<in> {0..1}" for \<xi>
proof (simp add: joinpaths_def, safe)
show "p (qq (2*\<xi>)) = f (pp (2*\<xi>))" if "\<xi>*2 \<le> 1"
using \<open>\<xi> \<in> {0..1}\<close> pqqeq that by auto
show "p (p' (f (\<gamma> (2*\<xi> - 1)))) = f (\<gamma> (2*\<xi> - 1))" if "\<not> \<xi>*2 \<le> 1"
apply (rule homeomorphism_apply2 [OF homUW' finW])
using that \<xi> by auto
qed
qed
with \<open>pathfinish \<gamma> = y'\<close> \<open>p' (f y') \<in> X\<close> show "y' \<in> l -` X"
unfolding pathfinish_join by (simp add: pathfinish_def)
qed
qed
qed
then show "continuous_on U l"
by (metis IntD1 IntD2 vimage_eq openin_subopen continuous_on_open_gen [OF LC])
qed
qed
corollary covering_space_lift_stronger:
fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
and f :: "'c::real_normed_vector \<Rightarrow> 'b"
assumes cov: "covering_space C p S" "a \<in> C" "z \<in> U"
and U: "path_connected U" "locally path_connected U"
and contf: "continuous_on U f" and fim: "f ` U \<subseteq> S"
and feq: "f z = p a"
and hom: "\<And>r. \<lbrakk>path r; path_image r \<subseteq> U; pathstart r = z; pathfinish r = z\<rbrakk>
\<Longrightarrow> \<exists>b. homotopic_paths S (f \<circ> r) (linepath b b)"
obtains g where "continuous_on U g" "g ` U \<subseteq> C" "g z = a" "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y"
proof (rule covering_space_lift_general [OF cov U contf fim feq])
fix r
assume "path r" "path_image r \<subseteq> U" "pathstart r = z" "pathfinish r = z"
then obtain b where b: "homotopic_paths S (f \<circ> r) (linepath b b)"
using hom by blast
then have "f (pathstart r) = b"
by (metis homotopic_paths_imp_pathstart pathstart_compose pathstart_linepath)
then have "homotopic_paths S (f \<circ> r) (linepath (f z) (f z))"
by (simp add: b \<open>pathstart r = z\<close>)
then have "homotopic_paths S (f \<circ> r) (p \<circ> linepath a a)"
by (simp add: o_def feq linepath_def)
then show "\<exists>q. path q \<and>
path_image q \<subseteq> C \<and>
pathstart q = a \<and> pathfinish q = a \<and> homotopic_paths S (f \<circ> r) (p \<circ> q)"
by (force simp: \<open>a \<in> C\<close>)
qed auto
corollary covering_space_lift_strong:
fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
and f :: "'c::real_normed_vector \<Rightarrow> 'b"
assumes cov: "covering_space C p S" "a \<in> C" "z \<in> U"
and scU: "simply_connected U" and lpcU: "locally path_connected U"
and contf: "continuous_on U f" and fim: "f ` U \<subseteq> S"
and feq: "f z = p a"
obtains g where "continuous_on U g" "g ` U \<subseteq> C" "g z = a" "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y"
proof (rule covering_space_lift_stronger [OF cov _ lpcU contf fim feq])
show "path_connected U"
using scU simply_connected_eq_contractible_loop_some by blast
fix r
assume r: "path r" "path_image r \<subseteq> U" "pathstart r = z" "pathfinish r = z"
have "linepath (f z) (f z) = f \<circ> linepath z z"
by (simp add: o_def linepath_def)
then have "homotopic_paths S (f \<circ> r) (linepath (f z) (f z))"
by (metis r contf fim homotopic_paths_continuous_image scU simply_connected_eq_contractible_path)
then show "\<exists>b. homotopic_paths S (f \<circ> r) (linepath b b)"
by blast
qed blast
corollary covering_space_lift:
fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
and f :: "'c::real_normed_vector \<Rightarrow> 'b"
assumes cov: "covering_space C p S"
and U: "simply_connected U" "locally path_connected U"
and contf: "continuous_on U f" and fim: "f ` U \<subseteq> S"
obtains g where "continuous_on U g" "g ` U \<subseteq> C" "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y"
proof (cases "U = {}")
case True
with that show ?thesis by auto
next
case False
then obtain z where "z \<in> U" by blast
then obtain a where "a \<in> C" "f z = p a"
by (metis cov covering_space_imp_surjective fim image_iff image_subset_iff)
then show ?thesis
by (metis that covering_space_lift_strong [OF cov _ \<open>z \<in> U\<close> U contf fim])
qed
subsection\<^marker>\<open>tag unimportant\<close> \<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: 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 "\<dots> 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 "\<dots> homeomorphic (path_image h)"
by (meson assms homeomorphic_def homeomorphism_arc)
finally show ?thesis .
qed
end
diff --git a/src/HOL/Analysis/Homotopy.thy b/src/HOL/Analysis/Homotopy.thy
--- a/src/HOL/Analysis/Homotopy.thy
+++ b/src/HOL/Analysis/Homotopy.thy
@@ -1,5709 +1,5709 @@
(* Title: HOL/Analysis/Path_Connected.thy
Authors: LC Paulson and Robert Himmelmann (TU Muenchen), based on material from HOL Light
*)
section \<open>Homotopy of Maps\<close>
theory Homotopy
imports Path_Connected Continuum_Not_Denumerable Product_Topology
begin
definition\<^marker>\<open>tag important\<close> homotopic_with
where
"homotopic_with P X Y f g \<equiv>
(\<exists>h. continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y h \<and>
(\<forall>x. h(0, x) = f x) \<and>
(\<forall>x. h(1, x) = g x) \<and>
(\<forall>t \<in> {0..1}. P(\<lambda>x. h(t,x))))"
text\<open>\<open>p\<close>, \<open>q\<close> are functions \<open>X \<rightarrow> Y\<close>, and the property \<open>P\<close> restricts all intermediate maps.
We often just want to require that \<open>P\<close> fixes some subset, but to include the case of a loop homotopy,
it is convenient to have a general property \<open>P\<close>.\<close>
abbreviation homotopic_with_canon ::
"[('a::topological_space \<Rightarrow> 'b::topological_space) \<Rightarrow> bool, 'a set, 'b set, 'a \<Rightarrow> 'b, 'a \<Rightarrow> 'b] \<Rightarrow> bool"
where
"homotopic_with_canon P S T p q \<equiv> homotopic_with P (top_of_set S) (top_of_set T) p q"
lemma split_01: "{0..1::real} = {0..1/2} \<union> {1/2..1}"
by force
lemma split_01_prod: "{0..1::real} \<times> X = ({0..1/2} \<times> X) \<union> ({1/2..1} \<times> X)"
by force
lemma image_Pair_const: "(\<lambda>x. (x, c)) ` A = A \<times> {c}"
by auto
lemma fst_o_paired [simp]: "fst \<circ> (\<lambda>(x,y). (f x y, g x y)) = (\<lambda>(x,y). f x y)"
by auto
lemma snd_o_paired [simp]: "snd \<circ> (\<lambda>(x,y). (f x y, g x y)) = (\<lambda>(x,y). g x y)"
by auto
lemma continuous_on_o_Pair: "\<lbrakk>continuous_on (T \<times> X) h; t \<in> T\<rbrakk> \<Longrightarrow> continuous_on X (h \<circ> Pair t)"
by (fast intro: continuous_intros elim!: continuous_on_subset)
lemma continuous_map_o_Pair:
assumes h: "continuous_map (prod_topology X Y) Z h" and t: "t \<in> topspace X"
shows "continuous_map Y Z (h \<circ> Pair t)"
apply (intro continuous_map_compose [OF _ h] continuous_map_id [unfolded id_def] continuous_intros)
apply (simp add: t)
done
subsection\<^marker>\<open>tag unimportant\<close>\<open>Trivial properties\<close>
text \<open>We often want to just localize the ending function equality or whatever.\<close>
text\<^marker>\<open>tag important\<close> \<open>%whitespace\<close>
proposition homotopic_with:
assumes "\<And>h k. (\<And>x. x \<in> topspace X \<Longrightarrow> h x = k x) \<Longrightarrow> (P h \<longleftrightarrow> P k)"
shows "homotopic_with P X Y p q \<longleftrightarrow>
(\<exists>h. continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y h \<and>
(\<forall>x \<in> topspace X. h(0,x) = p x) \<and>
(\<forall>x \<in> topspace X. h(1,x) = q x) \<and>
(\<forall>t \<in> {0..1}. P(\<lambda>x. h(t, x))))"
unfolding homotopic_with_def
apply (rule iffI, blast, clarify)
apply (rule_tac x="\<lambda>(u,v). if v \<in> topspace X then h(u,v) else if u = 0 then p v else q v" in exI)
apply auto
using continuous_map_eq apply fastforce
apply (drule_tac x=t in bspec, force)
apply (subst assms; simp)
done
lemma homotopic_with_mono:
assumes hom: "homotopic_with P X Y f g"
and Q: "\<And>h. \<lbrakk>continuous_map X Y h; P h\<rbrakk> \<Longrightarrow> Q h"
shows "homotopic_with Q X Y f g"
using hom
apply (simp add: homotopic_with_def)
apply (erule ex_forward)
apply (force simp: o_def dest: continuous_map_o_Pair intro: Q)
done
lemma homotopic_with_imp_continuous_maps:
assumes "homotopic_with P X Y f g"
shows "continuous_map X Y f \<and> continuous_map X Y g"
proof -
obtain h
where conth: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y h"
and h: "\<forall>x. h (0, x) = f x" "\<forall>x. h (1, x) = g x"
using assms by (auto simp: homotopic_with_def)
have *: "t \<in> {0..1} \<Longrightarrow> continuous_map X Y (h \<circ> (\<lambda>x. (t,x)))" for t
by (rule continuous_map_compose [OF _ conth]) (simp add: o_def continuous_map_pairwise)
show ?thesis
using h *[of 0] *[of 1] by (simp add: continuous_map_eq)
qed
lemma homotopic_with_imp_continuous:
assumes "homotopic_with_canon P X Y f g"
shows "continuous_on X f \<and> continuous_on X g"
by (meson assms continuous_map_subtopology_eu homotopic_with_imp_continuous_maps)
lemma homotopic_with_imp_property:
assumes "homotopic_with P X Y f g"
shows "P f \<and> P g"
proof
obtain h where h: "\<And>x. h(0, x) = f x" "\<And>x. h(1, x) = g x" and P: "\<And>t. t \<in> {0..1::real} \<Longrightarrow> P(\<lambda>x. h(t,x))"
using assms by (force simp: homotopic_with_def)
show "P f" "P g"
using P [of 0] P [of 1] by (force simp: h)+
qed
lemma homotopic_with_equal:
assumes "P f" "P g" and contf: "continuous_map X Y f" and fg: "\<And>x. x \<in> topspace X \<Longrightarrow> f x = g x"
shows "homotopic_with P X Y f g"
unfolding homotopic_with_def
proof (intro exI conjI allI ballI)
let ?h = "\<lambda>(t::real,x). if t = 1 then g x else f x"
show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y ?h"
proof (rule continuous_map_eq)
show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y (f \<circ> snd)"
by (simp add: contf continuous_map_of_snd)
qed (auto simp: fg)
show "P (\<lambda>x. ?h (t, x))" if "t \<in> {0..1}" for t
by (cases "t = 1") (simp_all add: assms)
qed auto
lemma homotopic_with_imp_subset1:
"homotopic_with_canon P X Y f g \<Longrightarrow> f ` X \<subseteq> Y"
by (simp add: homotopic_with_def image_subset_iff) (metis atLeastAtMost_iff order_refl zero_le_one)
lemma homotopic_with_imp_subset2:
"homotopic_with_canon P X Y f g \<Longrightarrow> g ` X \<subseteq> Y"
by (simp add: homotopic_with_def image_subset_iff) (metis atLeastAtMost_iff order_refl zero_le_one)
lemma homotopic_with_subset_left:
"\<lbrakk>homotopic_with_canon P X Y f g; Z \<subseteq> X\<rbrakk> \<Longrightarrow> homotopic_with_canon P Z Y f g"
apply (simp add: homotopic_with_def)
apply (fast elim!: continuous_on_subset ex_forward)
done
lemma homotopic_with_subset_right:
"\<lbrakk>homotopic_with_canon P X Y f g; Y \<subseteq> Z\<rbrakk> \<Longrightarrow> homotopic_with_canon P X Z f g"
apply (simp add: homotopic_with_def)
apply (fast elim!: continuous_on_subset ex_forward)
done
subsection\<open>Homotopy with P is an equivalence relation\<close>
text \<open>(on continuous functions mapping X into Y that satisfy P, though this only affects reflexivity)\<close>
lemma homotopic_with_refl [simp]: "homotopic_with P X Y f f \<longleftrightarrow> continuous_map X Y f \<and> P f"
by (auto simp: homotopic_with_imp_continuous_maps intro: homotopic_with_equal dest: homotopic_with_imp_property)
lemma homotopic_with_symD:
assumes "homotopic_with P X Y f g"
shows "homotopic_with P X Y g f"
proof -
let ?I01 = "subtopology euclideanreal {0..1}"
let ?j = "\<lambda>y. (1 - fst y, snd y)"
have 1: "continuous_map (prod_topology ?I01 X) (prod_topology euclideanreal X) ?j"
apply (intro continuous_intros)
apply (simp_all add: prod_topology_subtopology continuous_map_from_subtopology [OF continuous_map_fst])
done
have *: "continuous_map (prod_topology ?I01 X) (prod_topology ?I01 X) ?j"
proof -
have "continuous_map (prod_topology ?I01 X) (subtopology (prod_topology euclideanreal X) ({0..1} \<times> topspace X)) ?j"
by (simp add: continuous_map_into_subtopology [OF 1] image_subset_iff)
then show ?thesis
by (simp add: prod_topology_subtopology(1))
qed
show ?thesis
using assms
apply (clarsimp simp add: homotopic_with_def)
apply (rename_tac h)
apply (rule_tac x="h \<circ> (\<lambda>y. (1 - fst y, snd y))" in exI)
apply (simp add: continuous_map_compose [OF *])
done
qed
lemma homotopic_with_sym:
"homotopic_with P X Y f g \<longleftrightarrow> homotopic_with P X Y g f"
by (metis homotopic_with_symD)
proposition homotopic_with_trans:
assumes "homotopic_with P X Y f g" "homotopic_with P X Y g h"
shows "homotopic_with P X Y f h"
proof -
let ?X01 = "prod_topology (subtopology euclideanreal {0..1}) X"
obtain k1 k2
where contk1: "continuous_map ?X01 Y k1" and contk2: "continuous_map ?X01 Y k2"
and k12: "\<forall>x. k1 (1, x) = g x" "\<forall>x. k2 (0, x) = g x"
"\<forall>x. k1 (0, x) = f x" "\<forall>x. k2 (1, x) = h x"
and P: "\<forall>t\<in>{0..1}. P (\<lambda>x. k1 (t, x))" "\<forall>t\<in>{0..1}. P (\<lambda>x. k2 (t, x))"
using assms by (auto simp: homotopic_with_def)
define k where "k \<equiv> \<lambda>y. if fst y \<le> 1/2
then (k1 \<circ> (\<lambda>x. (2 *\<^sub>R fst x, snd x))) y
else (k2 \<circ> (\<lambda>x. (2 *\<^sub>R fst x -1, snd x))) y"
have keq: "k1 (2 * u, v) = k2 (2 * u -1, v)" if "u = 1/2" for u v
by (simp add: k12 that)
show ?thesis
unfolding homotopic_with_def
proof (intro exI conjI)
show "continuous_map ?X01 Y k"
unfolding k_def
proof (rule continuous_map_cases_le)
show fst: "continuous_map ?X01 euclideanreal fst"
using continuous_map_fst continuous_map_in_subtopology by blast
show "continuous_map ?X01 euclideanreal (\<lambda>x. 1/2)"
by simp
show "continuous_map (subtopology ?X01 {y \<in> topspace ?X01. fst y \<le> 1/2}) Y
(k1 \<circ> (\<lambda>x. (2 *\<^sub>R fst x, snd x)))"
apply (rule fst continuous_map_compose [OF _ contk1] continuous_intros continuous_map_into_subtopology | simp)+
apply (intro continuous_intros fst continuous_map_from_subtopology)
apply (force simp: prod_topology_subtopology)
using continuous_map_snd continuous_map_from_subtopology by blast
show "continuous_map (subtopology ?X01 {y \<in> topspace ?X01. 1/2 \<le> fst y}) Y
(k2 \<circ> (\<lambda>x. (2 *\<^sub>R fst x -1, snd x)))"
apply (rule fst continuous_map_compose [OF _ contk2] continuous_intros continuous_map_into_subtopology | simp)+
apply (rule continuous_intros fst continuous_map_from_subtopology | simp)+
apply (force simp: prod_topology_subtopology)
using continuous_map_snd continuous_map_from_subtopology by blast
show "(k1 \<circ> (\<lambda>x. (2 *\<^sub>R fst x, snd x))) y = (k2 \<circ> (\<lambda>x. (2 *\<^sub>R fst x -1, snd x))) y"
if "y \<in> topspace ?X01" and "fst y = 1/2" for y
using that by (simp add: keq)
qed
show "\<forall>x. k (0, x) = f x"
by (simp add: k12 k_def)
show "\<forall>x. k (1, x) = h x"
by (simp add: k12 k_def)
show "\<forall>t\<in>{0..1}. P (\<lambda>x. k (t, x))"
using P
apply (clarsimp simp add: k_def)
apply (case_tac "t \<le> 1/2", auto)
done
qed
qed
lemma homotopic_with_id2:
"(\<And>x. x \<in> topspace X \<Longrightarrow> g (f x) = x) \<Longrightarrow> homotopic_with (\<lambda>x. True) X X (g \<circ> f) id"
by (metis comp_apply continuous_map_id eq_id_iff homotopic_with_equal homotopic_with_symD)
subsection\<open>Continuity lemmas\<close>
lemma homotopic_with_compose_continuous_map_left:
"\<lbrakk>homotopic_with p X1 X2 f g; continuous_map X2 X3 h; \<And>j. p j \<Longrightarrow> q(h \<circ> j)\<rbrakk>
\<Longrightarrow> homotopic_with q X1 X3 (h \<circ> f) (h \<circ> g)"
unfolding homotopic_with_def
apply clarify
apply (rename_tac k)
apply (rule_tac x="h \<circ> k" in exI)
apply (rule conjI continuous_map_compose | simp add: o_def)+
done
lemma homotopic_compose_continuous_map_left:
"\<lbrakk>homotopic_with (\<lambda>k. True) X1 X2 f g; continuous_map X2 X3 h\<rbrakk>
\<Longrightarrow> homotopic_with (\<lambda>k. True) X1 X3 (h \<circ> f) (h \<circ> g)"
by (simp add: homotopic_with_compose_continuous_map_left)
lemma homotopic_with_compose_continuous_map_right:
assumes hom: "homotopic_with p X2 X3 f g" and conth: "continuous_map X1 X2 h"
and q: "\<And>j. p j \<Longrightarrow> q(j \<circ> h)"
shows "homotopic_with q X1 X3 (f \<circ> h) (g \<circ> h)"
proof -
obtain k
where contk: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X2) X3 k"
and k: "\<forall>x. k (0, x) = f x" "\<forall>x. k (1, x) = g x" and p: "\<And>t. t\<in>{0..1} \<Longrightarrow> p (\<lambda>x. k (t, x))"
using hom unfolding homotopic_with_def by blast
have hsnd: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X1) X2 (h \<circ> snd)"
by (rule continuous_map_compose [OF continuous_map_snd conth])
let ?h = "k \<circ> (\<lambda>(t,x). (t,h x))"
show ?thesis
unfolding homotopic_with_def
proof (intro exI conjI allI ballI)
have "continuous_map (prod_topology (top_of_set {0..1}) X1)
(prod_topology (top_of_set {0..1::real}) X2) (\<lambda>(t, x). (t, h x))"
by (metis (mono_tags, lifting) case_prod_beta' comp_def continuous_map_eq continuous_map_fst continuous_map_pairedI hsnd)
then show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X1) X3 ?h"
by (intro conjI continuous_map_compose [OF _ contk])
show "q (\<lambda>x. ?h (t, x))" if "t \<in> {0..1}" for t
using q [OF p [OF that]] by (simp add: o_def)
qed (auto simp: k)
qed
lemma homotopic_compose_continuous_map_right:
"\<lbrakk>homotopic_with (\<lambda>k. True) X2 X3 f g; continuous_map X1 X2 h\<rbrakk>
\<Longrightarrow> homotopic_with (\<lambda>k. True) X1 X3 (f \<circ> h) (g \<circ> h)"
by (meson homotopic_with_compose_continuous_map_right)
corollary homotopic_compose:
shows "\<lbrakk>homotopic_with (\<lambda>x. True) X Y f f'; homotopic_with (\<lambda>x. True) Y Z g g'\<rbrakk>
\<Longrightarrow> homotopic_with (\<lambda>x. True) X Z (g \<circ> f) (g' \<circ> f')"
apply (rule homotopic_with_trans [where g = "g \<circ> f'"])
apply (simp add: homotopic_compose_continuous_map_left homotopic_with_imp_continuous_maps)
by (simp add: homotopic_compose_continuous_map_right homotopic_with_imp_continuous_maps)
proposition homotopic_with_compose_continuous_right:
"\<lbrakk>homotopic_with_canon (\<lambda>f. p (f \<circ> h)) X Y f g; continuous_on W h; h ` W \<subseteq> X\<rbrakk>
\<Longrightarrow> homotopic_with_canon p W Y (f \<circ> h) (g \<circ> h)"
apply (clarsimp simp add: homotopic_with_def)
apply (rename_tac k)
apply (rule_tac x="k \<circ> (\<lambda>y. (fst y, h (snd y)))" in exI)
apply (rule conjI continuous_intros continuous_on_compose [where f=snd and g=h, unfolded o_def] | simp)+
apply (erule continuous_on_subset)
apply (fastforce simp: o_def)+
done
proposition homotopic_compose_continuous_right:
"\<lbrakk>homotopic_with_canon (\<lambda>f. True) X Y f g; continuous_on W h; h ` W \<subseteq> X\<rbrakk>
\<Longrightarrow> homotopic_with_canon (\<lambda>f. True) W Y (f \<circ> h) (g \<circ> h)"
using homotopic_with_compose_continuous_right by fastforce
proposition homotopic_with_compose_continuous_left:
"\<lbrakk>homotopic_with_canon (\<lambda>f. p (h \<circ> f)) X Y f g; continuous_on Y h; h ` Y \<subseteq> Z\<rbrakk>
\<Longrightarrow> homotopic_with_canon p X Z (h \<circ> f) (h \<circ> g)"
apply (clarsimp simp add: homotopic_with_def)
apply (rename_tac k)
apply (rule_tac x="h \<circ> k" in exI)
apply (rule conjI continuous_intros continuous_on_compose [where f=snd and g=h, unfolded o_def] | simp)+
apply (erule continuous_on_subset)
apply (fastforce simp: o_def)+
done
proposition homotopic_compose_continuous_left:
"\<lbrakk>homotopic_with_canon (\<lambda>_. True) X Y f g;
continuous_on Y h; h ` Y \<subseteq> Z\<rbrakk>
\<Longrightarrow> homotopic_with_canon (\<lambda>f. True) X Z (h \<circ> f) (h \<circ> g)"
using homotopic_with_compose_continuous_left by fastforce
lemma homotopic_from_subtopology:
"homotopic_with P X X' f g \<Longrightarrow> homotopic_with P (subtopology X s) X' f g"
unfolding homotopic_with_def
apply (erule ex_forward)
by (simp add: continuous_map_from_subtopology prod_topology_subtopology(2))
lemma homotopic_on_emptyI:
assumes "topspace X = {}" "P f" "P g"
shows "homotopic_with P X X' f g"
unfolding homotopic_with_def
proof (intro exI conjI ballI)
show "P (\<lambda>x. (\<lambda>(t,x). if t = 0 then f x else g x) (t, x))" if "t \<in> {0..1}" for t::real
by (cases "t = 0", auto simp: assms)
qed (auto simp: continuous_map_atin assms)
lemma homotopic_on_empty:
"topspace X = {} \<Longrightarrow> (homotopic_with P X X' f g \<longleftrightarrow> P f \<and> P g)"
using homotopic_on_emptyI homotopic_with_imp_property by metis
lemma homotopic_with_canon_on_empty [simp]: "homotopic_with_canon (\<lambda>x. True) {} t f g"
by (auto intro: homotopic_with_equal)
lemma homotopic_constant_maps:
"homotopic_with (\<lambda>x. True) X X' (\<lambda>x. a) (\<lambda>x. b) \<longleftrightarrow>
topspace X = {} \<or> path_component_of X' a b" (is "?lhs = ?rhs")
proof (cases "topspace X = {}")
case False
then obtain c where c: "c \<in> topspace X"
by blast
have "\<exists>g. continuous_map (top_of_set {0..1::real}) X' g \<and> g 0 = a \<and> g 1 = b"
if "x \<in> topspace X" and hom: "homotopic_with (\<lambda>x. True) X X' (\<lambda>x. a) (\<lambda>x. b)" for x
proof -
obtain h :: "real \<times> 'a \<Rightarrow> 'b"
where conth: "continuous_map (prod_topology (top_of_set {0..1}) X) X' h"
and h: "\<And>x. h (0, x) = a" "\<And>x. h (1, x) = b"
using hom by (auto simp: homotopic_with_def)
have cont: "continuous_map (top_of_set {0..1}) X' (h \<circ> (\<lambda>t. (t, c)))"
apply (rule continuous_map_compose [OF _ conth])
apply (rule continuous_intros c | simp)+
done
then show ?thesis
by (force simp: h)
qed
moreover have "homotopic_with (\<lambda>x. True) X X' (\<lambda>x. g 0) (\<lambda>x. g 1)"
if "x \<in> topspace X" "a = g 0" "b = g 1" "continuous_map (top_of_set {0..1}) X' g"
for x and g :: "real \<Rightarrow> 'b"
unfolding homotopic_with_def
by (force intro!: continuous_map_compose continuous_intros c that)
ultimately show ?thesis
using False by (auto simp: path_component_of_def pathin_def)
qed (simp add: homotopic_on_empty)
proposition homotopic_with_eq:
assumes h: "homotopic_with P X Y f g"
and f': "\<And>x. x \<in> topspace X \<Longrightarrow> f' x = f x"
and g': "\<And>x. x \<in> topspace X \<Longrightarrow> g' x = g x"
and P: "(\<And>h k. (\<And>x. x \<in> topspace X \<Longrightarrow> h x = k x) \<Longrightarrow> P h \<longleftrightarrow> P k)"
shows "homotopic_with P X Y f' g'"
using h unfolding homotopic_with_def
apply safe
apply (rule_tac x="\<lambda>(u,v). if v \<in> topspace X then h(u,v) else if u = 0 then f' v else g' v" in exI)
apply (simp add: f' g', safe)
apply (fastforce intro: continuous_map_eq)
apply (subst P; fastforce)
done
lemma homotopic_with_prod_topology:
assumes "homotopic_with p X1 Y1 f f'" and "homotopic_with q X2 Y2 g g'"
and r: "\<And>i j. \<lbrakk>p i; q j\<rbrakk> \<Longrightarrow> r(\<lambda>(x,y). (i x, j y))"
shows "homotopic_with r (prod_topology X1 X2) (prod_topology Y1 Y2)
(\<lambda>z. (f(fst z),g(snd z))) (\<lambda>z. (f'(fst z), g'(snd z)))"
proof -
obtain h
where h: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X1) Y1 h"
and h0: "\<And>x. h (0, x) = f x"
and h1: "\<And>x. h (1, x) = f' x"
and p: "\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> p (\<lambda>x. h (t,x))"
using assms unfolding homotopic_with_def by auto
obtain k
where k: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X2) Y2 k"
and k0: "\<And>x. k (0, x) = g x"
and k1: "\<And>x. k (1, x) = g' x"
and q: "\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> q (\<lambda>x. k (t,x))"
using assms unfolding homotopic_with_def by auto
let ?hk = "\<lambda>(t,x,y). (h(t,x), k(t,y))"
show ?thesis
unfolding homotopic_with_def
proof (intro conjI allI exI)
show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) (prod_topology X1 X2))
(prod_topology Y1 Y2) ?hk"
unfolding continuous_map_pairwise case_prod_unfold
by (rule conjI continuous_map_pairedI continuous_intros continuous_map_id [unfolded id_def]
continuous_map_fst_of [unfolded o_def] continuous_map_snd_of [unfolded o_def]
continuous_map_compose [OF _ h, unfolded o_def]
continuous_map_compose [OF _ k, unfolded o_def])+
next
fix x
show "?hk (0, x) = (f (fst x), g (snd x))" "?hk (1, x) = (f' (fst x), g' (snd x))"
by (auto simp: case_prod_beta h0 k0 h1 k1)
qed (auto simp: p q r)
qed
lemma homotopic_with_product_topology:
assumes ht: "\<And>i. i \<in> I \<Longrightarrow> homotopic_with (p i) (X i) (Y i) (f i) (g i)"
and pq: "\<And>h. (\<And>i. i \<in> I \<Longrightarrow> p i (h i)) \<Longrightarrow> q(\<lambda>x. (\<lambda>i\<in>I. h i (x i)))"
shows "homotopic_with q (product_topology X I) (product_topology Y I)
(\<lambda>z. (\<lambda>i\<in>I. (f i) (z i))) (\<lambda>z. (\<lambda>i\<in>I. (g i) (z i)))"
proof -
obtain h
where h: "\<And>i. i \<in> I \<Longrightarrow> continuous_map (prod_topology (subtopology euclideanreal {0..1}) (X i)) (Y i) (h i)"
and h0: "\<And>i x. i \<in> I \<Longrightarrow> h i (0, x) = f i x"
and h1: "\<And>i x. i \<in> I \<Longrightarrow> h i (1, x) = g i x"
and p: "\<And>i t. \<lbrakk>i \<in> I; t \<in> {0..1}\<rbrakk> \<Longrightarrow> p i (\<lambda>x. h i (t,x))"
using ht unfolding homotopic_with_def by metis
show ?thesis
unfolding homotopic_with_def
proof (intro conjI allI exI)
let ?h = "\<lambda>(t,z). \<lambda>i\<in>I. h i (t,z i)"
have "continuous_map (prod_topology (subtopology euclideanreal {0..1}) (product_topology X I))
(Y i) (\<lambda>x. h i (fst x, snd x i))" if "i \<in> I" for i
unfolding continuous_map_pairwise case_prod_unfold
apply (intro conjI that continuous_intros continuous_map_compose [OF _ h, unfolded o_def])
using continuous_map_componentwise continuous_map_snd that apply fastforce
done
then show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) (product_topology X I))
(product_topology Y I) ?h"
by (auto simp: continuous_map_componentwise case_prod_beta)
show "?h (0, x) = (\<lambda>i\<in>I. f i (x i))" "?h (1, x) = (\<lambda>i\<in>I. g i (x i))" for x
by (auto simp: case_prod_beta h0 h1)
show "\<forall>t\<in>{0..1}. q (\<lambda>x. ?h (t, x))"
by (force intro: p pq)
qed
qed
text\<open>Homotopic triviality implicitly incorporates path-connectedness.\<close>
lemma homotopic_triviality:
shows "(\<forall>f g. continuous_on S f \<and> f ` S \<subseteq> T \<and>
continuous_on S g \<and> g ` S \<subseteq> T
\<longrightarrow> homotopic_with_canon (\<lambda>x. True) S T f g) \<longleftrightarrow>
(S = {} \<or> path_connected T) \<and>
(\<forall>f. continuous_on S f \<and> f ` S \<subseteq> T \<longrightarrow> (\<exists>c. homotopic_with_canon (\<lambda>x. True) S T f (\<lambda>x. c)))"
(is "?lhs = ?rhs")
proof (cases "S = {} \<or> T = {}")
case True then show ?thesis
by (auto simp: homotopic_on_emptyI)
next
case False show ?thesis
proof
assume LHS [rule_format]: ?lhs
have pab: "path_component T a b" if "a \<in> T" "b \<in> T" for a b
proof -
have "homotopic_with_canon (\<lambda>x. True) S T (\<lambda>x. a) (\<lambda>x. b)"
by (simp add: LHS image_subset_iff that)
then show ?thesis
using False homotopic_constant_maps [of "top_of_set S" "top_of_set T" a b] by auto
qed
moreover
have "\<exists>c. homotopic_with_canon (\<lambda>x. True) S T f (\<lambda>x. c)" if "continuous_on S f" "f ` S \<subseteq> T" for f
using False LHS continuous_on_const that by blast
ultimately show ?rhs
by (simp add: path_connected_component)
next
assume RHS: ?rhs
with False have T: "path_connected T"
by blast
show ?lhs
proof clarify
fix f g
assume "continuous_on S f" "f ` S \<subseteq> T" "continuous_on S g" "g ` S \<subseteq> T"
obtain c d where c: "homotopic_with_canon (\<lambda>x. True) S T f (\<lambda>x. c)" and d: "homotopic_with_canon (\<lambda>x. True) S T g (\<lambda>x. d)"
using False \<open>continuous_on S f\<close> \<open>f ` S \<subseteq> T\<close> RHS \<open>continuous_on S g\<close> \<open>g ` S \<subseteq> T\<close> by blast
then have "c \<in> T" "d \<in> T"
using False homotopic_with_imp_continuous_maps by fastforce+
with T have "path_component T c d"
using path_connected_component by blast
then have "homotopic_with_canon (\<lambda>x. True) S T (\<lambda>x. c) (\<lambda>x. d)"
by (simp add: homotopic_constant_maps)
with c d show "homotopic_with_canon (\<lambda>x. True) S T f g"
by (meson homotopic_with_symD homotopic_with_trans)
qed
qed
qed
subsection\<open>Homotopy of paths, maintaining the same endpoints\<close>
definition\<^marker>\<open>tag important\<close> homotopic_paths :: "['a set, real \<Rightarrow> 'a, real \<Rightarrow> 'a::topological_space] \<Rightarrow> bool"
where
"homotopic_paths s p q \<equiv>
homotopic_with_canon (\<lambda>r. pathstart r = pathstart p \<and> pathfinish r = pathfinish p) {0..1} s p q"
lemma homotopic_paths:
"homotopic_paths s p q \<longleftrightarrow>
(\<exists>h. continuous_on ({0..1} \<times> {0..1}) h \<and>
h ` ({0..1} \<times> {0..1}) \<subseteq> s \<and>
(\<forall>x \<in> {0..1}. h(0,x) = p x) \<and>
(\<forall>x \<in> {0..1}. h(1,x) = q x) \<and>
(\<forall>t \<in> {0..1::real}. pathstart(h \<circ> Pair t) = pathstart p \<and>
pathfinish(h \<circ> Pair t) = pathfinish p))"
by (auto simp: homotopic_paths_def homotopic_with pathstart_def pathfinish_def)
proposition homotopic_paths_imp_pathstart:
"homotopic_paths s p q \<Longrightarrow> pathstart p = pathstart q"
by (metis (mono_tags, lifting) homotopic_paths_def homotopic_with_imp_property)
proposition homotopic_paths_imp_pathfinish:
"homotopic_paths s p q \<Longrightarrow> pathfinish p = pathfinish q"
by (metis (mono_tags, lifting) homotopic_paths_def homotopic_with_imp_property)
lemma homotopic_paths_imp_path:
"homotopic_paths s p q \<Longrightarrow> path p \<and> path q"
using homotopic_paths_def homotopic_with_imp_continuous_maps path_def continuous_map_subtopology_eu by blast
lemma homotopic_paths_imp_subset:
"homotopic_paths s p q \<Longrightarrow> path_image p \<subseteq> s \<and> path_image q \<subseteq> s"
by (metis (mono_tags) continuous_map_subtopology_eu homotopic_paths_def homotopic_with_imp_continuous_maps path_image_def)
proposition homotopic_paths_refl [simp]: "homotopic_paths s p p \<longleftrightarrow> path p \<and> path_image p \<subseteq> s"
by (simp add: homotopic_paths_def path_def path_image_def)
proposition homotopic_paths_sym: "homotopic_paths s p q \<Longrightarrow> homotopic_paths s q p"
by (metis (mono_tags) homotopic_paths_def homotopic_paths_imp_pathfinish homotopic_paths_imp_pathstart homotopic_with_symD)
proposition homotopic_paths_sym_eq: "homotopic_paths s p q \<longleftrightarrow> homotopic_paths s q p"
by (metis homotopic_paths_sym)
proposition homotopic_paths_trans [trans]:
assumes "homotopic_paths s p q" "homotopic_paths s q r"
shows "homotopic_paths s p r"
proof -
have "pathstart q = pathstart p" "pathfinish q = pathfinish p"
using assms by (simp_all add: homotopic_paths_imp_pathstart homotopic_paths_imp_pathfinish)
then have "homotopic_with_canon (\<lambda>f. pathstart f = pathstart p \<and> pathfinish f = pathfinish p) {0..1} s q r"
using \<open>homotopic_paths s q r\<close> homotopic_paths_def by force
then show ?thesis
using assms homotopic_paths_def homotopic_with_trans by blast
qed
proposition homotopic_paths_eq:
"\<lbrakk>path p; path_image p \<subseteq> s; \<And>t. t \<in> {0..1} \<Longrightarrow> p t = q t\<rbrakk> \<Longrightarrow> homotopic_paths s p q"
apply (simp add: homotopic_paths_def)
apply (rule homotopic_with_eq)
apply (auto simp: path_def pathstart_def pathfinish_def path_image_def elim: continuous_on_eq)
done
proposition homotopic_paths_reparametrize:
assumes "path p"
and pips: "path_image p \<subseteq> s"
and contf: "continuous_on {0..1} f"
and f01:"f ` {0..1} \<subseteq> {0..1}"
and [simp]: "f(0) = 0" "f(1) = 1"
and q: "\<And>t. t \<in> {0..1} \<Longrightarrow> q(t) = p(f t)"
shows "homotopic_paths s p q"
proof -
have contp: "continuous_on {0..1} p"
by (metis \<open>path p\<close> path_def)
then have "continuous_on {0..1} (p \<circ> f)"
using contf continuous_on_compose continuous_on_subset f01 by blast
then have "path q"
by (simp add: path_def) (metis q continuous_on_cong)
have piqs: "path_image q \<subseteq> s"
by (metis (no_types, hide_lams) pips f01 image_subset_iff path_image_def q)
have fb0: "\<And>a b. \<lbrakk>0 \<le> a; a \<le> 1; 0 \<le> b; b \<le> 1\<rbrakk> \<Longrightarrow> 0 \<le> (1 - a) * f b + a * b"
using f01 by force
have fb1: "\<lbrakk>0 \<le> a; a \<le> 1; 0 \<le> b; b \<le> 1\<rbrakk> \<Longrightarrow> (1 - a) * f b + a * b \<le> 1" for a b
using f01 [THEN subsetD, of "f b"] by (simp add: convex_bound_le)
have "homotopic_paths s q p"
proof (rule homotopic_paths_trans)
show "homotopic_paths s q (p \<circ> f)"
using q by (force intro: homotopic_paths_eq [OF \<open>path q\<close> piqs])
next
show "homotopic_paths s (p \<circ> f) p"
apply (simp add: homotopic_paths_def homotopic_with_def)
apply (rule_tac x="p \<circ> (\<lambda>y. (1 - (fst y)) *\<^sub>R ((f \<circ> snd) y) + (fst y) *\<^sub>R snd y)" in exI)
apply (rule conjI contf continuous_intros continuous_on_subset [OF contp] | simp)+
using pips [unfolded path_image_def]
apply (auto simp: fb0 fb1 pathstart_def pathfinish_def)
done
qed
then show ?thesis
by (simp add: homotopic_paths_sym)
qed
lemma homotopic_paths_subset: "\<lbrakk>homotopic_paths s p q; s \<subseteq> t\<rbrakk> \<Longrightarrow> homotopic_paths t p q"
unfolding homotopic_paths by fast
text\<open> A slightly ad-hoc but useful lemma in constructing homotopies.\<close>
lemma homotopic_join_lemma:
fixes q :: "[real,real] \<Rightarrow> 'a::topological_space"
assumes p: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>y. p (fst y) (snd y))"
and q: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>y. q (fst y) (snd y))"
and pf: "\<And>t. t \<in> {0..1} \<Longrightarrow> pathfinish(p t) = pathstart(q t)"
shows "continuous_on ({0..1} \<times> {0..1}) (\<lambda>y. (p(fst y) +++ q(fst y)) (snd y))"
proof -
have 1: "(\<lambda>y. p (fst y) (2 * snd y)) = (\<lambda>y. p (fst y) (snd y)) \<circ> (\<lambda>y. (fst y, 2 * snd y))"
by (rule ext) (simp)
have 2: "(\<lambda>y. q (fst y) (2 * snd y - 1)) = (\<lambda>y. q (fst y) (snd y)) \<circ> (\<lambda>y. (fst y, 2 * snd y - 1))"
by (rule ext) (simp)
show ?thesis
apply (simp add: joinpaths_def)
apply (rule continuous_on_cases_le)
apply (simp_all only: 1 2)
apply (rule continuous_intros continuous_on_subset [OF p] continuous_on_subset [OF q] | force)+
using pf
apply (auto simp: mult.commute pathstart_def pathfinish_def)
done
qed
text\<open> Congruence properties of homotopy w.r.t. path-combining operations.\<close>
lemma homotopic_paths_reversepath_D:
assumes "homotopic_paths s p q"
shows "homotopic_paths s (reversepath p) (reversepath q)"
using assms
apply (simp add: homotopic_paths_def homotopic_with_def, clarify)
apply (rule_tac x="h \<circ> (\<lambda>x. (fst x, 1 - snd x))" in exI)
apply (rule conjI continuous_intros)+
apply (auto simp: reversepath_def pathstart_def pathfinish_def elim!: continuous_on_subset)
done
proposition homotopic_paths_reversepath:
"homotopic_paths s (reversepath p) (reversepath q) \<longleftrightarrow> homotopic_paths s p q"
using homotopic_paths_reversepath_D by force
proposition homotopic_paths_join:
"\<lbrakk>homotopic_paths s p p'; homotopic_paths s q q'; pathfinish p = pathstart q\<rbrakk> \<Longrightarrow> homotopic_paths s (p +++ q) (p' +++ q')"
apply (simp add: homotopic_paths_def homotopic_with_def, clarify)
apply (rename_tac k1 k2)
apply (rule_tac x="(\<lambda>y. ((k1 \<circ> Pair (fst y)) +++ (k2 \<circ> Pair (fst y))) (snd y))" in exI)
apply (rule conjI continuous_intros homotopic_join_lemma)+
apply (auto simp: joinpaths_def pathstart_def pathfinish_def path_image_def)
done
proposition homotopic_paths_continuous_image:
"\<lbrakk>homotopic_paths s f g; continuous_on s h; h ` s \<subseteq> t\<rbrakk> \<Longrightarrow> homotopic_paths t (h \<circ> f) (h \<circ> g)"
unfolding homotopic_paths_def
by (simp add: homotopic_with_compose_continuous_map_left pathfinish_compose pathstart_compose)
subsection\<open>Group properties for homotopy of paths\<close>
text\<^marker>\<open>tag important\<close>\<open>So taking equivalence classes under homotopy would give the fundamental group\<close>
proposition homotopic_paths_rid:
"\<lbrakk>path p; path_image p \<subseteq> s\<rbrakk> \<Longrightarrow> homotopic_paths s (p +++ linepath (pathfinish p) (pathfinish p)) p"
apply (subst homotopic_paths_sym)
apply (rule homotopic_paths_reparametrize [where f = "\<lambda>t. if t \<le> 1 / 2 then 2 *\<^sub>R t else 1"])
apply (simp_all del: le_divide_eq_numeral1)
apply (subst split_01)
apply (rule continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def)+
done
proposition homotopic_paths_lid:
"\<lbrakk>path p; path_image p \<subseteq> s\<rbrakk> \<Longrightarrow> homotopic_paths s (linepath (pathstart p) (pathstart p) +++ p) p"
using homotopic_paths_rid [of "reversepath p" s]
by (metis homotopic_paths_reversepath path_image_reversepath path_reversepath pathfinish_linepath
pathfinish_reversepath reversepath_joinpaths reversepath_linepath)
proposition homotopic_paths_assoc:
"\<lbrakk>path p; path_image p \<subseteq> s; path q; path_image q \<subseteq> s; path r; path_image r \<subseteq> s; pathfinish p = pathstart q;
pathfinish q = pathstart r\<rbrakk>
\<Longrightarrow> homotopic_paths s (p +++ (q +++ r)) ((p +++ q) +++ r)"
apply (subst homotopic_paths_sym)
apply (rule homotopic_paths_reparametrize
[where f = "\<lambda>t. if t \<le> 1 / 2 then inverse 2 *\<^sub>R t
else if t \<le> 3 / 4 then t - (1 / 4)
else 2 *\<^sub>R t - 1"])
apply (simp_all del: le_divide_eq_numeral1)
apply (simp add: subset_path_image_join)
apply (rule continuous_on_cases_1 continuous_intros)+
apply (auto simp: joinpaths_def)
done
proposition homotopic_paths_rinv:
assumes "path p" "path_image p \<subseteq> s"
shows "homotopic_paths s (p +++ reversepath p) (linepath (pathstart p) (pathstart p))"
proof -
have "continuous_on ({0..1} \<times> {0..1}) (\<lambda>x. (subpath 0 (fst x) p +++ reversepath (subpath 0 (fst x) p)) (snd x))"
using assms
apply (simp add: joinpaths_def subpath_def reversepath_def path_def del: le_divide_eq_numeral1)
apply (rule continuous_on_cases_le)
apply (rule_tac [2] continuous_on_compose [of _ _ p, unfolded o_def])
apply (rule continuous_on_compose [of _ _ p, unfolded o_def])
apply (auto intro!: continuous_intros simp del: eq_divide_eq_numeral1)
apply (force elim!: continuous_on_subset simp add: mult_le_one)+
done
then show ?thesis
using assms
apply (subst homotopic_paths_sym_eq)
unfolding homotopic_paths_def homotopic_with_def
apply (rule_tac x="(\<lambda>y. (subpath 0 (fst y) p +++ reversepath(subpath 0 (fst y) p)) (snd y))" in exI)
apply (simp add: path_defs joinpaths_def subpath_def reversepath_def)
apply (force simp: mult_le_one)
done
qed
proposition homotopic_paths_linv:
assumes "path p" "path_image p \<subseteq> s"
shows "homotopic_paths s (reversepath p +++ p) (linepath (pathfinish p) (pathfinish p))"
using homotopic_paths_rinv [of "reversepath p" s] assms by simp
subsection\<open>Homotopy of loops without requiring preservation of endpoints\<close>
definition\<^marker>\<open>tag important\<close> homotopic_loops :: "'a::topological_space set \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> bool" where
"homotopic_loops s p q \<equiv>
homotopic_with_canon (\<lambda>r. pathfinish r = pathstart r) {0..1} s p q"
lemma homotopic_loops:
"homotopic_loops s p q \<longleftrightarrow>
(\<exists>h. continuous_on ({0..1::real} \<times> {0..1}) h \<and>
image h ({0..1} \<times> {0..1}) \<subseteq> s \<and>
(\<forall>x \<in> {0..1}. h(0,x) = p x) \<and>
(\<forall>x \<in> {0..1}. h(1,x) = q x) \<and>
(\<forall>t \<in> {0..1}. pathfinish(h \<circ> Pair t) = pathstart(h \<circ> Pair t)))"
by (simp add: homotopic_loops_def pathstart_def pathfinish_def homotopic_with)
proposition homotopic_loops_imp_loop:
"homotopic_loops s p q \<Longrightarrow> pathfinish p = pathstart p \<and> pathfinish q = pathstart q"
using homotopic_with_imp_property homotopic_loops_def by blast
proposition homotopic_loops_imp_path:
"homotopic_loops s p q \<Longrightarrow> path p \<and> path q"
unfolding homotopic_loops_def path_def
using homotopic_with_imp_continuous_maps continuous_map_subtopology_eu by blast
proposition homotopic_loops_imp_subset:
"homotopic_loops s p q \<Longrightarrow> path_image p \<subseteq> s \<and> path_image q \<subseteq> s"
unfolding homotopic_loops_def path_image_def
by (meson continuous_map_subtopology_eu homotopic_with_imp_continuous_maps)
proposition homotopic_loops_refl:
"homotopic_loops s p p \<longleftrightarrow>
path p \<and> path_image p \<subseteq> s \<and> pathfinish p = pathstart p"
by (simp add: homotopic_loops_def path_image_def path_def)
proposition homotopic_loops_sym: "homotopic_loops s p q \<Longrightarrow> homotopic_loops s q p"
by (simp add: homotopic_loops_def homotopic_with_sym)
proposition homotopic_loops_sym_eq: "homotopic_loops s p q \<longleftrightarrow> homotopic_loops s q p"
by (metis homotopic_loops_sym)
proposition homotopic_loops_trans:
"\<lbrakk>homotopic_loops s p q; homotopic_loops s q r\<rbrakk> \<Longrightarrow> homotopic_loops s p r"
unfolding homotopic_loops_def by (blast intro: homotopic_with_trans)
proposition homotopic_loops_subset:
"\<lbrakk>homotopic_loops s p q; s \<subseteq> t\<rbrakk> \<Longrightarrow> homotopic_loops t p q"
by (fastforce simp add: homotopic_loops)
proposition homotopic_loops_eq:
"\<lbrakk>path p; path_image p \<subseteq> s; pathfinish p = pathstart p; \<And>t. t \<in> {0..1} \<Longrightarrow> p(t) = q(t)\<rbrakk>
\<Longrightarrow> homotopic_loops s p q"
unfolding homotopic_loops_def
apply (rule homotopic_with_eq)
apply (rule homotopic_with_refl [where f = p, THEN iffD2])
apply (simp_all add: path_image_def path_def pathstart_def pathfinish_def)
done
proposition homotopic_loops_continuous_image:
"\<lbrakk>homotopic_loops s f g; continuous_on s h; h ` s \<subseteq> t\<rbrakk> \<Longrightarrow> homotopic_loops t (h \<circ> f) (h \<circ> g)"
unfolding homotopic_loops_def
by (simp add: homotopic_with_compose_continuous_map_left pathfinish_def pathstart_def)
subsection\<open>Relations between the two variants of homotopy\<close>
proposition homotopic_paths_imp_homotopic_loops:
"\<lbrakk>homotopic_paths s p q; pathfinish p = pathstart p; pathfinish q = pathstart p\<rbrakk> \<Longrightarrow> homotopic_loops s p q"
by (auto simp: homotopic_with_def homotopic_paths_def homotopic_loops_def)
proposition homotopic_loops_imp_homotopic_paths_null:
assumes "homotopic_loops s p (linepath a a)"
shows "homotopic_paths s p (linepath (pathstart p) (pathstart p))"
proof -
have "path p" by (metis assms homotopic_loops_imp_path)
have ploop: "pathfinish p = pathstart p" by (metis assms homotopic_loops_imp_loop)
have pip: "path_image p \<subseteq> s" by (metis assms homotopic_loops_imp_subset)
obtain h where conth: "continuous_on ({0..1::real} \<times> {0..1}) h"
and hs: "h ` ({0..1} \<times> {0..1}) \<subseteq> s"
and [simp]: "\<And>x. x \<in> {0..1} \<Longrightarrow> h(0,x) = p x"
and [simp]: "\<And>x. x \<in> {0..1} \<Longrightarrow> h(1,x) = a"
and ends: "\<And>t. t \<in> {0..1} \<Longrightarrow> pathfinish (h \<circ> Pair t) = pathstart (h \<circ> Pair t)"
using assms by (auto simp: homotopic_loops homotopic_with)
have conth0: "path (\<lambda>u. h (u, 0))"
unfolding path_def
apply (rule continuous_on_compose [of _ _ h, unfolded o_def])
apply (force intro: continuous_intros continuous_on_subset [OF conth])+
done
have pih0: "path_image (\<lambda>u. h (u, 0)) \<subseteq> s"
using hs by (force simp: path_image_def)
have c1: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>x. h (fst x * snd x, 0))"
apply (rule continuous_on_compose [of _ _ h, unfolded o_def])
apply (force simp: mult_le_one intro: continuous_intros continuous_on_subset [OF conth])+
done
have c2: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>x. h (fst x - fst x * snd x, 0))"
apply (rule continuous_on_compose [of _ _ h, unfolded o_def])
apply (force simp: mult_left_le mult_le_one intro: continuous_intros continuous_on_subset [OF conth])+
apply (rule continuous_on_subset [OF conth])
apply (auto simp: algebra_simps add_increasing2 mult_left_le)
done
have [simp]: "\<And>t. \<lbrakk>0 \<le> t \<and> t \<le> 1\<rbrakk> \<Longrightarrow> h (t, 1) = h (t, 0)"
using ends by (simp add: pathfinish_def pathstart_def)
have adhoc_le: "c * 4 \<le> 1 + c * (d * 4)" if "\<not> d * 4 \<le> 3" "0 \<le> c" "c \<le> 1" for c d::real
proof -
have "c * 3 \<le> c * (d * 4)" using that less_eq_real_def by auto
with \<open>c \<le> 1\<close> show ?thesis by fastforce
qed
have *: "\<And>p x. (path p \<and> path(reversepath p)) \<and>
(path_image p \<subseteq> s \<and> path_image(reversepath p) \<subseteq> s) \<and>
(pathfinish p = pathstart(linepath a a +++ reversepath p) \<and>
pathstart(reversepath p) = a) \<and> pathstart p = x
\<Longrightarrow> homotopic_paths s (p +++ linepath a a +++ reversepath p) (linepath x x)"
by (metis homotopic_paths_lid homotopic_paths_join
homotopic_paths_trans homotopic_paths_sym homotopic_paths_rinv)
have 1: "homotopic_paths s p (p +++ linepath (pathfinish p) (pathfinish p))"
using \<open>path p\<close> homotopic_paths_rid homotopic_paths_sym pip by blast
moreover have "homotopic_paths s (p +++ linepath (pathfinish p) (pathfinish p))
(linepath (pathstart p) (pathstart p) +++ p +++ linepath (pathfinish p) (pathfinish p))"
apply (rule homotopic_paths_sym)
using homotopic_paths_lid [of "p +++ linepath (pathfinish p) (pathfinish p)" s]
by (metis 1 homotopic_paths_imp_path homotopic_paths_imp_pathstart homotopic_paths_imp_subset)
moreover have "homotopic_paths s (linepath (pathstart p) (pathstart p) +++ p +++ linepath (pathfinish p) (pathfinish p))
((\<lambda>u. h (u, 0)) +++ linepath a a +++ reversepath (\<lambda>u. h (u, 0)))"
apply (simp add: homotopic_paths_def homotopic_with_def)
apply (rule_tac x="\<lambda>y. (subpath 0 (fst y) (\<lambda>u. h (u, 0)) +++ (\<lambda>u. h (Pair (fst y) u)) +++ subpath (fst y) 0 (\<lambda>u. h (u, 0))) (snd y)" in exI)
apply (simp add: subpath_reversepath)
apply (intro conjI homotopic_join_lemma)
using ploop
apply (simp_all add: path_defs joinpaths_def o_def subpath_def conth c1 c2)
apply (force simp: algebra_simps mult_le_one mult_left_le intro: hs [THEN subsetD] adhoc_le)
done
moreover have "homotopic_paths s ((\<lambda>u. h (u, 0)) +++ linepath a a +++ reversepath (\<lambda>u. h (u, 0)))
(linepath (pathstart p) (pathstart p))"
apply (rule *)
apply (simp add: pih0 pathstart_def pathfinish_def conth0)
apply (simp add: reversepath_def joinpaths_def)
done
ultimately show ?thesis
by (blast intro: homotopic_paths_trans)
qed
proposition homotopic_loops_conjugate:
fixes s :: "'a::real_normed_vector set"
assumes "path p" "path q" and pip: "path_image p \<subseteq> s" and piq: "path_image q \<subseteq> s"
and papp: "pathfinish p = pathstart q" and qloop: "pathfinish q = pathstart q"
shows "homotopic_loops s (p +++ q +++ reversepath p) q"
proof -
have contp: "continuous_on {0..1} p" using \<open>path p\<close> [unfolded path_def] by blast
have contq: "continuous_on {0..1} q" using \<open>path q\<close> [unfolded path_def] by blast
have c1: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>x. p ((1 - fst x) * snd x + fst x))"
apply (rule continuous_on_compose [of _ _ p, unfolded o_def])
apply (force simp: mult_le_one intro!: continuous_intros)
apply (rule continuous_on_subset [OF contp])
apply (auto simp: algebra_simps add_increasing2 mult_right_le_one_le sum_le_prod1)
done
have c2: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>x. p ((fst x - 1) * snd x + 1))"
apply (rule continuous_on_compose [of _ _ p, unfolded o_def])
apply (force simp: mult_le_one intro!: continuous_intros)
apply (rule continuous_on_subset [OF contp])
apply (auto simp: algebra_simps add_increasing2 mult_left_le_one_le)
done
have ps1: "\<And>a b. \<lbrakk>b * 2 \<le> 1; 0 \<le> b; 0 \<le> a; a \<le> 1\<rbrakk> \<Longrightarrow> p ((1 - a) * (2 * b) + a) \<in> s"
using sum_le_prod1
by (force simp: algebra_simps add_increasing2 mult_left_le intro: pip [unfolded path_image_def, THEN subsetD])
have ps2: "\<And>a b. \<lbrakk>\<not> 4 * b \<le> 3; b \<le> 1; 0 \<le> a; a \<le> 1\<rbrakk> \<Longrightarrow> p ((a - 1) * (4 * b - 3) + 1) \<in> s"
apply (rule pip [unfolded path_image_def, THEN subsetD])
apply (rule image_eqI, blast)
apply (simp add: algebra_simps)
- by (metis add_mono_thms_linordered_semiring(1) affine_ineq linear mult.commute mult.left_neutral mult_right_mono not_le
+ by (metis add_mono_thms_linordered_semiring(1) affine_ineq linear mult.commute mult.left_neutral mult_right_mono
add.commute zero_le_numeral)
have qs: "\<And>a b. \<lbrakk>4 * b \<le> 3; \<not> b * 2 \<le> 1\<rbrakk> \<Longrightarrow> q (4 * b - 2) \<in> s"
using path_image_def piq by fastforce
have "homotopic_loops s (p +++ q +++ reversepath p)
(linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q))"
apply (simp add: homotopic_loops_def homotopic_with_def)
apply (rule_tac x="(\<lambda>y. (subpath (fst y) 1 p +++ q +++ subpath 1 (fst y) p) (snd y))" in exI)
apply (simp add: subpath_refl subpath_reversepath)
apply (intro conjI homotopic_join_lemma)
using papp qloop
apply (simp_all add: path_defs joinpaths_def o_def subpath_def c1 c2)
apply (force simp: contq intro: continuous_on_compose [of _ _ q, unfolded o_def] continuous_on_id continuous_on_snd)
apply (auto simp: ps1 ps2 qs)
done
moreover have "homotopic_loops s (linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q)) q"
proof -
have "homotopic_paths s (linepath (pathfinish q) (pathfinish q) +++ q) q"
using \<open>path q\<close> homotopic_paths_lid qloop piq by auto
hence 1: "\<And>f. homotopic_paths s f q \<or> \<not> homotopic_paths s f (linepath (pathfinish q) (pathfinish q) +++ q)"
using homotopic_paths_trans by blast
hence "homotopic_paths s (linepath (pathfinish q) (pathfinish q) +++ q +++ linepath (pathfinish q) (pathfinish q)) q"
proof -
have "homotopic_paths s (q +++ linepath (pathfinish q) (pathfinish q)) q"
by (simp add: \<open>path q\<close> homotopic_paths_rid piq)
thus ?thesis
by (metis (no_types) 1 \<open>path q\<close> homotopic_paths_join homotopic_paths_rinv homotopic_paths_sym
homotopic_paths_trans qloop pathfinish_linepath piq)
qed
thus ?thesis
by (metis (no_types) qloop homotopic_loops_sym homotopic_paths_imp_homotopic_loops homotopic_paths_imp_pathfinish homotopic_paths_sym)
qed
ultimately show ?thesis
by (blast intro: homotopic_loops_trans)
qed
lemma homotopic_paths_loop_parts:
assumes loops: "homotopic_loops S (p +++ reversepath q) (linepath a a)" and "path q"
shows "homotopic_paths S p q"
proof -
have paths: "homotopic_paths S (p +++ reversepath q) (linepath (pathstart p) (pathstart p))"
using homotopic_loops_imp_homotopic_paths_null [OF loops] by simp
then have "path p"
using \<open>path q\<close> homotopic_loops_imp_path loops path_join path_join_path_ends path_reversepath by blast
show ?thesis
proof (cases "pathfinish p = pathfinish q")
case True
have pipq: "path_image p \<subseteq> S" "path_image q \<subseteq> S"
by (metis Un_subset_iff paths \<open>path p\<close> \<open>path q\<close> homotopic_loops_imp_subset homotopic_paths_imp_path loops
path_image_join path_image_reversepath path_imp_reversepath path_join_eq)+
have "homotopic_paths S p (p +++ (linepath (pathfinish p) (pathfinish p)))"
using \<open>path p\<close> \<open>path_image p \<subseteq> S\<close> homotopic_paths_rid homotopic_paths_sym by blast
moreover have "homotopic_paths S (p +++ (linepath (pathfinish p) (pathfinish p))) (p +++ (reversepath q +++ q))"
by (simp add: True \<open>path p\<close> \<open>path q\<close> pipq homotopic_paths_join homotopic_paths_linv homotopic_paths_sym)
moreover have "homotopic_paths S (p +++ (reversepath q +++ q)) ((p +++ reversepath q) +++ q)"
by (simp add: True \<open>path p\<close> \<open>path q\<close> homotopic_paths_assoc pipq)
moreover have "homotopic_paths S ((p +++ reversepath q) +++ q) (linepath (pathstart p) (pathstart p) +++ q)"
by (simp add: \<open>path q\<close> homotopic_paths_join paths pipq)
moreover then have "homotopic_paths S (linepath (pathstart p) (pathstart p) +++ q) q"
by (metis \<open>path q\<close> homotopic_paths_imp_path homotopic_paths_lid linepath_trivial path_join_path_ends pathfinish_def pipq(2))
ultimately show ?thesis
using homotopic_paths_trans by metis
next
case False
then show ?thesis
using \<open>path q\<close> homotopic_loops_imp_path loops path_join_path_ends by fastforce
qed
qed
subsection\<^marker>\<open>tag unimportant\<close>\<open>Homotopy of "nearby" function, paths and loops\<close>
lemma homotopic_with_linear:
fixes f g :: "_ \<Rightarrow> 'b::real_normed_vector"
assumes contf: "continuous_on s f"
and contg:"continuous_on s g"
and sub: "\<And>x. x \<in> s \<Longrightarrow> closed_segment (f x) (g x) \<subseteq> t"
shows "homotopic_with_canon (\<lambda>z. True) s t f g"
apply (simp add: homotopic_with_def)
apply (rule_tac x="\<lambda>y. ((1 - (fst y)) *\<^sub>R f(snd y) + (fst y) *\<^sub>R g(snd y))" in exI)
apply (intro conjI)
apply (rule subset_refl continuous_intros continuous_on_subset [OF contf] continuous_on_compose2 [where g=f]
continuous_on_subset [OF contg] continuous_on_compose2 [where g=g]| simp)+
using sub closed_segment_def apply fastforce+
done
lemma homotopic_paths_linear:
fixes g h :: "real \<Rightarrow> 'a::real_normed_vector"
assumes "path g" "path h" "pathstart h = pathstart g" "pathfinish h = pathfinish g"
"\<And>t. t \<in> {0..1} \<Longrightarrow> closed_segment (g t) (h t) \<subseteq> s"
shows "homotopic_paths s g h"
using assms
unfolding path_def
apply (simp add: closed_segment_def pathstart_def pathfinish_def homotopic_paths_def homotopic_with_def)
apply (rule_tac x="\<lambda>y. ((1 - (fst y)) *\<^sub>R (g \<circ> snd) y + (fst y) *\<^sub>R (h \<circ> snd) y)" in exI)
apply (intro conjI subsetI continuous_intros; force)
done
lemma homotopic_loops_linear:
fixes g h :: "real \<Rightarrow> 'a::real_normed_vector"
assumes "path g" "path h" "pathfinish g = pathstart g" "pathfinish h = pathstart h"
"\<And>t x. t \<in> {0..1} \<Longrightarrow> closed_segment (g t) (h t) \<subseteq> s"
shows "homotopic_loops s g h"
using assms
unfolding path_def
apply (simp add: pathstart_def pathfinish_def homotopic_loops_def homotopic_with_def)
apply (rule_tac x="\<lambda>y. ((1 - (fst y)) *\<^sub>R g(snd y) + (fst y) *\<^sub>R h(snd y))" in exI)
apply (auto intro!: continuous_intros intro: continuous_on_compose2 [where g=g] continuous_on_compose2 [where g=h])
apply (force simp: closed_segment_def)
done
lemma homotopic_paths_nearby_explicit:
assumes "path g" "path h" "pathstart h = pathstart g" "pathfinish h = pathfinish g"
and no: "\<And>t x. \<lbrakk>t \<in> {0..1}; x \<notin> s\<rbrakk> \<Longrightarrow> norm(h t - g t) < norm(g t - x)"
shows "homotopic_paths s g h"
apply (rule homotopic_paths_linear [OF assms(1-4)])
by (metis no segment_bound(1) subsetI norm_minus_commute not_le)
lemma homotopic_loops_nearby_explicit:
assumes "path g" "path h" "pathfinish g = pathstart g" "pathfinish h = pathstart h"
and no: "\<And>t x. \<lbrakk>t \<in> {0..1}; x \<notin> s\<rbrakk> \<Longrightarrow> norm(h t - g t) < norm(g t - x)"
shows "homotopic_loops s g h"
apply (rule homotopic_loops_linear [OF assms(1-4)])
by (metis no segment_bound(1) subsetI norm_minus_commute not_le)
lemma homotopic_nearby_paths:
fixes g h :: "real \<Rightarrow> 'a::euclidean_space"
assumes "path g" "open s" "path_image g \<subseteq> s"
shows "\<exists>e. 0 < e \<and>
(\<forall>h. path h \<and>
pathstart h = pathstart g \<and> pathfinish h = pathfinish g \<and>
(\<forall>t \<in> {0..1}. norm(h t - g t) < e) \<longrightarrow> homotopic_paths s g h)"
proof -
obtain e where "e > 0" and e: "\<And>x y. x \<in> path_image g \<Longrightarrow> y \<in> - s \<Longrightarrow> e \<le> dist x y"
using separate_compact_closed [of "path_image g" "-s"] assms by force
show ?thesis
apply (intro exI conjI)
using e [unfolded dist_norm]
apply (auto simp: intro!: homotopic_paths_nearby_explicit assms \<open>e > 0\<close>)
by (metis atLeastAtMost_iff imageI le_less_trans not_le path_image_def)
qed
lemma homotopic_nearby_loops:
fixes g h :: "real \<Rightarrow> 'a::euclidean_space"
assumes "path g" "open s" "path_image g \<subseteq> s" "pathfinish g = pathstart g"
shows "\<exists>e. 0 < e \<and>
(\<forall>h. path h \<and> pathfinish h = pathstart h \<and>
(\<forall>t \<in> {0..1}. norm(h t - g t) < e) \<longrightarrow> homotopic_loops s g h)"
proof -
obtain e where "e > 0" and e: "\<And>x y. x \<in> path_image g \<Longrightarrow> y \<in> - s \<Longrightarrow> e \<le> dist x y"
using separate_compact_closed [of "path_image g" "-s"] assms by force
show ?thesis
apply (intro exI conjI)
using e [unfolded dist_norm]
apply (auto simp: intro!: homotopic_loops_nearby_explicit assms \<open>e > 0\<close>)
by (metis atLeastAtMost_iff imageI le_less_trans not_le path_image_def)
qed
subsection\<open> Homotopy and subpaths\<close>
lemma homotopic_join_subpaths1:
assumes "path g" and pag: "path_image g \<subseteq> s"
and u: "u \<in> {0..1}" and v: "v \<in> {0..1}" and w: "w \<in> {0..1}" "u \<le> v" "v \<le> w"
shows "homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)"
proof -
have 1: "t * 2 \<le> 1 \<Longrightarrow> u + t * (v * 2) \<le> v + t * (u * 2)" for t
using affine_ineq \<open>u \<le> v\<close> by fastforce
have 2: "t * 2 > 1 \<Longrightarrow> u + (2*t - 1) * v \<le> v + (2*t - 1) * w" for t
by (metis add_mono_thms_linordered_semiring(1) diff_gt_0_iff_gt less_eq_real_def mult.commute mult_right_mono \<open>u \<le> v\<close> \<open>v \<le> w\<close>)
have t2: "\<And>t::real. t*2 = 1 \<Longrightarrow> t = 1/2" by auto
show ?thesis
apply (rule homotopic_paths_subset [OF _ pag])
using assms
apply (cases "w = u")
using homotopic_paths_rinv [of "subpath u v g" "path_image g"]
apply (force simp: closed_segment_eq_real_ivl image_mono path_image_def subpath_refl)
apply (rule homotopic_paths_sym)
apply (rule homotopic_paths_reparametrize
[where f = "\<lambda>t. if t \<le> 1 / 2
then inverse((w - u)) *\<^sub>R (2 * (v - u)) *\<^sub>R t
else inverse((w - u)) *\<^sub>R ((v - u) + (w - v) *\<^sub>R (2 *\<^sub>R t - 1))"])
using \<open>path g\<close> path_subpath u w apply blast
using \<open>path g\<close> path_image_subpath_subset u w(1) apply blast
apply simp_all
apply (subst split_01)
apply (rule continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def)+
apply (simp_all add: field_simps not_le)
apply (force dest!: t2)
apply (force simp: algebra_simps mult_left_mono affine_ineq dest!: 1 2)
apply (simp add: joinpaths_def subpath_def)
apply (force simp: algebra_simps)
done
qed
lemma homotopic_join_subpaths2:
assumes "homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)"
shows "homotopic_paths s (subpath w v g +++ subpath v u g) (subpath w u g)"
by (metis assms homotopic_paths_reversepath_D pathfinish_subpath pathstart_subpath reversepath_joinpaths reversepath_subpath)
lemma homotopic_join_subpaths3:
assumes hom: "homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)"
and "path g" and pag: "path_image g \<subseteq> s"
and u: "u \<in> {0..1}" and v: "v \<in> {0..1}" and w: "w \<in> {0..1}"
shows "homotopic_paths s (subpath v w g +++ subpath w u g) (subpath v u g)"
proof -
have "homotopic_paths s (subpath u w g +++ subpath w v g) ((subpath u v g +++ subpath v w g) +++ subpath w v g)"
apply (rule homotopic_paths_join)
using hom homotopic_paths_sym_eq apply blast
apply (metis \<open>path g\<close> homotopic_paths_eq pag path_image_subpath_subset path_subpath subset_trans v w, simp)
done
also have "homotopic_paths s ((subpath u v g +++ subpath v w g) +++ subpath w v g) (subpath u v g +++ subpath v w g +++ subpath w v g)"
apply (rule homotopic_paths_sym [OF homotopic_paths_assoc])
using assms by (simp_all add: path_image_subpath_subset [THEN order_trans])
also have "homotopic_paths s (subpath u v g +++ subpath v w g +++ subpath w v g)
(subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g)))"
apply (rule homotopic_paths_join)
apply (metis \<open>path g\<close> homotopic_paths_eq order.trans pag path_image_subpath_subset path_subpath u v)
apply (metis (no_types, lifting) \<open>path g\<close> homotopic_paths_linv order_trans pag path_image_subpath_subset path_subpath pathfinish_subpath reversepath_subpath v w)
apply simp
done
also have "homotopic_paths s (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g))) (subpath u v g)"
apply (rule homotopic_paths_rid)
using \<open>path g\<close> path_subpath u v apply blast
apply (meson \<open>path g\<close> order.trans pag path_image_subpath_subset u v)
done
finally have "homotopic_paths s (subpath u w g +++ subpath w v g) (subpath u v g)" .
then show ?thesis
using homotopic_join_subpaths2 by blast
qed
proposition homotopic_join_subpaths:
"\<lbrakk>path g; path_image g \<subseteq> s; u \<in> {0..1}; v \<in> {0..1}; w \<in> {0..1}\<rbrakk>
\<Longrightarrow> homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)"
apply (rule le_cases3 [of u v w])
using homotopic_join_subpaths1 homotopic_join_subpaths2 homotopic_join_subpaths3 by metis+
text\<open>Relating homotopy of trivial loops to path-connectedness.\<close>
lemma path_component_imp_homotopic_points:
"path_component S a b \<Longrightarrow> homotopic_loops S (linepath a a) (linepath b b)"
apply (simp add: path_component_def homotopic_loops_def homotopic_with_def
pathstart_def pathfinish_def path_image_def path_def, clarify)
apply (rule_tac x="g \<circ> fst" in exI)
apply (intro conjI continuous_intros continuous_on_compose)+
apply (auto elim!: continuous_on_subset)
done
lemma homotopic_loops_imp_path_component_value:
"\<lbrakk>homotopic_loops S p q; 0 \<le> t; t \<le> 1\<rbrakk>
\<Longrightarrow> path_component S (p t) (q t)"
apply (simp add: path_component_def homotopic_loops_def homotopic_with_def
pathstart_def pathfinish_def path_image_def path_def, clarify)
apply (rule_tac x="h \<circ> (\<lambda>u. (u, t))" in exI)
apply (intro conjI continuous_intros continuous_on_compose)+
apply (auto elim!: continuous_on_subset)
done
lemma homotopic_points_eq_path_component:
"homotopic_loops S (linepath a a) (linepath b b) \<longleftrightarrow>
path_component S a b"
by (auto simp: path_component_imp_homotopic_points
dest: homotopic_loops_imp_path_component_value [where t=1])
lemma path_connected_eq_homotopic_points:
"path_connected S \<longleftrightarrow>
(\<forall>a b. a \<in> S \<and> b \<in> S \<longrightarrow> homotopic_loops S (linepath a a) (linepath b b))"
by (auto simp: path_connected_def path_component_def homotopic_points_eq_path_component)
subsection\<open>Simply connected sets\<close>
text\<^marker>\<open>tag important\<close>\<open>defined as "all loops are homotopic (as loops)\<close>
definition\<^marker>\<open>tag important\<close> simply_connected where
"simply_connected S \<equiv>
\<forall>p q. path p \<and> pathfinish p = pathstart p \<and> path_image p \<subseteq> S \<and>
path q \<and> pathfinish q = pathstart q \<and> path_image q \<subseteq> S
\<longrightarrow> homotopic_loops S p q"
lemma simply_connected_empty [iff]: "simply_connected {}"
by (simp add: simply_connected_def)
lemma simply_connected_imp_path_connected:
fixes S :: "_::real_normed_vector set"
shows "simply_connected S \<Longrightarrow> path_connected S"
by (simp add: simply_connected_def path_connected_eq_homotopic_points)
lemma simply_connected_imp_connected:
fixes S :: "_::real_normed_vector set"
shows "simply_connected S \<Longrightarrow> connected S"
by (simp add: path_connected_imp_connected simply_connected_imp_path_connected)
lemma simply_connected_eq_contractible_loop_any:
fixes S :: "_::real_normed_vector set"
shows "simply_connected S \<longleftrightarrow>
(\<forall>p a. path p \<and> path_image p \<subseteq> S \<and>
pathfinish p = pathstart p \<and> a \<in> S
\<longrightarrow> homotopic_loops S p (linepath a a))"
apply (simp add: simply_connected_def)
apply (rule iffI, force, clarify)
apply (rule_tac q = "linepath (pathstart p) (pathstart p)" in homotopic_loops_trans)
apply (fastforce simp add:)
using homotopic_loops_sym apply blast
done
lemma simply_connected_eq_contractible_loop_some:
fixes S :: "_::real_normed_vector set"
shows "simply_connected S \<longleftrightarrow>
path_connected S \<and>
(\<forall>p. path p \<and> path_image p \<subseteq> S \<and> pathfinish p = pathstart p
\<longrightarrow> (\<exists>a. a \<in> S \<and> homotopic_loops S p (linepath a a)))"
apply (rule iffI)
apply (fastforce simp: simply_connected_imp_path_connected simply_connected_eq_contractible_loop_any)
apply (clarsimp simp add: simply_connected_eq_contractible_loop_any)
apply (drule_tac x=p in spec)
using homotopic_loops_trans path_connected_eq_homotopic_points
apply blast
done
lemma simply_connected_eq_contractible_loop_all:
fixes S :: "_::real_normed_vector set"
shows "simply_connected S \<longleftrightarrow>
S = {} \<or>
(\<exists>a \<in> S. \<forall>p. path p \<and> path_image p \<subseteq> S \<and> pathfinish p = pathstart p
\<longrightarrow> homotopic_loops S p (linepath a a))"
(is "?lhs = ?rhs")
proof (cases "S = {}")
case True then show ?thesis by force
next
case False
then obtain a where "a \<in> S" by blast
show ?thesis
proof
assume "simply_connected S"
then show ?rhs
using \<open>a \<in> S\<close> \<open>simply_connected S\<close> simply_connected_eq_contractible_loop_any
by blast
next
assume ?rhs
then show "simply_connected S"
apply (simp add: simply_connected_eq_contractible_loop_any False)
by (meson homotopic_loops_refl homotopic_loops_sym homotopic_loops_trans
path_component_imp_homotopic_points path_component_refl)
qed
qed
lemma simply_connected_eq_contractible_path:
fixes S :: "_::real_normed_vector set"
shows "simply_connected S \<longleftrightarrow>
path_connected S \<and>
(\<forall>p. path p \<and> path_image p \<subseteq> S \<and> pathfinish p = pathstart p
\<longrightarrow> homotopic_paths S p (linepath (pathstart p) (pathstart p)))"
apply (rule iffI)
apply (simp add: simply_connected_imp_path_connected)
apply (metis simply_connected_eq_contractible_loop_some homotopic_loops_imp_homotopic_paths_null)
by (meson homotopic_paths_imp_homotopic_loops pathfinish_linepath pathstart_in_path_image
simply_connected_eq_contractible_loop_some subset_iff)
lemma simply_connected_eq_homotopic_paths:
fixes S :: "_::real_normed_vector set"
shows "simply_connected S \<longleftrightarrow>
path_connected S \<and>
(\<forall>p q. path p \<and> path_image p \<subseteq> S \<and>
path q \<and> path_image q \<subseteq> S \<and>
pathstart q = pathstart p \<and> pathfinish q = pathfinish p
\<longrightarrow> homotopic_paths S p q)"
(is "?lhs = ?rhs")
proof
assume ?lhs
then have pc: "path_connected S"
and *: "\<And>p. \<lbrakk>path p; path_image p \<subseteq> S;
pathfinish p = pathstart p\<rbrakk>
\<Longrightarrow> homotopic_paths S p (linepath (pathstart p) (pathstart p))"
by (auto simp: simply_connected_eq_contractible_path)
have "homotopic_paths S p q"
if "path p" "path_image p \<subseteq> S" "path q"
"path_image q \<subseteq> S" "pathstart q = pathstart p"
"pathfinish q = pathfinish p" for p q
proof -
have "homotopic_paths S p (p +++ linepath (pathfinish p) (pathfinish p))"
by (simp add: homotopic_paths_rid homotopic_paths_sym that)
also have "homotopic_paths S (p +++ linepath (pathfinish p) (pathfinish p))
(p +++ reversepath q +++ q)"
using that
by (metis homotopic_paths_join homotopic_paths_linv homotopic_paths_refl homotopic_paths_sym_eq pathstart_linepath)
also have "homotopic_paths S (p +++ reversepath q +++ q)
((p +++ reversepath q) +++ q)"
by (simp add: that homotopic_paths_assoc)
also have "homotopic_paths S ((p +++ reversepath q) +++ q)
(linepath (pathstart q) (pathstart q) +++ q)"
using * [of "p +++ reversepath q"] that
by (simp add: homotopic_paths_join path_image_join)
also have "homotopic_paths S (linepath (pathstart q) (pathstart q) +++ q) q"
using that homotopic_paths_lid by blast
finally show ?thesis .
qed
then show ?rhs
by (blast intro: pc *)
next
assume ?rhs
then show ?lhs
by (force simp: simply_connected_eq_contractible_path)
qed
proposition simply_connected_Times:
fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set"
assumes S: "simply_connected S" and T: "simply_connected T"
shows "simply_connected(S \<times> T)"
proof -
have "homotopic_loops (S \<times> T) p (linepath (a, b) (a, b))"
if "path p" "path_image p \<subseteq> S \<times> T" "p 1 = p 0" "a \<in> S" "b \<in> T"
for p a b
proof -
have "path (fst \<circ> p)"
apply (rule Path_Connected.path_continuous_image [OF \<open>path p\<close>])
apply (rule continuous_intros)+
done
moreover have "path_image (fst \<circ> p) \<subseteq> S"
using that apply (simp add: path_image_def) by force
ultimately have p1: "homotopic_loops S (fst \<circ> p) (linepath a a)"
using S that
apply (simp add: simply_connected_eq_contractible_loop_any)
apply (drule_tac x="fst \<circ> p" in spec)
apply (drule_tac x=a in spec)
apply (auto simp: pathstart_def pathfinish_def)
done
have "path (snd \<circ> p)"
apply (rule Path_Connected.path_continuous_image [OF \<open>path p\<close>])
apply (rule continuous_intros)+
done
moreover have "path_image (snd \<circ> p) \<subseteq> T"
using that apply (simp add: path_image_def) by force
ultimately have p2: "homotopic_loops T (snd \<circ> p) (linepath b b)"
using T that
apply (simp add: simply_connected_eq_contractible_loop_any)
apply (drule_tac x="snd \<circ> p" in spec)
apply (drule_tac x=b in spec)
apply (auto simp: pathstart_def pathfinish_def)
done
show ?thesis
using p1 p2
apply (simp add: homotopic_loops, clarify)
apply (rename_tac h k)
apply (rule_tac x="\<lambda>z. Pair (h z) (k z)" in exI)
apply (intro conjI continuous_intros | assumption)+
apply (auto simp: pathstart_def pathfinish_def)
done
qed
with assms show ?thesis
by (simp add: simply_connected_eq_contractible_loop_any pathfinish_def pathstart_def)
qed
subsection\<open>Contractible sets\<close>
definition\<^marker>\<open>tag important\<close> contractible where
"contractible S \<equiv> \<exists>a. homotopic_with_canon (\<lambda>x. True) S S id (\<lambda>x. a)"
proposition contractible_imp_simply_connected:
fixes S :: "_::real_normed_vector set"
assumes "contractible S" shows "simply_connected S"
proof (cases "S = {}")
case True then show ?thesis by force
next
case False
obtain a where a: "homotopic_with_canon (\<lambda>x. True) S S id (\<lambda>x. a)"
using assms by (force simp: contractible_def)
then have "a \<in> S"
by (metis False homotopic_constant_maps homotopic_with_symD homotopic_with_trans path_component_in_topspace topspace_euclidean_subtopology)
show ?thesis
apply (simp add: simply_connected_eq_contractible_loop_all False)
apply (rule bexI [OF _ \<open>a \<in> S\<close>])
using a apply (simp add: homotopic_loops_def homotopic_with_def path_def path_image_def pathfinish_def pathstart_def, clarify)
apply (rule_tac x="(h \<circ> (\<lambda>y. (fst y, (p \<circ> snd) y)))" in exI)
apply (intro conjI continuous_on_compose continuous_intros)
apply (erule continuous_on_subset | force)+
done
qed
corollary contractible_imp_connected:
fixes S :: "_::real_normed_vector set"
shows "contractible S \<Longrightarrow> connected S"
by (simp add: contractible_imp_simply_connected simply_connected_imp_connected)
lemma contractible_imp_path_connected:
fixes S :: "_::real_normed_vector set"
shows "contractible S \<Longrightarrow> path_connected S"
by (simp add: contractible_imp_simply_connected simply_connected_imp_path_connected)
lemma nullhomotopic_through_contractible:
fixes S :: "_::topological_space set"
assumes f: "continuous_on S f" "f ` S \<subseteq> T"
and g: "continuous_on T g" "g ` T \<subseteq> U"
and T: "contractible T"
obtains c where "homotopic_with_canon (\<lambda>h. True) S U (g \<circ> f) (\<lambda>x. c)"
proof -
obtain b where b: "homotopic_with_canon (\<lambda>x. True) T T id (\<lambda>x. b)"
using assms by (force simp: contractible_def)
have "homotopic_with_canon (\<lambda>f. True) T U (g \<circ> id) (g \<circ> (\<lambda>x. b))"
by (metis Abstract_Topology.continuous_map_subtopology_eu b g homotopic_compose_continuous_map_left)
then have "homotopic_with_canon (\<lambda>f. True) S U (g \<circ> id \<circ> f) (g \<circ> (\<lambda>x. b) \<circ> f)"
by (simp add: f homotopic_with_compose_continuous_map_right)
then show ?thesis
by (simp add: comp_def that)
qed
lemma nullhomotopic_into_contractible:
assumes f: "continuous_on S f" "f ` S \<subseteq> T"
and T: "contractible T"
obtains c where "homotopic_with_canon (\<lambda>h. True) S T f (\<lambda>x. c)"
apply (rule nullhomotopic_through_contractible [OF f, of id T])
using assms
apply (auto)
done
lemma nullhomotopic_from_contractible:
assumes f: "continuous_on S f" "f ` S \<subseteq> T"
and S: "contractible S"
obtains c where "homotopic_with_canon (\<lambda>h. True) S T f (\<lambda>x. c)"
apply (rule nullhomotopic_through_contractible [OF continuous_on_id _ f S, of S])
using assms
apply (auto simp: comp_def)
done
lemma homotopic_through_contractible:
fixes S :: "_::real_normed_vector set"
assumes "continuous_on S f1" "f1 ` S \<subseteq> T"
"continuous_on T g1" "g1 ` T \<subseteq> U"
"continuous_on S f2" "f2 ` S \<subseteq> T"
"continuous_on T g2" "g2 ` T \<subseteq> U"
"contractible T" "path_connected U"
shows "homotopic_with_canon (\<lambda>h. True) S U (g1 \<circ> f1) (g2 \<circ> f2)"
proof -
obtain c1 where c1: "homotopic_with_canon (\<lambda>h. True) S U (g1 \<circ> f1) (\<lambda>x. c1)"
apply (rule nullhomotopic_through_contractible [of S f1 T g1 U])
using assms apply auto
done
obtain c2 where c2: "homotopic_with_canon (\<lambda>h. True) S U (g2 \<circ> f2) (\<lambda>x. c2)"
apply (rule nullhomotopic_through_contractible [of S f2 T g2 U])
using assms apply auto
done
have *: "S = {} \<or> (\<exists>t. path_connected t \<and> t \<subseteq> U \<and> c2 \<in> t \<and> c1 \<in> t)"
proof (cases "S = {}")
case True then show ?thesis by force
next
case False
with c1 c2 have "c1 \<in> U" "c2 \<in> U"
using homotopic_with_imp_continuous_maps by fastforce+
with \<open>path_connected U\<close> show ?thesis by blast
qed
show ?thesis
apply (rule homotopic_with_trans [OF c1])
apply (rule homotopic_with_symD)
apply (rule homotopic_with_trans [OF c2])
apply (simp add: path_component homotopic_constant_maps *)
done
qed
lemma homotopic_into_contractible:
fixes S :: "'a::real_normed_vector set" and T:: "'b::real_normed_vector set"
assumes f: "continuous_on S f" "f ` S \<subseteq> T"
and g: "continuous_on S g" "g ` S \<subseteq> T"
and T: "contractible T"
shows "homotopic_with_canon (\<lambda>h. True) S T f g"
using homotopic_through_contractible [of S f T id T g id]
by (simp add: assms contractible_imp_path_connected)
lemma homotopic_from_contractible:
fixes S :: "'a::real_normed_vector set" and T:: "'b::real_normed_vector set"
assumes f: "continuous_on S f" "f ` S \<subseteq> T"
and g: "continuous_on S g" "g ` S \<subseteq> T"
and "contractible S" "path_connected T"
shows "homotopic_with_canon (\<lambda>h. True) S T f g"
using homotopic_through_contractible [of S id S f T id g]
by (simp add: assms contractible_imp_path_connected)
subsection\<open>Starlike sets\<close>
definition\<^marker>\<open>tag important\<close> "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 convex_imp_starlike:
"convex S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> starlike S"
unfolding convex_contains_segment starlike_def by auto
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
lemma starlike_imp_contractible_gen:
fixes S :: "'a::real_normed_vector set"
assumes S: "starlike S"
and P: "\<And>a T. \<lbrakk>a \<in> S; 0 \<le> T; T \<le> 1\<rbrakk> \<Longrightarrow> P(\<lambda>x. (1 - T) *\<^sub>R x + T *\<^sub>R a)"
obtains a where "homotopic_with_canon P S S (\<lambda>x. x) (\<lambda>x. a)"
proof -
obtain a where "a \<in> S" and a: "\<And>x. x \<in> S \<Longrightarrow> closed_segment a x \<subseteq> S"
using S by (auto simp: starlike_def)
have "(\<lambda>y. (1 - fst y) *\<^sub>R snd y + fst y *\<^sub>R a) ` ({0..1} \<times> S) \<subseteq> S"
apply clarify
apply (erule a [unfolded closed_segment_def, THEN subsetD], simp)
apply (metis add_diff_cancel_right' diff_ge_0_iff_ge le_add_diff_inverse pth_c(1))
done
then show ?thesis
apply (rule_tac a=a in that)
using \<open>a \<in> S\<close>
apply (simp add: homotopic_with_def)
apply (rule_tac x="\<lambda>y. (1 - (fst y)) *\<^sub>R snd y + (fst y) *\<^sub>R a" in exI)
apply (intro conjI ballI continuous_on_compose continuous_intros)
apply (simp_all add: P)
done
qed
lemma starlike_imp_contractible:
fixes S :: "'a::real_normed_vector set"
shows "starlike S \<Longrightarrow> contractible S"
using starlike_imp_contractible_gen contractible_def by (fastforce simp: id_def)
lemma contractible_UNIV [simp]: "contractible (UNIV :: 'a::real_normed_vector set)"
by (simp add: starlike_imp_contractible)
lemma starlike_imp_simply_connected:
fixes S :: "'a::real_normed_vector set"
shows "starlike S \<Longrightarrow> simply_connected S"
by (simp add: contractible_imp_simply_connected starlike_imp_contractible)
lemma convex_imp_simply_connected:
fixes S :: "'a::real_normed_vector set"
shows "convex S \<Longrightarrow> simply_connected S"
using convex_imp_starlike starlike_imp_simply_connected by blast
lemma starlike_imp_path_connected:
fixes S :: "'a::real_normed_vector set"
shows "starlike S \<Longrightarrow> path_connected S"
by (simp add: simply_connected_imp_path_connected starlike_imp_simply_connected)
lemma starlike_imp_connected:
fixes S :: "'a::real_normed_vector set"
shows "starlike S \<Longrightarrow> connected S"
by (simp add: path_connected_imp_connected starlike_imp_path_connected)
lemma is_interval_simply_connected_1:
fixes S :: "real set"
shows "is_interval S \<longleftrightarrow> simply_connected S"
using convex_imp_simply_connected is_interval_convex_1 is_interval_path_connected_1 simply_connected_imp_path_connected by auto
lemma contractible_empty [simp]: "contractible {}"
by (simp add: contractible_def homotopic_on_emptyI)
lemma contractible_convex_tweak_boundary_points:
fixes S :: "'a::euclidean_space set"
assumes "convex S" and TS: "rel_interior S \<subseteq> T" "T \<subseteq> closure S"
shows "contractible T"
proof (cases "S = {}")
case True
with assms show ?thesis
by (simp add: subsetCE)
next
case False
show ?thesis
apply (rule starlike_imp_contractible)
apply (rule starlike_convex_tweak_boundary_points [OF \<open>convex S\<close> False TS])
done
qed
lemma convex_imp_contractible:
fixes S :: "'a::real_normed_vector set"
shows "convex S \<Longrightarrow> contractible S"
using contractible_empty convex_imp_starlike starlike_imp_contractible by blast
lemma contractible_sing [simp]:
fixes a :: "'a::real_normed_vector"
shows "contractible {a}"
by (rule convex_imp_contractible [OF convex_singleton])
lemma is_interval_contractible_1:
fixes S :: "real set"
shows "is_interval S \<longleftrightarrow> contractible S"
using contractible_imp_simply_connected convex_imp_contractible is_interval_convex_1
is_interval_simply_connected_1 by auto
lemma contractible_Times:
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
assumes S: "contractible S" and T: "contractible T"
shows "contractible (S \<times> T)"
proof -
obtain a h where conth: "continuous_on ({0..1} \<times> S) h"
and hsub: "h ` ({0..1} \<times> S) \<subseteq> S"
and [simp]: "\<And>x. x \<in> S \<Longrightarrow> h (0, x) = x"
and [simp]: "\<And>x. x \<in> S \<Longrightarrow> h (1::real, x) = a"
using S by (auto simp: contractible_def homotopic_with)
obtain b k where contk: "continuous_on ({0..1} \<times> T) k"
and ksub: "k ` ({0..1} \<times> T) \<subseteq> T"
and [simp]: "\<And>x. x \<in> T \<Longrightarrow> k (0, x) = x"
and [simp]: "\<And>x. x \<in> T \<Longrightarrow> k (1::real, x) = b"
using T by (auto simp: contractible_def homotopic_with)
show ?thesis
apply (simp add: contractible_def homotopic_with)
apply (rule exI [where x=a])
apply (rule exI [where x=b])
apply (rule exI [where x = "\<lambda>z. (h (fst z, fst(snd z)), k (fst z, snd(snd z)))"])
apply (intro conjI ballI continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF contk])
using hsub ksub
apply auto
done
qed
subsection\<open>Local versions of topological properties in general\<close>
definition\<^marker>\<open>tag important\<close> locally :: "('a::topological_space set \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
where
"locally P S \<equiv>
\<forall>w x. openin (top_of_set S) w \<and> x \<in> w
\<longrightarrow> (\<exists>u v. openin (top_of_set S) u \<and> P v \<and>
x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> w)"
lemma locallyI:
assumes "\<And>w x. \<lbrakk>openin (top_of_set S) w; x \<in> w\<rbrakk>
\<Longrightarrow> \<exists>u v. openin (top_of_set S) u \<and> P v \<and>
x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> w"
shows "locally P S"
using assms by (force simp: locally_def)
lemma locallyE:
assumes "locally P S" "openin (top_of_set S) w" "x \<in> w"
obtains u v where "openin (top_of_set S) u"
"P v" "x \<in> u" "u \<subseteq> v" "v \<subseteq> w"
using assms unfolding locally_def by meson
lemma locally_mono:
assumes "locally P S" "\<And>t. P t \<Longrightarrow> Q t"
shows "locally Q S"
by (metis assms locally_def)
lemma locally_open_subset:
assumes "locally P S" "openin (top_of_set S) t"
shows "locally P t"
using assms
apply (simp add: locally_def)
apply (erule all_forward)+
apply (rule impI)
apply (erule impCE)
using openin_trans apply blast
apply (erule ex_forward)
by (metis (no_types, hide_lams) Int_absorb1 Int_lower1 Int_subset_iff openin_open openin_subtopology_Int_subset)
lemma locally_diff_closed:
"\<lbrakk>locally P S; closedin (top_of_set S) t\<rbrakk> \<Longrightarrow> locally P (S - t)"
using locally_open_subset closedin_def by fastforce
lemma locally_empty [iff]: "locally P {}"
by (simp add: locally_def openin_subtopology)
lemma locally_singleton [iff]:
fixes a :: "'a::metric_space"
shows "locally P {a} \<longleftrightarrow> P {a}"
apply (simp add: locally_def openin_euclidean_subtopology_iff subset_singleton_iff conj_disj_distribR cong: conj_cong)
using zero_less_one by blast
lemma locally_iff:
"locally P S \<longleftrightarrow>
(\<forall>T x. open T \<and> x \<in> S \<inter> T \<longrightarrow> (\<exists>U. open U \<and> (\<exists>v. P v \<and> x \<in> S \<inter> U \<and> S \<inter> U \<subseteq> v \<and> v \<subseteq> S \<inter> T)))"
apply (simp add: le_inf_iff locally_def openin_open, safe)
apply (metis IntE IntI le_inf_iff)
apply (metis IntI Int_subset_iff)
done
lemma locally_Int:
assumes S: "locally P S" and t: "locally P t"
and P: "\<And>S t. P S \<and> P t \<Longrightarrow> P(S \<inter> t)"
shows "locally P (S \<inter> t)"
using S t unfolding locally_iff
apply clarify
apply (drule_tac x=T in spec)+
apply (drule_tac x=x in spec)+
apply clarsimp
apply (rename_tac U1 U2 V1 V2)
apply (rule_tac x="U1 \<inter> U2" in exI)
apply (simp add: open_Int)
apply (rule_tac x="V1 \<inter> V2" in exI)
apply (auto intro: P)
done
lemma locally_Times:
fixes S :: "('a::metric_space) set" and T :: "('b::metric_space) set"
assumes PS: "locally P S" and QT: "locally Q T" and R: "\<And>S T. P S \<and> Q T \<Longrightarrow> R(S \<times> T)"
shows "locally R (S \<times> T)"
unfolding locally_def
proof (clarify)
fix W x y
assume W: "openin (top_of_set (S \<times> T)) W" and xy: "(x, y) \<in> W"
then obtain U V where "openin (top_of_set S) U" "x \<in> U"
"openin (top_of_set T) V" "y \<in> V" "U \<times> V \<subseteq> W"
using Times_in_interior_subtopology by metis
then obtain U1 U2 V1 V2
where opeS: "openin (top_of_set S) U1 \<and> P U2 \<and> x \<in> U1 \<and> U1 \<subseteq> U2 \<and> U2 \<subseteq> U"
and opeT: "openin (top_of_set T) V1 \<and> Q V2 \<and> y \<in> V1 \<and> V1 \<subseteq> V2 \<and> V2 \<subseteq> V"
by (meson PS QT locallyE)
with \<open>U \<times> V \<subseteq> W\<close> show "\<exists>u v. openin (top_of_set (S \<times> T)) u \<and> R v \<and> (x,y) \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> W"
apply (rule_tac x="U1 \<times> V1" in exI)
apply (rule_tac x="U2 \<times> V2" in exI)
apply (auto simp: openin_Times R openin_Times_eq)
done
qed
proposition homeomorphism_locally_imp:
fixes S :: "'a::metric_space set" and t :: "'b::t2_space set"
assumes S: "locally P S" and hom: "homeomorphism S t f g"
and Q: "\<And>S S'. \<lbrakk>P S; homeomorphism S S' f g\<rbrakk> \<Longrightarrow> Q S'"
shows "locally Q t"
proof (clarsimp simp: locally_def)
fix W y
assume "y \<in> W" and "openin (top_of_set t) W"
then obtain T where T: "open T" "W = t \<inter> T"
by (force simp: openin_open)
then have "W \<subseteq> t" by auto
have f: "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x" "f ` S = t" "continuous_on S f"
and g: "\<And>y. y \<in> t \<Longrightarrow> f(g y) = y" "g ` t = S" "continuous_on t g"
using hom by (auto simp: homeomorphism_def)
have gw: "g ` W = S \<inter> f -` W"
using \<open>W \<subseteq> t\<close>
apply auto
using \<open>g ` t = S\<close> \<open>W \<subseteq> t\<close> apply blast
using g \<open>W \<subseteq> t\<close> apply auto[1]
by (simp add: f rev_image_eqI)
have \<circ>: "openin (top_of_set S) (g ` W)"
proof -
have "continuous_on S f"
using f(3) by blast
then show "openin (top_of_set S) (g ` W)"
by (simp add: gw Collect_conj_eq \<open>openin (top_of_set t) W\<close> continuous_on_open f(2))
qed
then obtain u v
where osu: "openin (top_of_set S) u" and uv: "P v" "g y \<in> u" "u \<subseteq> v" "v \<subseteq> g ` W"
using S [unfolded locally_def, rule_format, of "g ` W" "g y"] \<open>y \<in> W\<close> by force
have "v \<subseteq> S" using uv by (simp add: gw)
have fv: "f ` v = t \<inter> {x. g x \<in> v}"
using \<open>f ` S = t\<close> f \<open>v \<subseteq> S\<close> by auto
have "f ` v \<subseteq> W"
using uv using Int_lower2 gw image_subsetI mem_Collect_eq subset_iff by auto
have contvf: "continuous_on v f"
using \<open>v \<subseteq> S\<close> continuous_on_subset f(3) by blast
have contvg: "continuous_on (f ` v) g"
using \<open>f ` v \<subseteq> W\<close> \<open>W \<subseteq> t\<close> continuous_on_subset [OF g(3)] by blast
have homv: "homeomorphism v (f ` v) f g"
using \<open>v \<subseteq> S\<close> \<open>W \<subseteq> t\<close> f
apply (simp add: homeomorphism_def contvf contvg, auto)
by (metis f(1) rev_image_eqI rev_subsetD)
have 1: "openin (top_of_set t) (t \<inter> g -` u)"
apply (rule continuous_on_open [THEN iffD1, rule_format])
apply (rule \<open>continuous_on t g\<close>)
using \<open>g ` t = S\<close> apply (simp add: osu)
done
have 2: "\<exists>V. Q V \<and> y \<in> (t \<inter> g -` u) \<and> (t \<inter> g -` u) \<subseteq> V \<and> V \<subseteq> W"
apply (rule_tac x="f ` v" in exI)
apply (intro conjI Q [OF \<open>P v\<close> homv])
using \<open>W \<subseteq> t\<close> \<open>y \<in> W\<close> \<open>f ` v \<subseteq> W\<close> uv apply (auto simp: fv)
done
show "\<exists>U. openin (top_of_set t) U \<and> (\<exists>v. Q v \<and> y \<in> U \<and> U \<subseteq> v \<and> v \<subseteq> W)"
by (meson 1 2)
qed
lemma homeomorphism_locally:
fixes f:: "'a::metric_space \<Rightarrow> 'b::metric_space"
assumes hom: "homeomorphism S t f g"
and eq: "\<And>S t. homeomorphism S t f g \<Longrightarrow> (P S \<longleftrightarrow> Q t)"
shows "locally P S \<longleftrightarrow> locally Q t"
apply (rule iffI)
apply (erule homeomorphism_locally_imp [OF _ hom])
apply (simp add: eq)
apply (erule homeomorphism_locally_imp)
using eq homeomorphism_sym homeomorphism_symD [OF hom] apply blast+
done
lemma homeomorphic_locally:
fixes S:: "'a::metric_space set" and T:: "'b::metric_space set"
assumes hom: "S homeomorphic T"
and iff: "\<And>X Y. X homeomorphic Y \<Longrightarrow> (P X \<longleftrightarrow> Q Y)"
shows "locally P S \<longleftrightarrow> locally Q T"
proof -
obtain f g where hom: "homeomorphism S T f g"
using assms by (force simp: homeomorphic_def)
then show ?thesis
using homeomorphic_def local.iff
by (blast intro!: homeomorphism_locally)
qed
lemma homeomorphic_local_compactness:
fixes S:: "'a::metric_space set" and T:: "'b::metric_space set"
shows "S homeomorphic T \<Longrightarrow> locally compact S \<longleftrightarrow> locally compact T"
by (simp add: homeomorphic_compactness homeomorphic_locally)
lemma locally_translation:
fixes P :: "'a :: real_normed_vector set \<Rightarrow> bool"
shows
"(\<And>S. P (image (\<lambda>x. a + x) S) \<longleftrightarrow> P S)
\<Longrightarrow> locally P (image (\<lambda>x. a + x) S) \<longleftrightarrow> locally P S"
apply (rule homeomorphism_locally [OF homeomorphism_translation])
apply (simp add: homeomorphism_def)
by metis
lemma locally_injective_linear_image:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes f: "linear f" "inj f" and iff: "\<And>S. P (f ` S) \<longleftrightarrow> Q S"
shows "locally P (f ` S) \<longleftrightarrow> locally Q S"
apply (rule linear_homeomorphism_image [OF f])
apply (rule_tac f=g and g = f in homeomorphism_locally, assumption)
by (metis iff homeomorphism_def)
lemma locally_open_map_image:
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
assumes P: "locally P S"
and f: "continuous_on S f"
and oo: "\<And>t. openin (top_of_set S) t
\<Longrightarrow> openin (top_of_set (f ` S)) (f ` t)"
and Q: "\<And>t. \<lbrakk>t \<subseteq> S; P t\<rbrakk> \<Longrightarrow> Q(f ` t)"
shows "locally Q (f ` S)"
proof (clarsimp simp add: locally_def)
fix W y
assume oiw: "openin (top_of_set (f ` S)) W" and "y \<in> W"
then have "W \<subseteq> f ` S" by (simp add: openin_euclidean_subtopology_iff)
have oivf: "openin (top_of_set S) (S \<inter> f -` W)"
by (rule continuous_on_open [THEN iffD1, rule_format, OF f oiw])
then obtain x where "x \<in> S" "f x = y"
using \<open>W \<subseteq> f ` S\<close> \<open>y \<in> W\<close> by blast
then obtain U V
where "openin (top_of_set S) U" "P V" "x \<in> U" "U \<subseteq> V" "V \<subseteq> S \<inter> f -` W"
using P [unfolded locally_def, rule_format, of "(S \<inter> f -` W)" x] oivf \<open>y \<in> W\<close>
by auto
then show "\<exists>X. openin (top_of_set (f ` S)) X \<and> (\<exists>Y. Q Y \<and> y \<in> X \<and> X \<subseteq> Y \<and> Y \<subseteq> W)"
apply (rule_tac x="f ` U" in exI)
apply (rule conjI, blast intro!: oo)
apply (rule_tac x="f ` V" in exI)
apply (force simp: \<open>f x = y\<close> rev_image_eqI intro: Q)
done
qed
subsection\<open>An induction principle for connected sets\<close>
proposition connected_induction:
assumes "connected S"
and opD: "\<And>T a. \<lbrakk>openin (top_of_set S) T; a \<in> T\<rbrakk> \<Longrightarrow> \<exists>z. z \<in> T \<and> P z"
and opI: "\<And>a. a \<in> S
\<Longrightarrow> \<exists>T. openin (top_of_set S) T \<and> a \<in> T \<and>
(\<forall>x \<in> T. \<forall>y \<in> T. P x \<and> P y \<and> Q x \<longrightarrow> Q y)"
and etc: "a \<in> S" "b \<in> S" "P a" "P b" "Q a"
shows "Q b"
proof -
have 1: "openin (top_of_set S)
{b. \<exists>T. openin (top_of_set S) T \<and>
b \<in> T \<and> (\<forall>x\<in>T. P x \<longrightarrow> Q x)}"
apply (subst openin_subopen, clarify)
apply (rule_tac x=T in exI, auto)
done
have 2: "openin (top_of_set S)
{b. \<exists>T. openin (top_of_set S) T \<and>
b \<in> T \<and> (\<forall>x\<in>T. P x \<longrightarrow> \<not> Q x)}"
apply (subst openin_subopen, clarify)
apply (rule_tac x=T in exI, auto)
done
show ?thesis
using \<open>connected S\<close>
apply (simp only: connected_openin HOL.not_ex HOL.de_Morgan_conj)
apply (elim disjE allE)
apply (blast intro: 1)
apply (blast intro: 2, simp_all)
apply clarify apply (metis opI)
using opD apply (blast intro: etc elim: dest:)
using opI etc apply meson+
done
qed
lemma connected_equivalence_relation_gen:
assumes "connected S"
and etc: "a \<in> S" "b \<in> S" "P a" "P b"
and trans: "\<And>x y z. \<lbrakk>R x y; R y z\<rbrakk> \<Longrightarrow> R x z"
and opD: "\<And>T a. \<lbrakk>openin (top_of_set S) T; a \<in> T\<rbrakk> \<Longrightarrow> \<exists>z. z \<in> T \<and> P z"
and opI: "\<And>a. a \<in> S
\<Longrightarrow> \<exists>T. openin (top_of_set S) T \<and> a \<in> T \<and>
(\<forall>x \<in> T. \<forall>y \<in> T. P x \<and> P y \<longrightarrow> R x y)"
shows "R a b"
proof -
have "\<And>a b c. \<lbrakk>a \<in> S; P a; b \<in> S; c \<in> S; P b; P c; R a b\<rbrakk> \<Longrightarrow> R a c"
apply (rule connected_induction [OF \<open>connected S\<close> opD], simp_all)
by (meson trans opI)
then show ?thesis by (metis etc opI)
qed
lemma connected_induction_simple:
assumes "connected S"
and etc: "a \<in> S" "b \<in> S" "P a"
and opI: "\<And>a. a \<in> S
\<Longrightarrow> \<exists>T. openin (top_of_set S) T \<and> a \<in> T \<and>
(\<forall>x \<in> T. \<forall>y \<in> T. P x \<longrightarrow> P y)"
shows "P b"
apply (rule connected_induction [OF \<open>connected S\<close> _, where P = "\<lambda>x. True"], blast)
apply (frule opI)
using etc apply simp_all
done
lemma connected_equivalence_relation:
assumes "connected S"
and etc: "a \<in> S" "b \<in> S"
and sym: "\<And>x y. \<lbrakk>R x y; x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> R y x"
and trans: "\<And>x y z. \<lbrakk>R x y; R y z; x \<in> S; y \<in> S; z \<in> S\<rbrakk> \<Longrightarrow> R x z"
and opI: "\<And>a. a \<in> S \<Longrightarrow> \<exists>T. openin (top_of_set S) T \<and> a \<in> T \<and> (\<forall>x \<in> T. R a x)"
shows "R a b"
proof -
have "\<And>a b c. \<lbrakk>a \<in> S; b \<in> S; c \<in> S; R a b\<rbrakk> \<Longrightarrow> R a c"
apply (rule connected_induction_simple [OF \<open>connected S\<close>], simp_all)
by (meson local.sym local.trans opI openin_imp_subset subsetCE)
then show ?thesis by (metis etc opI)
qed
lemma locally_constant_imp_constant:
assumes "connected S"
and opI: "\<And>a. a \<in> S
\<Longrightarrow> \<exists>T. openin (top_of_set S) T \<and> a \<in> T \<and> (\<forall>x \<in> T. f x = f a)"
shows "f constant_on S"
proof -
have "\<And>x y. x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> f x = f y"
apply (rule connected_equivalence_relation [OF \<open>connected S\<close>], simp_all)
by (metis opI)
then show ?thesis
by (metis constant_on_def)
qed
lemma locally_constant:
"connected S \<Longrightarrow> locally (\<lambda>U. f constant_on U) S \<longleftrightarrow> f constant_on S"
apply (simp add: locally_def)
apply (rule iffI)
apply (rule locally_constant_imp_constant, assumption)
apply (metis (mono_tags, hide_lams) constant_on_def constant_on_subset openin_subtopology_self)
by (meson constant_on_subset openin_imp_subset order_refl)
subsection\<open>Basic properties of local compactness\<close>
proposition locally_compact:
fixes s :: "'a :: metric_space set"
shows
"locally compact s \<longleftrightarrow>
(\<forall>x \<in> s. \<exists>u v. x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<and>
openin (top_of_set s) u \<and> compact v)"
(is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
apply clarify
apply (erule_tac w = "s \<inter> ball x 1" in locallyE)
by auto
next
assume r [rule_format]: ?rhs
have *: "\<exists>u v.
openin (top_of_set s) u \<and>
compact v \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<inter> T"
if "open T" "x \<in> s" "x \<in> T" for x T
proof -
obtain u v where uv: "x \<in> u" "u \<subseteq> v" "v \<subseteq> s" "compact v" "openin (top_of_set s) u"
using r [OF \<open>x \<in> s\<close>] by auto
obtain e where "e>0" and e: "cball x e \<subseteq> T"
using open_contains_cball \<open>open T\<close> \<open>x \<in> T\<close> by blast
show ?thesis
apply (rule_tac x="(s \<inter> ball x e) \<inter> u" in exI)
apply (rule_tac x="cball x e \<inter> v" in exI)
using that \<open>e > 0\<close> e uv
apply auto
done
qed
show ?lhs
apply (rule locallyI)
apply (subst (asm) openin_open)
apply (blast intro: *)
done
qed
lemma locally_compactE:
fixes s :: "'a :: metric_space set"
assumes "locally compact s"
obtains u v where "\<And>x. x \<in> s \<Longrightarrow> x \<in> u x \<and> u x \<subseteq> v x \<and> v x \<subseteq> s \<and>
openin (top_of_set s) (u x) \<and> compact (v x)"
using assms
unfolding locally_compact by metis
lemma locally_compact_alt:
fixes s :: "'a :: heine_borel set"
shows "locally compact s \<longleftrightarrow>
(\<forall>x \<in> s. \<exists>u. x \<in> u \<and>
openin (top_of_set s) u \<and> compact(closure u) \<and> closure u \<subseteq> s)"
apply (simp add: locally_compact)
apply (intro ball_cong ex_cong refl iffI)
apply (metis bounded_subset closure_eq closure_mono compact_eq_bounded_closed dual_order.trans)
by (meson closure_subset compact_closure)
lemma locally_compact_Int_cball:
fixes s :: "'a :: heine_borel set"
shows "locally compact s \<longleftrightarrow> (\<forall>x \<in> s. \<exists>e. 0 < e \<and> closed(cball x e \<inter> s))"
(is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
apply (simp add: locally_compact openin_contains_cball)
apply (clarify | assumption | drule bspec)+
by (metis (no_types, lifting) compact_cball compact_imp_closed compact_Int inf.absorb_iff2 inf.orderE inf_sup_aci(2))
next
assume ?rhs
then show ?lhs
apply (simp add: locally_compact openin_contains_cball)
apply (clarify | assumption | drule bspec)+
apply (rule_tac x="ball x e \<inter> s" in exI, simp)
apply (rule_tac x="cball x e \<inter> s" in exI)
using compact_eq_bounded_closed
apply auto
apply (metis open_ball le_infI1 mem_ball open_contains_cball_eq)
done
qed
lemma locally_compact_compact:
fixes s :: "'a :: heine_borel set"
shows "locally compact s \<longleftrightarrow>
(\<forall>k. k \<subseteq> s \<and> compact k
\<longrightarrow> (\<exists>u v. k \<subseteq> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<and>
openin (top_of_set s) u \<and> compact v))"
(is "?lhs = ?rhs")
proof
assume ?lhs
then obtain u v where
uv: "\<And>x. x \<in> s \<Longrightarrow> x \<in> u x \<and> u x \<subseteq> v x \<and> v x \<subseteq> s \<and>
openin (top_of_set s) (u x) \<and> compact (v x)"
by (metis locally_compactE)
have *: "\<exists>u v. k \<subseteq> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<and> openin (top_of_set s) u \<and> compact v"
if "k \<subseteq> s" "compact k" for k
proof -
have "\<And>C. (\<forall>c\<in>C. openin (top_of_set k) c) \<and> k \<subseteq> \<Union>C \<Longrightarrow>
\<exists>D\<subseteq>C. finite D \<and> k \<subseteq> \<Union>D"
using that by (simp add: compact_eq_openin_cover)
moreover have "\<forall>c \<in> (\<lambda>x. k \<inter> u x) ` k. openin (top_of_set k) c"
using that by clarify (metis subsetD inf.absorb_iff2 openin_subset openin_subtopology_Int_subset topspace_euclidean_subtopology uv)
moreover have "k \<subseteq> \<Union>((\<lambda>x. k \<inter> u x) ` k)"
using that by clarsimp (meson subsetCE uv)
ultimately obtain D where "D \<subseteq> (\<lambda>x. k \<inter> u x) ` k" "finite D" "k \<subseteq> \<Union>D"
by metis
then obtain T where T: "T \<subseteq> k" "finite T" "k \<subseteq> \<Union>((\<lambda>x. k \<inter> u x) ` T)"
by (metis finite_subset_image)
have Tuv: "\<Union>(u ` T) \<subseteq> \<Union>(v ` T)"
using T that by (force simp: dest!: uv)
show ?thesis
apply (rule_tac x="\<Union>(u ` T)" in exI)
apply (rule_tac x="\<Union>(v ` T)" in exI)
apply (simp add: Tuv)
using T that
apply (auto simp: dest!: uv)
done
qed
show ?rhs
by (blast intro: *)
next
assume ?rhs
then show ?lhs
apply (clarsimp simp add: locally_compact)
apply (drule_tac x="{x}" in spec, simp)
done
qed
lemma open_imp_locally_compact:
fixes s :: "'a :: heine_borel set"
assumes "open s"
shows "locally compact s"
proof -
have *: "\<exists>u v. x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<and> openin (top_of_set s) u \<and> compact v"
if "x \<in> s" for x
proof -
obtain e where "e>0" and e: "cball x e \<subseteq> s"
using open_contains_cball assms \<open>x \<in> s\<close> by blast
have ope: "openin (top_of_set s) (ball x e)"
by (meson e open_ball ball_subset_cball dual_order.trans open_subset)
show ?thesis
apply (rule_tac x="ball x e" in exI)
apply (rule_tac x="cball x e" in exI)
using \<open>e > 0\<close> e apply (auto simp: ope)
done
qed
show ?thesis
unfolding locally_compact
by (blast intro: *)
qed
lemma closed_imp_locally_compact:
fixes s :: "'a :: heine_borel set"
assumes "closed s"
shows "locally compact s"
proof -
have *: "\<exists>u v. x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<and>
openin (top_of_set s) u \<and> compact v"
if "x \<in> s" for x
proof -
show ?thesis
apply (rule_tac x = "s \<inter> ball x 1" in exI)
apply (rule_tac x = "s \<inter> cball x 1" in exI)
using \<open>x \<in> s\<close> assms apply auto
done
qed
show ?thesis
unfolding locally_compact
by (blast intro: *)
qed
lemma locally_compact_UNIV: "locally compact (UNIV :: 'a :: heine_borel set)"
by (simp add: closed_imp_locally_compact)
lemma locally_compact_Int:
fixes s :: "'a :: t2_space set"
shows "\<lbrakk>locally compact s; locally compact t\<rbrakk> \<Longrightarrow> locally compact (s \<inter> t)"
by (simp add: compact_Int locally_Int)
lemma locally_compact_closedin:
fixes s :: "'a :: heine_borel set"
shows "\<lbrakk>closedin (top_of_set s) t; locally compact s\<rbrakk>
\<Longrightarrow> locally compact t"
unfolding closedin_closed
using closed_imp_locally_compact locally_compact_Int by blast
lemma locally_compact_delete:
fixes s :: "'a :: t1_space set"
shows "locally compact s \<Longrightarrow> locally compact (s - {a})"
by (auto simp: openin_delete locally_open_subset)
lemma locally_closed:
fixes s :: "'a :: heine_borel set"
shows "locally closed s \<longleftrightarrow> locally compact s"
(is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
apply (simp only: locally_def)
apply (erule all_forward imp_forward asm_rl exE)+
apply (rule_tac x = "u \<inter> ball x 1" in exI)
apply (rule_tac x = "v \<inter> cball x 1" in exI)
apply (force intro: openin_trans)
done
next
assume ?rhs then show ?lhs
using compact_eq_bounded_closed locally_mono by blast
qed
lemma locally_compact_openin_Un:
fixes S :: "'a::euclidean_space set"
assumes LCS: "locally compact S" and LCT:"locally compact T"
and opS: "openin (top_of_set (S \<union> T)) S"
and opT: "openin (top_of_set (S \<union> T)) T"
shows "locally compact (S \<union> T)"
proof -
have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if "x \<in> S" for x
proof -
obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> S)"
using LCS \<open>x \<in> S\<close> unfolding locally_compact_Int_cball by blast
moreover obtain e2 where "e2 > 0" and e2: "cball x e2 \<inter> (S \<union> T) \<subseteq> S"
by (meson \<open>x \<in> S\<close> opS openin_contains_cball)
then have "cball x e2 \<inter> (S \<union> T) = cball x e2 \<inter> S"
by force
ultimately show ?thesis
apply (rule_tac x="min e1 e2" in exI)
apply (auto simp: cball_min_Int \<open>e2 > 0\<close> inf_assoc closed_Int)
by (metis closed_Int closed_cball inf_left_commute)
qed
moreover have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if "x \<in> T" for x
proof -
obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> T)"
using LCT \<open>x \<in> T\<close> unfolding locally_compact_Int_cball by blast
moreover obtain e2 where "e2 > 0" and e2: "cball x e2 \<inter> (S \<union> T) \<subseteq> T"
by (meson \<open>x \<in> T\<close> opT openin_contains_cball)
then have "cball x e2 \<inter> (S \<union> T) = cball x e2 \<inter> T"
by force
ultimately show ?thesis
apply (rule_tac x="min e1 e2" in exI)
apply (auto simp: cball_min_Int \<open>e2 > 0\<close> inf_assoc closed_Int)
by (metis closed_Int closed_cball inf_left_commute)
qed
ultimately show ?thesis
by (force simp: locally_compact_Int_cball)
qed
lemma locally_compact_closedin_Un:
fixes S :: "'a::euclidean_space set"
assumes LCS: "locally compact S" and LCT:"locally compact T"
and clS: "closedin (top_of_set (S \<union> T)) S"
and clT: "closedin (top_of_set (S \<union> T)) T"
shows "locally compact (S \<union> T)"
proof -
have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if "x \<in> S" "x \<in> T" for x
proof -
obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> S)"
using LCS \<open>x \<in> S\<close> unfolding locally_compact_Int_cball by blast
moreover
obtain e2 where "e2 > 0" and e2: "closed (cball x e2 \<inter> T)"
using LCT \<open>x \<in> T\<close> unfolding locally_compact_Int_cball by blast
ultimately show ?thesis
apply (rule_tac x="min e1 e2" in exI)
apply (auto simp: cball_min_Int \<open>e2 > 0\<close> inf_assoc closed_Int Int_Un_distrib)
by (metis closed_Int closed_Un closed_cball inf_left_commute)
qed
moreover
have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if x: "x \<in> S" "x \<notin> T" for x
proof -
obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> S)"
using LCS \<open>x \<in> S\<close> unfolding locally_compact_Int_cball by blast
moreover
obtain e2 where "e2>0" and "cball x e2 \<inter> (S \<union> T) \<subseteq> S - T"
using clT x by (fastforce simp: openin_contains_cball closedin_def)
then have "closed (cball x e2 \<inter> T)"
proof -
have "{} = T - (T - cball x e2)"
using Diff_subset Int_Diff \<open>cball x e2 \<inter> (S \<union> T) \<subseteq> S - T\<close> by auto
then show ?thesis
by (simp add: Diff_Diff_Int inf_commute)
qed
ultimately show ?thesis
apply (rule_tac x="min e1 e2" in exI)
apply (auto simp: cball_min_Int \<open>e2 > 0\<close> inf_assoc closed_Int Int_Un_distrib)
by (metis closed_Int closed_Un closed_cball inf_left_commute)
qed
moreover
have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if x: "x \<notin> S" "x \<in> T" for x
proof -
obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> T)"
using LCT \<open>x \<in> T\<close> unfolding locally_compact_Int_cball by blast
moreover
obtain e2 where "e2>0" and "cball x e2 \<inter> (S \<union> T) \<subseteq> S \<union> T - S"
using clS x by (fastforce simp: openin_contains_cball closedin_def)
then have "closed (cball x e2 \<inter> S)"
by (metis Diff_disjoint Int_empty_right closed_empty inf.left_commute inf.orderE inf_sup_absorb)
ultimately show ?thesis
apply (rule_tac x="min e1 e2" in exI)
apply (auto simp: cball_min_Int \<open>e2 > 0\<close> inf_assoc closed_Int Int_Un_distrib)
by (metis closed_Int closed_Un closed_cball inf_left_commute)
qed
ultimately show ?thesis
by (auto simp: locally_compact_Int_cball)
qed
lemma locally_compact_Times:
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
shows "\<lbrakk>locally compact S; locally compact T\<rbrakk> \<Longrightarrow> locally compact (S \<times> T)"
by (auto simp: compact_Times locally_Times)
lemma locally_compact_compact_subopen:
fixes S :: "'a :: heine_borel set"
shows
"locally compact S \<longleftrightarrow>
(\<forall>K T. K \<subseteq> S \<and> compact K \<and> open T \<and> K \<subseteq> T
\<longrightarrow> (\<exists>U V. K \<subseteq> U \<and> U \<subseteq> V \<and> U \<subseteq> T \<and> V \<subseteq> S \<and>
openin (top_of_set S) U \<and> compact V))"
(is "?lhs = ?rhs")
proof
assume L: ?lhs
show ?rhs
proof clarify
fix K :: "'a set" and T :: "'a set"
assume "K \<subseteq> S" and "compact K" and "open T" and "K \<subseteq> T"
obtain U V where "K \<subseteq> U" "U \<subseteq> V" "V \<subseteq> S" "compact V"
and ope: "openin (top_of_set S) U"
using L unfolding locally_compact_compact by (meson \<open>K \<subseteq> S\<close> \<open>compact K\<close>)
show "\<exists>U V. K \<subseteq> U \<and> U \<subseteq> V \<and> U \<subseteq> T \<and> V \<subseteq> S \<and>
openin (top_of_set S) U \<and> compact V"
proof (intro exI conjI)
show "K \<subseteq> U \<inter> T"
by (simp add: \<open>K \<subseteq> T\<close> \<open>K \<subseteq> U\<close>)
show "U \<inter> T \<subseteq> closure(U \<inter> T)"
by (rule closure_subset)
show "closure (U \<inter> T) \<subseteq> S"
by (metis \<open>U \<subseteq> V\<close> \<open>V \<subseteq> S\<close> \<open>compact V\<close> closure_closed closure_mono compact_imp_closed inf.cobounded1 subset_trans)
show "openin (top_of_set S) (U \<inter> T)"
by (simp add: \<open>open T\<close> ope openin_Int_open)
show "compact (closure (U \<inter> T))"
by (meson Int_lower1 \<open>U \<subseteq> V\<close> \<open>compact V\<close> bounded_subset compact_closure compact_eq_bounded_closed)
qed auto
qed
next
assume ?rhs then show ?lhs
unfolding locally_compact_compact
by (metis open_openin openin_topspace subtopology_superset top.extremum topspace_euclidean_subtopology)
qed
subsection\<open>Sura-Bura's results about compact components of sets\<close>
proposition Sura_Bura_compact:
fixes S :: "'a::euclidean_space set"
assumes "compact S" and C: "C \<in> components S"
shows "C = \<Inter>{T. C \<subseteq> T \<and> openin (top_of_set S) T \<and>
closedin (top_of_set S) T}"
(is "C = \<Inter>?\<T>")
proof
obtain x where x: "C = connected_component_set S x" and "x \<in> S"
using C by (auto simp: components_def)
have "C \<subseteq> S"
by (simp add: C in_components_subset)
have "\<Inter>?\<T> \<subseteq> connected_component_set S x"
proof (rule connected_component_maximal)
have "x \<in> C"
by (simp add: \<open>x \<in> S\<close> x)
then show "x \<in> \<Inter>?\<T>"
by blast
have clo: "closed (\<Inter>?\<T>)"
by (simp add: \<open>compact S\<close> closed_Inter closedin_compact_eq compact_imp_closed)
have False
if K1: "closedin (top_of_set (\<Inter>?\<T>)) K1" and
K2: "closedin (top_of_set (\<Inter>?\<T>)) K2" and
K12_Int: "K1 \<inter> K2 = {}" and K12_Un: "K1 \<union> K2 = \<Inter>?\<T>" and "K1 \<noteq> {}" "K2 \<noteq> {}"
for K1 K2
proof -
have "closed K1" "closed K2"
using closedin_closed_trans clo K1 K2 by blast+
then obtain V1 V2 where "open V1" "open V2" "K1 \<subseteq> V1" "K2 \<subseteq> V2" and V12: "V1 \<inter> V2 = {}"
using separation_normal \<open>K1 \<inter> K2 = {}\<close> by metis
have SV12_ne: "(S - (V1 \<union> V2)) \<inter> (\<Inter>?\<T>) \<noteq> {}"
proof (rule compact_imp_fip)
show "compact (S - (V1 \<union> V2))"
by (simp add: \<open>open V1\<close> \<open>open V2\<close> \<open>compact S\<close> compact_diff open_Un)
show clo\<T>: "closed T" if "T \<in> ?\<T>" for T
using that \<open>compact S\<close>
by (force intro: closedin_closed_trans simp add: compact_imp_closed)
show "(S - (V1 \<union> V2)) \<inter> \<Inter>\<F> \<noteq> {}" if "finite \<F>" and \<F>: "\<F> \<subseteq> ?\<T>" for \<F>
proof
assume djo: "(S - (V1 \<union> V2)) \<inter> \<Inter>\<F> = {}"
obtain D where opeD: "openin (top_of_set S) D"
and cloD: "closedin (top_of_set S) D"
and "C \<subseteq> D" and DV12: "D \<subseteq> V1 \<union> V2"
proof (cases "\<F> = {}")
case True
with \<open>C \<subseteq> S\<close> djo that show ?thesis
by force
next
case False show ?thesis
proof
show ope: "openin (top_of_set S) (\<Inter>\<F>)"
using openin_Inter \<open>finite \<F>\<close> False \<F> by blast
then show "closedin (top_of_set S) (\<Inter>\<F>)"
by (meson clo\<T> \<F> closed_Inter closed_subset openin_imp_subset subset_eq)
show "C \<subseteq> \<Inter>\<F>"
using \<F> by auto
show "\<Inter>\<F> \<subseteq> V1 \<union> V2"
using ope djo openin_imp_subset by fastforce
qed
qed
have "connected C"
by (simp add: x)
have "closed D"
using \<open>compact S\<close> cloD closedin_closed_trans compact_imp_closed by blast
have cloV1: "closedin (top_of_set D) (D \<inter> closure V1)"
and cloV2: "closedin (top_of_set D) (D \<inter> closure V2)"
by (simp_all add: closedin_closed_Int)
moreover have "D \<inter> closure V1 = D \<inter> V1" "D \<inter> closure V2 = D \<inter> V2"
apply safe
using \<open>D \<subseteq> V1 \<union> V2\<close> \<open>open V1\<close> \<open>open V2\<close> V12
apply (simp_all add: closure_subset [THEN subsetD] closure_iff_nhds_not_empty, blast+)
done
ultimately have cloDV1: "closedin (top_of_set D) (D \<inter> V1)"
and cloDV2: "closedin (top_of_set D) (D \<inter> V2)"
by metis+
then obtain U1 U2 where "closed U1" "closed U2"
and D1: "D \<inter> V1 = D \<inter> U1" and D2: "D \<inter> V2 = D \<inter> U2"
by (auto simp: closedin_closed)
have "D \<inter> U1 \<inter> C \<noteq> {}"
proof
assume "D \<inter> U1 \<inter> C = {}"
then have *: "C \<subseteq> D \<inter> V2"
using D1 DV12 \<open>C \<subseteq> D\<close> by auto
have "\<Inter>?\<T> \<subseteq> D \<inter> V2"
apply (rule Inter_lower)
using * apply simp
by (meson cloDV2 \<open>open V2\<close> cloD closedin_trans le_inf_iff opeD openin_Int_open)
then show False
using K1 V12 \<open>K1 \<noteq> {}\<close> \<open>K1 \<subseteq> V1\<close> closedin_imp_subset by blast
qed
moreover have "D \<inter> U2 \<inter> C \<noteq> {}"
proof
assume "D \<inter> U2 \<inter> C = {}"
then have *: "C \<subseteq> D \<inter> V1"
using D2 DV12 \<open>C \<subseteq> D\<close> by auto
have "\<Inter>?\<T> \<subseteq> D \<inter> V1"
apply (rule Inter_lower)
using * apply simp
by (meson cloDV1 \<open>open V1\<close> cloD closedin_trans le_inf_iff opeD openin_Int_open)
then show False
using K2 V12 \<open>K2 \<noteq> {}\<close> \<open>K2 \<subseteq> V2\<close> closedin_imp_subset by blast
qed
ultimately show False
using \<open>connected C\<close> unfolding connected_closed
apply (simp only: not_ex)
apply (drule_tac x="D \<inter> U1" in spec)
apply (drule_tac x="D \<inter> U2" in spec)
using \<open>C \<subseteq> D\<close> D1 D2 V12 DV12 \<open>closed U1\<close> \<open>closed U2\<close> \<open>closed D\<close>
by blast
qed
qed
show False
by (metis (full_types) DiffE UnE Un_upper2 SV12_ne \<open>K1 \<subseteq> V1\<close> \<open>K2 \<subseteq> V2\<close> disjoint_iff_not_equal subsetCE sup_ge1 K12_Un)
qed
then show "connected (\<Inter>?\<T>)"
by (auto simp: connected_closedin_eq)
show "\<Inter>?\<T> \<subseteq> S"
by (fastforce simp: C in_components_subset)
qed
with x show "\<Inter>?\<T> \<subseteq> C" by simp
qed auto
corollary Sura_Bura_clopen_subset:
fixes S :: "'a::euclidean_space set"
assumes S: "locally compact S" and C: "C \<in> components S" and "compact C"
and U: "open U" "C \<subseteq> U"
obtains K where "openin (top_of_set S) K" "compact K" "C \<subseteq> K" "K \<subseteq> U"
proof (rule ccontr)
assume "\<not> thesis"
with that have neg: "\<nexists>K. openin (top_of_set S) K \<and> compact K \<and> C \<subseteq> K \<and> K \<subseteq> U"
by metis
obtain V K where "C \<subseteq> V" "V \<subseteq> U" "V \<subseteq> K" "K \<subseteq> S" "compact K"
and opeSV: "openin (top_of_set S) V"
using S U \<open>compact C\<close>
apply (simp add: locally_compact_compact_subopen)
by (meson C in_components_subset)
let ?\<T> = "{T. C \<subseteq> T \<and> openin (top_of_set K) T \<and> compact T \<and> T \<subseteq> K}"
have CK: "C \<in> components K"
by (meson C \<open>C \<subseteq> V\<close> \<open>K \<subseteq> S\<close> \<open>V \<subseteq> K\<close> components_intermediate_subset subset_trans)
with \<open>compact K\<close>
have "C = \<Inter>{T. C \<subseteq> T \<and> openin (top_of_set K) T \<and> closedin (top_of_set K) T}"
by (simp add: Sura_Bura_compact)
then have Ceq: "C = \<Inter>?\<T>"
by (simp add: closedin_compact_eq \<open>compact K\<close>)
obtain W where "open W" and W: "V = S \<inter> W"
using opeSV by (auto simp: openin_open)
have "-(U \<inter> W) \<inter> \<Inter>?\<T> \<noteq> {}"
proof (rule closed_imp_fip_compact)
show "- (U \<inter> W) \<inter> \<Inter>\<F> \<noteq> {}"
if "finite \<F>" and \<F>: "\<F> \<subseteq> ?\<T>" for \<F>
proof (cases "\<F> = {}")
case True
have False if "U = UNIV" "W = UNIV"
proof -
have "V = S"
by (simp add: W \<open>W = UNIV\<close>)
with neg show False
using \<open>C \<subseteq> V\<close> \<open>K \<subseteq> S\<close> \<open>V \<subseteq> K\<close> \<open>V \<subseteq> U\<close> \<open>compact K\<close> by auto
qed
with True show ?thesis
by auto
next
case False
show ?thesis
proof
assume "- (U \<inter> W) \<inter> \<Inter>\<F> = {}"
then have FUW: "\<Inter>\<F> \<subseteq> U \<inter> W"
by blast
have "C \<subseteq> \<Inter>\<F>"
using \<F> by auto
moreover have "compact (\<Inter>\<F>)"
by (metis (no_types, lifting) compact_Inter False mem_Collect_eq subsetCE \<F>)
moreover have "\<Inter>\<F> \<subseteq> K"
using False that(2) by fastforce
moreover have opeKF: "openin (top_of_set K) (\<Inter>\<F>)"
using False \<F> \<open>finite \<F>\<close> by blast
then have opeVF: "openin (top_of_set V) (\<Inter>\<F>)"
using W \<open>K \<subseteq> S\<close> \<open>V \<subseteq> K\<close> opeKF \<open>\<Inter>\<F> \<subseteq> K\<close> FUW openin_subset_trans by fastforce
then have "openin (top_of_set S) (\<Inter>\<F>)"
by (metis opeSV openin_trans)
moreover have "\<Inter>\<F> \<subseteq> U"
by (meson \<open>V \<subseteq> U\<close> opeVF dual_order.trans openin_imp_subset)
ultimately show False
using neg by blast
qed
qed
qed (use \<open>open W\<close> \<open>open U\<close> in auto)
with W Ceq \<open>C \<subseteq> V\<close> \<open>C \<subseteq> U\<close> show False
by auto
qed
corollary Sura_Bura_clopen_subset_alt:
fixes S :: "'a::euclidean_space set"
assumes S: "locally compact S" and C: "C \<in> components S" and "compact C"
and opeSU: "openin (top_of_set S) U" and "C \<subseteq> U"
obtains K where "openin (top_of_set S) K" "compact K" "C \<subseteq> K" "K \<subseteq> U"
proof -
obtain V where "open V" "U = S \<inter> V"
using opeSU by (auto simp: openin_open)
with \<open>C \<subseteq> U\<close> have "C \<subseteq> V"
by auto
then show ?thesis
using Sura_Bura_clopen_subset [OF S C \<open>compact C\<close> \<open>open V\<close>]
by (metis \<open>U = S \<inter> V\<close> inf.bounded_iff openin_imp_subset that)
qed
corollary Sura_Bura:
fixes S :: "'a::euclidean_space set"
assumes "locally compact S" "C \<in> components S" "compact C"
shows "C = \<Inter> {K. C \<subseteq> K \<and> compact K \<and> openin (top_of_set S) K}"
(is "C = ?rhs")
proof
show "?rhs \<subseteq> C"
proof (clarsimp, rule ccontr)
fix x
assume *: "\<forall>X. C \<subseteq> X \<and> compact X \<and> openin (top_of_set S) X \<longrightarrow> x \<in> X"
and "x \<notin> C"
obtain U V where "open U" "open V" "{x} \<subseteq> U" "C \<subseteq> V" "U \<inter> V = {}"
using separation_normal [of "{x}" C]
by (metis Int_empty_left \<open>x \<notin> C\<close> \<open>compact C\<close> closed_empty closed_insert compact_imp_closed insert_disjoint(1))
have "x \<notin> V"
using \<open>U \<inter> V = {}\<close> \<open>{x} \<subseteq> U\<close> by blast
then show False
by (meson "*" Sura_Bura_clopen_subset \<open>C \<subseteq> V\<close> \<open>open V\<close> assms(1) assms(2) assms(3) subsetCE)
qed
qed blast
subsection\<open>Special cases of local connectedness and path connectedness\<close>
lemma locally_connected_1:
assumes
"\<And>v x. \<lbrakk>openin (top_of_set S) v; x \<in> v\<rbrakk>
\<Longrightarrow> \<exists>u. openin (top_of_set S) u \<and>
connected u \<and> x \<in> u \<and> u \<subseteq> v"
shows "locally connected S"
apply (clarsimp simp add: locally_def)
apply (drule assms; blast)
done
lemma locally_connected_2:
assumes "locally connected S"
"openin (top_of_set S) t"
"x \<in> t"
shows "openin (top_of_set S) (connected_component_set t x)"
proof -
{ fix y :: 'a
let ?SS = "top_of_set S"
assume 1: "openin ?SS t"
"\<forall>w x. openin ?SS w \<and> x \<in> w \<longrightarrow> (\<exists>u. openin ?SS u \<and> (\<exists>v. connected v \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> w))"
and "connected_component t x y"
then have "y \<in> t" and y: "y \<in> connected_component_set t x"
using connected_component_subset by blast+
obtain F where
"\<forall>x y. (\<exists>w. openin ?SS w \<and> (\<exists>u. connected u \<and> x \<in> w \<and> w \<subseteq> u \<and> u \<subseteq> y)) = (openin ?SS (F x y) \<and> (\<exists>u. connected u \<and> x \<in> F x y \<and> F x y \<subseteq> u \<and> u \<subseteq> y))"
by moura
then obtain G where
"\<forall>a A. (\<exists>U. openin ?SS U \<and> (\<exists>V. connected V \<and> a \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> A)) = (openin ?SS (F a A) \<and> connected (G a A) \<and> a \<in> F a A \<and> F a A \<subseteq> G a A \<and> G a A \<subseteq> A)"
by moura
then have *: "openin ?SS (F y t) \<and> connected (G y t) \<and> y \<in> F y t \<and> F y t \<subseteq> G y t \<and> G y t \<subseteq> t"
using 1 \<open>y \<in> t\<close> by presburger
have "G y t \<subseteq> connected_component_set t y"
by (metis (no_types) * connected_component_eq_self connected_component_mono contra_subsetD)
then have "\<exists>A. openin ?SS A \<and> y \<in> A \<and> A \<subseteq> connected_component_set t x"
by (metis (no_types) * connected_component_eq dual_order.trans y)
}
then show ?thesis
using assms openin_subopen by (force simp: locally_def)
qed
lemma locally_connected_3:
assumes "\<And>t x. \<lbrakk>openin (top_of_set S) t; x \<in> t\<rbrakk>
\<Longrightarrow> openin (top_of_set S)
(connected_component_set t x)"
"openin (top_of_set S) v" "x \<in> v"
shows "\<exists>u. openin (top_of_set S) u \<and> connected u \<and> x \<in> u \<and> u \<subseteq> v"
using assms connected_component_subset by fastforce
lemma locally_connected:
"locally connected S \<longleftrightarrow>
(\<forall>v x. openin (top_of_set S) v \<and> x \<in> v
\<longrightarrow> (\<exists>u. openin (top_of_set S) u \<and> connected u \<and> x \<in> u \<and> u \<subseteq> v))"
by (metis locally_connected_1 locally_connected_2 locally_connected_3)
lemma locally_connected_open_connected_component:
"locally connected S \<longleftrightarrow>
(\<forall>t x. openin (top_of_set S) t \<and> x \<in> t
\<longrightarrow> openin (top_of_set S) (connected_component_set t x))"
by (metis locally_connected_1 locally_connected_2 locally_connected_3)
lemma locally_path_connected_1:
assumes
"\<And>v x. \<lbrakk>openin (top_of_set S) v; x \<in> v\<rbrakk>
\<Longrightarrow> \<exists>u. openin (top_of_set S) u \<and> path_connected u \<and> x \<in> u \<and> u \<subseteq> v"
shows "locally path_connected S"
apply (clarsimp simp add: locally_def)
apply (drule assms; blast)
done
lemma locally_path_connected_2:
assumes "locally path_connected S"
"openin (top_of_set S) t"
"x \<in> t"
shows "openin (top_of_set S) (path_component_set t x)"
proof -
{ fix y :: 'a
let ?SS = "top_of_set S"
assume 1: "openin ?SS t"
"\<forall>w x. openin ?SS w \<and> x \<in> w \<longrightarrow> (\<exists>u. openin ?SS u \<and> (\<exists>v. path_connected v \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> w))"
and "path_component t x y"
then have "y \<in> t" and y: "y \<in> path_component_set t x"
using path_component_mem(2) by blast+
obtain F where
"\<forall>x y. (\<exists>w. openin ?SS w \<and> (\<exists>u. path_connected u \<and> x \<in> w \<and> w \<subseteq> u \<and> u \<subseteq> y)) = (openin ?SS (F x y) \<and> (\<exists>u. path_connected u \<and> x \<in> F x y \<and> F x y \<subseteq> u \<and> u \<subseteq> y))"
by moura
then obtain G where
"\<forall>a A. (\<exists>U. openin ?SS U \<and> (\<exists>V. path_connected V \<and> a \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> A)) = (openin ?SS (F a A) \<and> path_connected (G a A) \<and> a \<in> F a A \<and> F a A \<subseteq> G a A \<and> G a A \<subseteq> A)"
by moura
then have *: "openin ?SS (F y t) \<and> path_connected (G y t) \<and> y \<in> F y t \<and> F y t \<subseteq> G y t \<and> G y t \<subseteq> t"
using 1 \<open>y \<in> t\<close> by presburger
have "G y t \<subseteq> path_component_set t y"
using * path_component_maximal rev_subsetD by blast
then have "\<exists>A. openin ?SS A \<and> y \<in> A \<and> A \<subseteq> path_component_set t x"
by (metis "*" \<open>G y t \<subseteq> path_component_set t y\<close> dual_order.trans path_component_eq y)
}
then show ?thesis
using assms openin_subopen by (force simp: locally_def)
qed
lemma locally_path_connected_3:
assumes "\<And>t x. \<lbrakk>openin (top_of_set S) t; x \<in> t\<rbrakk>
\<Longrightarrow> openin (top_of_set S) (path_component_set t x)"
"openin (top_of_set S) v" "x \<in> v"
shows "\<exists>u. openin (top_of_set S) u \<and> path_connected u \<and> x \<in> u \<and> u \<subseteq> v"
proof -
have "path_component v x x"
by (meson assms(3) path_component_refl)
then show ?thesis
by (metis assms(1) assms(2) assms(3) mem_Collect_eq path_component_subset path_connected_path_component)
qed
proposition locally_path_connected:
"locally path_connected S \<longleftrightarrow>
(\<forall>v x. openin (top_of_set S) v \<and> x \<in> v
\<longrightarrow> (\<exists>u. openin (top_of_set S) u \<and> path_connected u \<and> x \<in> u \<and> u \<subseteq> v))"
by (metis locally_path_connected_1 locally_path_connected_2 locally_path_connected_3)
proposition locally_path_connected_open_path_component:
"locally path_connected S \<longleftrightarrow>
(\<forall>t x. openin (top_of_set S) t \<and> x \<in> t
\<longrightarrow> openin (top_of_set S) (path_component_set t x))"
by (metis locally_path_connected_1 locally_path_connected_2 locally_path_connected_3)
lemma locally_connected_open_component:
"locally connected S \<longleftrightarrow>
(\<forall>t c. openin (top_of_set S) t \<and> c \<in> components t
\<longrightarrow> openin (top_of_set S) c)"
by (metis components_iff locally_connected_open_connected_component)
proposition locally_connected_im_kleinen:
"locally connected S \<longleftrightarrow>
(\<forall>v x. openin (top_of_set S) v \<and> x \<in> v
\<longrightarrow> (\<exists>u. openin (top_of_set S) u \<and>
x \<in> u \<and> u \<subseteq> v \<and>
(\<forall>y. y \<in> u \<longrightarrow> (\<exists>c. connected c \<and> c \<subseteq> v \<and> x \<in> c \<and> y \<in> c))))"
(is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
by (fastforce simp add: locally_connected)
next
assume ?rhs
have *: "\<exists>T. openin (top_of_set S) T \<and> x \<in> T \<and> T \<subseteq> c"
if "openin (top_of_set S) t" and c: "c \<in> components t" and "x \<in> c" for t c x
proof -
from that \<open>?rhs\<close> [rule_format, of t x]
obtain u where u:
"openin (top_of_set S) u \<and> x \<in> u \<and> u \<subseteq> t \<and>
(\<forall>y. y \<in> u \<longrightarrow> (\<exists>c. connected c \<and> c \<subseteq> t \<and> x \<in> c \<and> y \<in> c))"
using in_components_subset by auto
obtain F :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a" where
"\<forall>x y. (\<exists>z. z \<in> x \<and> y = connected_component_set x z) = (F x y \<in> x \<and> y = connected_component_set x (F x y))"
by moura
then have F: "F t c \<in> t \<and> c = connected_component_set t (F t c)"
by (meson components_iff c)
obtain G :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a" where
G: "\<forall>x y. (\<exists>z. z \<in> y \<and> z \<notin> x) = (G x y \<in> y \<and> G x y \<notin> x)"
by moura
have "G c u \<notin> u \<or> G c u \<in> c"
using F by (metis (full_types) u connected_componentI connected_component_eq mem_Collect_eq that(3))
then show ?thesis
using G u by auto
qed
show ?lhs
apply (clarsimp simp add: locally_connected_open_component)
apply (subst openin_subopen)
apply (blast intro: *)
done
qed
proposition locally_path_connected_im_kleinen:
"locally path_connected S \<longleftrightarrow>
(\<forall>v x. openin (top_of_set S) v \<and> x \<in> v
\<longrightarrow> (\<exists>u. openin (top_of_set S) u \<and>
x \<in> u \<and> u \<subseteq> v \<and>
(\<forall>y. y \<in> u \<longrightarrow> (\<exists>p. path p \<and> path_image p \<subseteq> v \<and>
pathstart p = x \<and> pathfinish p = y))))"
(is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
apply (simp add: locally_path_connected path_connected_def)
apply (erule all_forward ex_forward imp_forward conjE | simp)+
by (meson dual_order.trans)
next
assume ?rhs
have *: "\<exists>T. openin (top_of_set S) T \<and>
x \<in> T \<and> T \<subseteq> path_component_set u z"
if "openin (top_of_set S) u" and "z \<in> u" and c: "path_component u z x" for u z x
proof -
have "x \<in> u"
by (meson c path_component_mem(2))
with that \<open>?rhs\<close> [rule_format, of u x]
obtain U where U:
"openin (top_of_set S) U \<and> x \<in> U \<and> U \<subseteq> u \<and>
(\<forall>y. y \<in> U \<longrightarrow> (\<exists>p. path p \<and> path_image p \<subseteq> u \<and> pathstart p = x \<and> pathfinish p = y))"
by blast
show ?thesis
apply (rule_tac x=U in exI)
apply (auto simp: U)
apply (metis U c path_component_trans path_component_def)
done
qed
show ?lhs
apply (clarsimp simp add: locally_path_connected_open_path_component)
apply (subst openin_subopen)
apply (blast intro: *)
done
qed
lemma locally_path_connected_imp_locally_connected:
"locally path_connected S \<Longrightarrow> locally connected S"
using locally_mono path_connected_imp_connected by blast
lemma locally_connected_components:
"\<lbrakk>locally connected S; c \<in> components S\<rbrakk> \<Longrightarrow> locally connected c"
by (meson locally_connected_open_component locally_open_subset openin_subtopology_self)
lemma locally_path_connected_components:
"\<lbrakk>locally path_connected S; c \<in> components S\<rbrakk> \<Longrightarrow> locally path_connected c"
by (meson locally_connected_open_component locally_open_subset locally_path_connected_imp_locally_connected openin_subtopology_self)
lemma locally_path_connected_connected_component:
"locally path_connected S \<Longrightarrow> locally path_connected (connected_component_set S x)"
by (metis components_iff connected_component_eq_empty locally_empty locally_path_connected_components)
lemma open_imp_locally_path_connected:
fixes S :: "'a :: real_normed_vector set"
shows "open S \<Longrightarrow> locally path_connected S"
apply (rule locally_mono [of convex])
apply (simp_all add: locally_def openin_open_eq convex_imp_path_connected)
apply (meson open_ball centre_in_ball convex_ball openE order_trans)
done
lemma open_imp_locally_connected:
fixes S :: "'a :: real_normed_vector set"
shows "open S \<Longrightarrow> locally connected S"
by (simp add: locally_path_connected_imp_locally_connected open_imp_locally_path_connected)
lemma locally_path_connected_UNIV: "locally path_connected (UNIV::'a :: real_normed_vector set)"
by (simp add: open_imp_locally_path_connected)
lemma locally_connected_UNIV: "locally connected (UNIV::'a :: real_normed_vector set)"
by (simp add: open_imp_locally_connected)
lemma openin_connected_component_locally_connected:
"locally connected S
\<Longrightarrow> openin (top_of_set S) (connected_component_set S x)"
apply (simp add: locally_connected_open_connected_component)
by (metis connected_component_eq_empty connected_component_subset open_empty open_subset openin_subtopology_self)
lemma openin_components_locally_connected:
"\<lbrakk>locally connected S; c \<in> components S\<rbrakk> \<Longrightarrow> openin (top_of_set S) c"
using locally_connected_open_component openin_subtopology_self by blast
lemma openin_path_component_locally_path_connected:
"locally path_connected S
\<Longrightarrow> openin (top_of_set S) (path_component_set S x)"
by (metis (no_types) empty_iff locally_path_connected_2 openin_subopen openin_subtopology_self path_component_eq_empty)
lemma closedin_path_component_locally_path_connected:
"locally path_connected S
\<Longrightarrow> closedin (top_of_set S) (path_component_set S x)"
apply (simp add: closedin_def path_component_subset complement_path_component_Union)
apply (rule openin_Union)
using openin_path_component_locally_path_connected by auto
lemma convex_imp_locally_path_connected:
fixes S :: "'a:: real_normed_vector set"
shows "convex S \<Longrightarrow> locally path_connected S"
apply (clarsimp simp add: locally_path_connected)
apply (subst (asm) openin_open)
apply clarify
apply (erule (1) openE)
apply (rule_tac x = "S \<inter> ball x e" in exI)
apply (force simp: convex_Int convex_imp_path_connected)
done
lemma convex_imp_locally_connected:
fixes S :: "'a:: real_normed_vector set"
shows "convex S \<Longrightarrow> locally connected S"
by (simp add: locally_path_connected_imp_locally_connected convex_imp_locally_path_connected)
subsection\<open>Relations between components and path components\<close>
lemma path_component_eq_connected_component:
assumes "locally path_connected S"
shows "(path_component S x = connected_component S x)"
proof (cases "x \<in> S")
case True
have "openin (top_of_set (connected_component_set S x)) (path_component_set S x)"
apply (rule openin_subset_trans [of S])
apply (intro conjI openin_path_component_locally_path_connected [OF assms])
using path_component_subset_connected_component apply (auto simp: connected_component_subset)
done
moreover have "closedin (top_of_set (connected_component_set S x)) (path_component_set S x)"
apply (rule closedin_subset_trans [of S])
apply (intro conjI closedin_path_component_locally_path_connected [OF assms])
using path_component_subset_connected_component apply (auto simp: connected_component_subset)
done
ultimately have *: "path_component_set S x = connected_component_set S x"
by (metis connected_connected_component connected_clopen True path_component_eq_empty)
then show ?thesis
by blast
next
case False then show ?thesis
by (metis Collect_empty_eq_bot connected_component_eq_empty path_component_eq_empty)
qed
lemma path_component_eq_connected_component_set:
"locally path_connected S \<Longrightarrow> (path_component_set S x = connected_component_set S x)"
by (simp add: path_component_eq_connected_component)
lemma locally_path_connected_path_component:
"locally path_connected S \<Longrightarrow> locally path_connected (path_component_set S x)"
using locally_path_connected_connected_component path_component_eq_connected_component by fastforce
lemma open_path_connected_component:
fixes S :: "'a :: real_normed_vector set"
shows "open S \<Longrightarrow> path_component S x = connected_component S x"
by (simp add: path_component_eq_connected_component open_imp_locally_path_connected)
lemma open_path_connected_component_set:
fixes S :: "'a :: real_normed_vector set"
shows "open S \<Longrightarrow> path_component_set S x = connected_component_set S x"
by (simp add: open_path_connected_component)
proposition locally_connected_quotient_image:
assumes lcS: "locally connected S"
and oo: "\<And>T. T \<subseteq> f ` S
\<Longrightarrow> openin (top_of_set S) (S \<inter> f -` T) \<longleftrightarrow>
openin (top_of_set (f ` S)) T"
shows "locally connected (f ` S)"
proof (clarsimp simp: locally_connected_open_component)
fix U C
assume opefSU: "openin (top_of_set (f ` S)) U" and "C \<in> components U"
then have "C \<subseteq> U" "U \<subseteq> f ` S"
by (meson in_components_subset openin_imp_subset)+
then have "openin (top_of_set (f ` S)) C \<longleftrightarrow>
openin (top_of_set S) (S \<inter> f -` C)"
by (auto simp: oo)
moreover have "openin (top_of_set S) (S \<inter> f -` C)"
proof (subst openin_subopen, clarify)
fix x
assume "x \<in> S" "f x \<in> C"
show "\<exists>T. openin (top_of_set S) T \<and> x \<in> T \<and> T \<subseteq> (S \<inter> f -` C)"
proof (intro conjI exI)
show "openin (top_of_set S) (connected_component_set (S \<inter> f -` U) x)"
proof (rule ccontr)
assume **: "\<not> openin (top_of_set S) (connected_component_set (S \<inter> f -` U) x)"
then have "x \<notin> (S \<inter> f -` U)"
using \<open>U \<subseteq> f ` S\<close> opefSU lcS locally_connected_2 oo by blast
with ** show False
by (metis (no_types) connected_component_eq_empty empty_iff openin_subopen)
qed
next
show "x \<in> connected_component_set (S \<inter> f -` U) x"
using \<open>C \<subseteq> U\<close> \<open>f x \<in> C\<close> \<open>x \<in> S\<close> by auto
next
have contf: "continuous_on S f"
by (simp add: continuous_on_open oo openin_imp_subset)
then have "continuous_on (connected_component_set (S \<inter> f -` U) x) f"
apply (rule continuous_on_subset)
using connected_component_subset apply blast
done
then have "connected (f ` connected_component_set (S \<inter> f -` U) x)"
by (rule connected_continuous_image [OF _ connected_connected_component])
moreover have "f ` connected_component_set (S \<inter> f -` U) x \<subseteq> U"
using connected_component_in by blast
moreover have "C \<inter> f ` connected_component_set (S \<inter> f -` U) x \<noteq> {}"
using \<open>C \<subseteq> U\<close> \<open>f x \<in> C\<close> \<open>x \<in> S\<close> by fastforce
ultimately have fC: "f ` (connected_component_set (S \<inter> f -` U) x) \<subseteq> C"
by (rule components_maximal [OF \<open>C \<in> components U\<close>])
have cUC: "connected_component_set (S \<inter> f -` U) x \<subseteq> (S \<inter> f -` C)"
using connected_component_subset fC by blast
have "connected_component_set (S \<inter> f -` U) x \<subseteq> connected_component_set (S \<inter> f -` C) x"
proof -
{ assume "x \<in> connected_component_set (S \<inter> f -` U) x"
then have ?thesis
using cUC connected_component_idemp connected_component_mono by blast }
then show ?thesis
using connected_component_eq_empty by auto
qed
also have "\<dots> \<subseteq> (S \<inter> f -` C)"
by (rule connected_component_subset)
finally show "connected_component_set (S \<inter> f -` U) x \<subseteq> (S \<inter> f -` C)" .
qed
qed
ultimately show "openin (top_of_set (f ` S)) C"
by metis
qed
text\<open>The proof resembles that above but is not identical!\<close>
proposition locally_path_connected_quotient_image:
assumes lcS: "locally path_connected S"
and oo: "\<And>T. T \<subseteq> f ` S
\<Longrightarrow> openin (top_of_set S) (S \<inter> f -` T) \<longleftrightarrow> openin (top_of_set (f ` S)) T"
shows "locally path_connected (f ` S)"
proof (clarsimp simp: locally_path_connected_open_path_component)
fix U y
assume opefSU: "openin (top_of_set (f ` S)) U" and "y \<in> U"
then have "path_component_set U y \<subseteq> U" "U \<subseteq> f ` S"
by (meson path_component_subset openin_imp_subset)+
then have "openin (top_of_set (f ` S)) (path_component_set U y) \<longleftrightarrow>
openin (top_of_set S) (S \<inter> f -` path_component_set U y)"
proof -
have "path_component_set U y \<subseteq> f ` S"
using \<open>U \<subseteq> f ` S\<close> \<open>path_component_set U y \<subseteq> U\<close> by blast
then show ?thesis
using oo by blast
qed
moreover have "openin (top_of_set S) (S \<inter> f -` path_component_set U y)"
proof (subst openin_subopen, clarify)
fix x
assume "x \<in> S" and Uyfx: "path_component U y (f x)"
then have "f x \<in> U"
using path_component_mem by blast
show "\<exists>T. openin (top_of_set S) T \<and> x \<in> T \<and> T \<subseteq> (S \<inter> f -` path_component_set U y)"
proof (intro conjI exI)
show "openin (top_of_set S) (path_component_set (S \<inter> f -` U) x)"
proof (rule ccontr)
assume **: "\<not> openin (top_of_set S) (path_component_set (S \<inter> f -` U) x)"
then have "x \<notin> (S \<inter> f -` U)"
by (metis (no_types, lifting) \<open>U \<subseteq> f ` S\<close> opefSU lcS oo locally_path_connected_open_path_component)
then show False
using ** \<open>path_component_set U y \<subseteq> U\<close> \<open>x \<in> S\<close> \<open>path_component U y (f x)\<close> by blast
qed
next
show "x \<in> path_component_set (S \<inter> f -` U) x"
by (simp add: \<open>f x \<in> U\<close> \<open>x \<in> S\<close> path_component_refl)
next
have contf: "continuous_on S f"
by (simp add: continuous_on_open oo openin_imp_subset)
then have "continuous_on (path_component_set (S \<inter> f -` U) x) f"
apply (rule continuous_on_subset)
using path_component_subset apply blast
done
then have "path_connected (f ` path_component_set (S \<inter> f -` U) x)"
by (simp add: path_connected_continuous_image)
moreover have "f ` path_component_set (S \<inter> f -` U) x \<subseteq> U"
using path_component_mem by fastforce
moreover have "f x \<in> f ` path_component_set (S \<inter> f -` U) x"
by (force simp: \<open>x \<in> S\<close> \<open>f x \<in> U\<close> path_component_refl_eq)
ultimately have "f ` (path_component_set (S \<inter> f -` U) x) \<subseteq> path_component_set U (f x)"
by (meson path_component_maximal)
also have "\<dots> \<subseteq> path_component_set U y"
by (simp add: Uyfx path_component_maximal path_component_subset path_component_sym)
finally have fC: "f ` (path_component_set (S \<inter> f -` U) x) \<subseteq> path_component_set U y" .
have cUC: "path_component_set (S \<inter> f -` U) x \<subseteq> (S \<inter> f -` path_component_set U y)"
using path_component_subset fC by blast
have "path_component_set (S \<inter> f -` U) x \<subseteq> path_component_set (S \<inter> f -` path_component_set U y) x"
proof -
have "\<And>a. path_component_set (path_component_set (S \<inter> f -` U) x) a \<subseteq> path_component_set (S \<inter> f -` path_component_set U y) a"
using cUC path_component_mono by blast
then show ?thesis
using path_component_path_component by blast
qed
also have "\<dots> \<subseteq> (S \<inter> f -` path_component_set U y)"
by (rule path_component_subset)
finally show "path_component_set (S \<inter> f -` U) x \<subseteq> (S \<inter> f -` path_component_set U y)" .
qed
qed
ultimately show "openin (top_of_set (f ` S)) (path_component_set U y)"
by metis
qed
subsection\<^marker>\<open>tag unimportant\<close>\<open>Components, continuity, openin, closedin\<close>
lemma continuous_on_components_gen:
fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
assumes "\<And>c. c \<in> components S \<Longrightarrow>
openin (top_of_set S) c \<and> continuous_on c f"
shows "continuous_on S f"
proof (clarsimp simp: continuous_openin_preimage_eq)
fix t :: "'b set"
assume "open t"
have *: "S \<inter> f -` t = (\<Union>c \<in> components S. c \<inter> f -` t)"
by auto
show "openin (top_of_set S) (S \<inter> f -` t)"
unfolding * using \<open>open t\<close> assms continuous_openin_preimage_gen openin_trans openin_Union by blast
qed
lemma continuous_on_components:
fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
assumes "locally connected S "
"\<And>c. c \<in> components S \<Longrightarrow> continuous_on c f"
shows "continuous_on S f"
apply (rule continuous_on_components_gen)
apply (auto simp: assms intro: openin_components_locally_connected)
done
lemma continuous_on_components_eq:
"locally connected S
\<Longrightarrow> (continuous_on S f \<longleftrightarrow> (\<forall>c \<in> components S. continuous_on c f))"
by (meson continuous_on_components continuous_on_subset in_components_subset)
lemma continuous_on_components_open:
fixes S :: "'a::real_normed_vector set"
assumes "open S "
"\<And>c. c \<in> components S \<Longrightarrow> continuous_on c f"
shows "continuous_on S f"
using continuous_on_components open_imp_locally_connected assms by blast
lemma continuous_on_components_open_eq:
fixes S :: "'a::real_normed_vector set"
shows "open S \<Longrightarrow> (continuous_on S f \<longleftrightarrow> (\<forall>c \<in> components S. continuous_on c f))"
using continuous_on_subset in_components_subset
by (blast intro: continuous_on_components_open)
lemma closedin_union_complement_components:
assumes u: "locally connected u"
and S: "closedin (top_of_set u) S"
and cuS: "c \<subseteq> components(u - S)"
shows "closedin (top_of_set u) (S \<union> \<Union>c)"
proof -
have di: "(\<And>S t. S \<in> c \<and> t \<in> c' \<Longrightarrow> disjnt S t) \<Longrightarrow> disjnt (\<Union> c) (\<Union> c')" for c'
by (simp add: disjnt_def) blast
have "S \<subseteq> u"
using S closedin_imp_subset by blast
moreover have "u - S = \<Union>c \<union> \<Union>(components (u - S) - c)"
by (metis Diff_partition Union_components Union_Un_distrib assms(3))
moreover have "disjnt (\<Union>c) (\<Union>(components (u - S) - c))"
apply (rule di)
by (metis DiffD1 DiffD2 assms(3) components_nonoverlap disjnt_def subsetCE)
ultimately have eq: "S \<union> \<Union>c = u - (\<Union>(components(u - S) - c))"
by (auto simp: disjnt_def)
have *: "openin (top_of_set u) (\<Union>(components (u - S) - c))"
apply (rule openin_Union)
apply (rule openin_trans [of "u - S"])
apply (simp add: u S locally_diff_closed openin_components_locally_connected)
apply (simp add: openin_diff S)
done
have "openin (top_of_set u) (u - (u - \<Union>(components (u - S) - c)))"
apply (rule openin_diff, simp)
apply (metis closedin_diff closedin_topspace topspace_euclidean_subtopology *)
done
then show ?thesis
by (force simp: eq closedin_def)
qed
lemma closed_union_complement_components:
fixes S :: "'a::real_normed_vector set"
assumes S: "closed S" and c: "c \<subseteq> components(- S)"
shows "closed(S \<union> \<Union> c)"
proof -
have "closedin (top_of_set UNIV) (S \<union> \<Union>c)"
apply (rule closedin_union_complement_components [OF locally_connected_UNIV])
using S c apply (simp_all add: Compl_eq_Diff_UNIV)
done
then show ?thesis by simp
qed
lemma closedin_Un_complement_component:
fixes S :: "'a::real_normed_vector set"
assumes u: "locally connected u"
and S: "closedin (top_of_set u) S"
and c: " c \<in> components(u - S)"
shows "closedin (top_of_set u) (S \<union> c)"
proof -
have "closedin (top_of_set u) (S \<union> \<Union>{c})"
using c by (blast intro: closedin_union_complement_components [OF u S])
then show ?thesis
by simp
qed
lemma closed_Un_complement_component:
fixes S :: "'a::real_normed_vector set"
assumes S: "closed S" and c: " c \<in> components(-S)"
shows "closed (S \<union> c)"
by (metis Compl_eq_Diff_UNIV S c closed_closedin closedin_Un_complement_component
locally_connected_UNIV subtopology_UNIV)
subsection\<open>Existence of isometry between subspaces of same dimension\<close>
lemma isometry_subset_subspace:
fixes S :: "'a::euclidean_space set"
and T :: "'b::euclidean_space set"
assumes S: "subspace S"
and T: "subspace T"
and d: "dim S \<le> dim T"
obtains f where "linear f" "f ` S \<subseteq> T" "\<And>x. x \<in> S \<Longrightarrow> norm(f x) = norm x"
proof -
obtain B where "B \<subseteq> S" and Borth: "pairwise orthogonal B"
and B1: "\<And>x. x \<in> B \<Longrightarrow> norm x = 1"
and "independent B" "finite B" "card B = dim S" "span B = S"
by (metis orthonormal_basis_subspace [OF S] independent_finite)
obtain C where "C \<subseteq> T" and Corth: "pairwise orthogonal C"
and C1:"\<And>x. x \<in> C \<Longrightarrow> norm x = 1"
and "independent C" "finite C" "card C = dim T" "span C = T"
by (metis orthonormal_basis_subspace [OF T] independent_finite)
obtain fb where "fb ` B \<subseteq> C" "inj_on fb B"
by (metis \<open>card B = dim S\<close> \<open>card C = dim T\<close> \<open>finite B\<close> \<open>finite C\<close> card_le_inj d)
then have pairwise_orth_fb: "pairwise (\<lambda>v j. orthogonal (fb v) (fb j)) B"
using Corth
apply (auto simp: pairwise_def orthogonal_clauses)
by (meson subsetD image_eqI inj_on_def)
obtain f where "linear f" and ffb: "\<And>x. x \<in> B \<Longrightarrow> f x = fb x"
using linear_independent_extend \<open>independent B\<close> by fastforce
have "span (f ` B) \<subseteq> span C"
by (metis \<open>fb ` B \<subseteq> C\<close> ffb image_cong span_mono)
then have "f ` S \<subseteq> T"
unfolding \<open>span B = S\<close> \<open>span C = T\<close> span_linear_image[OF \<open>linear f\<close>] .
have [simp]: "\<And>x. x \<in> B \<Longrightarrow> norm (fb x) = norm x"
using B1 C1 \<open>fb ` B \<subseteq> C\<close> by auto
have "norm (f x) = norm x" if "x \<in> S" for x
proof -
interpret linear f by fact
obtain a where x: "x = (\<Sum>v \<in> B. a v *\<^sub>R v)"
using \<open>finite B\<close> \<open>span B = S\<close> \<open>x \<in> S\<close> span_finite by fastforce
have "norm (f x)^2 = norm (\<Sum>v\<in>B. a v *\<^sub>R fb v)^2" by (simp add: sum scale ffb x)
also have "\<dots> = (\<Sum>v\<in>B. norm ((a v *\<^sub>R fb v))^2)"
apply (rule norm_sum_Pythagorean [OF \<open>finite B\<close>])
apply (rule pairwise_ortho_scaleR [OF pairwise_orth_fb])
done
also have "\<dots> = norm x ^2"
by (simp add: x pairwise_ortho_scaleR Borth norm_sum_Pythagorean [OF \<open>finite B\<close>])
finally show ?thesis
by (simp add: norm_eq_sqrt_inner)
qed
then show ?thesis
by (rule that [OF \<open>linear f\<close> \<open>f ` S \<subseteq> T\<close>])
qed
proposition isometries_subspaces:
fixes S :: "'a::euclidean_space set"
and T :: "'b::euclidean_space set"
assumes S: "subspace S"
and T: "subspace T"
and d: "dim S = dim T"
obtains f g where "linear f" "linear g" "f ` S = T" "g ` T = S"
"\<And>x. x \<in> S \<Longrightarrow> norm(f x) = norm x"
"\<And>x. x \<in> T \<Longrightarrow> norm(g x) = norm x"
"\<And>x. x \<in> S \<Longrightarrow> g(f x) = x"
"\<And>x. x \<in> T \<Longrightarrow> f(g x) = x"
proof -
obtain B where "B \<subseteq> S" and Borth: "pairwise orthogonal B"
and B1: "\<And>x. x \<in> B \<Longrightarrow> norm x = 1"
and "independent B" "finite B" "card B = dim S" "span B = S"
by (metis orthonormal_basis_subspace [OF S] independent_finite)
obtain C where "C \<subseteq> T" and Corth: "pairwise orthogonal C"
and C1:"\<And>x. x \<in> C \<Longrightarrow> norm x = 1"
and "independent C" "finite C" "card C = dim T" "span C = T"
by (metis orthonormal_basis_subspace [OF T] independent_finite)
obtain fb where "bij_betw fb B C"
by (metis \<open>finite B\<close> \<open>finite C\<close> bij_betw_iff_card \<open>card B = dim S\<close> \<open>card C = dim T\<close> d)
then have pairwise_orth_fb: "pairwise (\<lambda>v j. orthogonal (fb v) (fb j)) B"
using Corth
apply (auto simp: pairwise_def orthogonal_clauses bij_betw_def)
by (meson subsetD image_eqI inj_on_def)
obtain f where "linear f" and ffb: "\<And>x. x \<in> B \<Longrightarrow> f x = fb x"
using linear_independent_extend \<open>independent B\<close> by fastforce
interpret f: linear f by fact
define gb where "gb \<equiv> inv_into B fb"
then have pairwise_orth_gb: "pairwise (\<lambda>v j. orthogonal (gb v) (gb j)) C"
using Borth
apply (auto simp: pairwise_def orthogonal_clauses bij_betw_def)
by (metis \<open>bij_betw fb B C\<close> bij_betw_imp_surj_on bij_betw_inv_into_right inv_into_into)
obtain g where "linear g" and ggb: "\<And>x. x \<in> C \<Longrightarrow> g x = gb x"
using linear_independent_extend \<open>independent C\<close> by fastforce
interpret g: linear g by fact
have "span (f ` B) \<subseteq> span C"
by (metis \<open>bij_betw fb B C\<close> bij_betw_imp_surj_on eq_iff ffb image_cong)
then have "f ` S \<subseteq> T"
unfolding \<open>span B = S\<close> \<open>span C = T\<close>
span_linear_image[OF \<open>linear f\<close>] .
have [simp]: "\<And>x. x \<in> B \<Longrightarrow> norm (fb x) = norm x"
using B1 C1 \<open>bij_betw fb B C\<close> bij_betw_imp_surj_on by fastforce
have f [simp]: "norm (f x) = norm x" "g (f x) = x" if "x \<in> S" for x
proof -
obtain a where x: "x = (\<Sum>v \<in> B. a v *\<^sub>R v)"
using \<open>finite B\<close> \<open>span B = S\<close> \<open>x \<in> S\<close> span_finite by fastforce
have "f x = (\<Sum>v \<in> B. f (a v *\<^sub>R v))"
using linear_sum [OF \<open>linear f\<close>] x by auto
also have "\<dots> = (\<Sum>v \<in> B. a v *\<^sub>R f v)"
by (simp add: f.sum f.scale)
also have "\<dots> = (\<Sum>v \<in> B. a v *\<^sub>R fb v)"
by (simp add: ffb cong: sum.cong)
finally have *: "f x = (\<Sum>v\<in>B. a v *\<^sub>R fb v)" .
then have "(norm (f x))\<^sup>2 = (norm (\<Sum>v\<in>B. a v *\<^sub>R fb v))\<^sup>2" by simp
also have "\<dots> = (\<Sum>v\<in>B. norm ((a v *\<^sub>R fb v))^2)"
apply (rule norm_sum_Pythagorean [OF \<open>finite B\<close>])
apply (rule pairwise_ortho_scaleR [OF pairwise_orth_fb])
done
also have "\<dots> = (norm x)\<^sup>2"
by (simp add: x pairwise_ortho_scaleR Borth norm_sum_Pythagorean [OF \<open>finite B\<close>])
finally show "norm (f x) = norm x"
by (simp add: norm_eq_sqrt_inner)
have "g (f x) = g (\<Sum>v\<in>B. a v *\<^sub>R fb v)" by (simp add: *)
also have "\<dots> = (\<Sum>v\<in>B. g (a v *\<^sub>R fb v))"
by (simp add: g.sum g.scale)
also have "\<dots> = (\<Sum>v\<in>B. a v *\<^sub>R g (fb v))"
by (simp add: g.scale)
also have "\<dots> = (\<Sum>v\<in>B. a v *\<^sub>R v)"
apply (rule sum.cong [OF refl])
using \<open>bij_betw fb B C\<close> gb_def bij_betwE bij_betw_inv_into_left gb_def ggb by fastforce
also have "\<dots> = x"
using x by blast
finally show "g (f x) = x" .
qed
have [simp]: "\<And>x. x \<in> C \<Longrightarrow> norm (gb x) = norm x"
by (metis B1 C1 \<open>bij_betw fb B C\<close> bij_betw_imp_surj_on gb_def inv_into_into)
have g [simp]: "f (g x) = x" if "x \<in> T" for x
proof -
obtain a where x: "x = (\<Sum>v \<in> C. a v *\<^sub>R v)"
using \<open>finite C\<close> \<open>span C = T\<close> \<open>x \<in> T\<close> span_finite by fastforce
have "g x = (\<Sum>v \<in> C. g (a v *\<^sub>R v))"
by (simp add: x g.sum)
also have "\<dots> = (\<Sum>v \<in> C. a v *\<^sub>R g v)"
by (simp add: g.scale)
also have "\<dots> = (\<Sum>v \<in> C. a v *\<^sub>R gb v)"
by (simp add: ggb cong: sum.cong)
finally have "f (g x) = f (\<Sum>v\<in>C. a v *\<^sub>R gb v)" by simp
also have "\<dots> = (\<Sum>v\<in>C. f (a v *\<^sub>R gb v))"
by (simp add: f.scale f.sum)
also have "\<dots> = (\<Sum>v\<in>C. a v *\<^sub>R f (gb v))"
by (simp add: f.scale f.sum)
also have "\<dots> = (\<Sum>v\<in>C. a v *\<^sub>R v)"
using \<open>bij_betw fb B C\<close>
by (simp add: bij_betw_def gb_def bij_betw_inv_into_right ffb inv_into_into)
also have "\<dots> = x"
using x by blast
finally show "f (g x) = x" .
qed
have gim: "g ` T = S"
by (metis (full_types) S T \<open>f ` S \<subseteq> T\<close> d dim_eq_span dim_image_le f(2) g.linear_axioms
image_iff linear_subspace_image span_eq_iff subset_iff)
have fim: "f ` S = T"
using \<open>g ` T = S\<close> image_iff by fastforce
have [simp]: "norm (g x) = norm x" if "x \<in> T" for x
using fim that by auto
show ?thesis
apply (rule that [OF \<open>linear f\<close> \<open>linear g\<close>])
apply (simp_all add: fim gim)
done
qed
corollary isometry_subspaces:
fixes S :: "'a::euclidean_space set"
and T :: "'b::euclidean_space set"
assumes S: "subspace S"
and T: "subspace T"
and d: "dim S = dim T"
obtains f where "linear f" "f ` S = T" "\<And>x. x \<in> S \<Longrightarrow> norm(f x) = norm x"
using isometries_subspaces [OF assms]
by metis
corollary isomorphisms_UNIV_UNIV:
assumes "DIM('M) = DIM('N)"
obtains f::"'M::euclidean_space \<Rightarrow>'N::euclidean_space" and g
where "linear f" "linear g"
"\<And>x. norm(f x) = norm x" "\<And>y. norm(g y) = norm y"
"\<And>x. g (f x) = x" "\<And>y. f(g y) = y"
using assms by (auto intro: isometries_subspaces [of "UNIV::'M set" "UNIV::'N set"])
lemma homeomorphic_subspaces:
fixes S :: "'a::euclidean_space set"
and T :: "'b::euclidean_space set"
assumes S: "subspace S"
and T: "subspace T"
and d: "dim S = dim T"
shows "S homeomorphic T"
proof -
obtain f g where "linear f" "linear g" "f ` S = T" "g ` T = S"
"\<And>x. x \<in> S \<Longrightarrow> g(f x) = x" "\<And>x. x \<in> T \<Longrightarrow> f(g x) = x"
by (blast intro: isometries_subspaces [OF assms])
then show ?thesis
apply (simp add: homeomorphic_def homeomorphism_def)
apply (rule_tac x=f in exI)
apply (rule_tac x=g in exI)
apply (auto simp: linear_continuous_on linear_conv_bounded_linear)
done
qed
lemma homeomorphic_affine_sets:
assumes "affine S" "affine T" "aff_dim S = aff_dim T"
shows "S homeomorphic T"
proof (cases "S = {} \<or> T = {}")
case True with assms aff_dim_empty homeomorphic_empty show ?thesis
by metis
next
case False
then obtain a b where ab: "a \<in> S" "b \<in> T" by auto
then have ss: "subspace ((+) (- a) ` S)" "subspace ((+) (- b) ` T)"
using affine_diffs_subspace assms by blast+
have dd: "dim ((+) (- a) ` S) = dim ((+) (- b) ` T)"
using assms ab by (simp add: aff_dim_eq_dim [OF hull_inc] image_def)
have "S homeomorphic ((+) (- a) ` S)"
by (fact homeomorphic_translation)
also have "\<dots> homeomorphic ((+) (- b) ` T)"
by (rule homeomorphic_subspaces [OF ss dd])
also have "\<dots> homeomorphic T"
using homeomorphic_translation [of T "- b"] by (simp add: homeomorphic_sym [of T])
finally show ?thesis .
qed
subsection\<open>Retracts, in a general sense, preserve (co)homotopic triviality)\<close>
locale\<^marker>\<open>tag important\<close> Retracts =
fixes s h t k
assumes conth: "continuous_on s h"
and imh: "h ` s = t"
and contk: "continuous_on t k"
and imk: "k ` t \<subseteq> s"
and idhk: "\<And>y. y \<in> t \<Longrightarrow> h(k y) = y"
begin
lemma homotopically_trivial_retraction_gen:
assumes P: "\<And>f. \<lbrakk>continuous_on u f; f ` u \<subseteq> t; Q f\<rbrakk> \<Longrightarrow> P(k \<circ> f)"
and Q: "\<And>f. \<lbrakk>continuous_on u f; f ` u \<subseteq> s; P f\<rbrakk> \<Longrightarrow> Q(h \<circ> f)"
and Qeq: "\<And>h k. (\<And>x. x \<in> u \<Longrightarrow> h x = k x) \<Longrightarrow> Q h = Q k"
and hom: "\<And>f g. \<lbrakk>continuous_on u f; f ` u \<subseteq> s; P f;
continuous_on u g; g ` u \<subseteq> s; P g\<rbrakk>
\<Longrightarrow> homotopic_with_canon P u s f g"
and contf: "continuous_on u f" and imf: "f ` u \<subseteq> t" and Qf: "Q f"
and contg: "continuous_on u g" and img: "g ` u \<subseteq> t" and Qg: "Q g"
shows "homotopic_with_canon Q u t f g"
proof -
have feq: "\<And>x. x \<in> u \<Longrightarrow> (h \<circ> (k \<circ> f)) x = f x" using idhk imf by auto
have geq: "\<And>x. x \<in> u \<Longrightarrow> (h \<circ> (k \<circ> g)) x = g x" using idhk img by auto
have "continuous_on u (k \<circ> f)"
using contf continuous_on_compose continuous_on_subset contk imf by blast
moreover have "(k \<circ> f) ` u \<subseteq> s"
using imf imk by fastforce
moreover have "P (k \<circ> f)"
by (simp add: P Qf contf imf)
moreover have "continuous_on u (k \<circ> g)"
using contg continuous_on_compose continuous_on_subset contk img by blast
moreover have "(k \<circ> g) ` u \<subseteq> s"
using img imk by fastforce
moreover have "P (k \<circ> g)"
by (simp add: P Qg contg img)
ultimately have "homotopic_with_canon P u s (k \<circ> f) (k \<circ> g)"
by (rule hom)
then have "homotopic_with_canon Q u t (h \<circ> (k \<circ> f)) (h \<circ> (k \<circ> g))"
apply (rule homotopic_with_compose_continuous_left [OF homotopic_with_mono])
using Q by (auto simp: conth imh)
then show ?thesis
apply (rule homotopic_with_eq)
using feq geq apply fastforce+
using Qeq topspace_euclidean_subtopology by blast
qed
lemma homotopically_trivial_retraction_null_gen:
assumes P: "\<And>f. \<lbrakk>continuous_on u f; f ` u \<subseteq> t; Q f\<rbrakk> \<Longrightarrow> P(k \<circ> f)"
and Q: "\<And>f. \<lbrakk>continuous_on u f; f ` u \<subseteq> s; P f\<rbrakk> \<Longrightarrow> Q(h \<circ> f)"
and Qeq: "\<And>h k. (\<And>x. x \<in> u \<Longrightarrow> h x = k x) \<Longrightarrow> Q h = Q k"
and hom: "\<And>f. \<lbrakk>continuous_on u f; f ` u \<subseteq> s; P f\<rbrakk>
\<Longrightarrow> \<exists>c. homotopic_with_canon P u s f (\<lambda>x. c)"
and contf: "continuous_on u f" and imf:"f ` u \<subseteq> t" and Qf: "Q f"
obtains c where "homotopic_with_canon Q u t f (\<lambda>x. c)"
proof -
have feq: "\<And>x. x \<in> u \<Longrightarrow> (h \<circ> (k \<circ> f)) x = f x" using idhk imf by auto
have "continuous_on u (k \<circ> f)"
using contf continuous_on_compose continuous_on_subset contk imf by blast
moreover have "(k \<circ> f) ` u \<subseteq> s"
using imf imk by fastforce
moreover have "P (k \<circ> f)"
by (simp add: P Qf contf imf)
ultimately obtain c where "homotopic_with_canon P u s (k \<circ> f) (\<lambda>x. c)"
by (metis hom)
then have "homotopic_with_canon Q u t (h \<circ> (k \<circ> f)) (h \<circ> (\<lambda>x. c))"
apply (rule homotopic_with_compose_continuous_left [OF homotopic_with_mono])
using Q by (auto simp: conth imh)
then show ?thesis
apply (rule_tac c = "h c" in that)
apply (erule homotopic_with_eq)
using feq apply auto[1]
apply simp
using Qeq topspace_euclidean_subtopology by blast
qed
lemma cohomotopically_trivial_retraction_gen:
assumes P: "\<And>f. \<lbrakk>continuous_on t f; f ` t \<subseteq> u; Q f\<rbrakk> \<Longrightarrow> P(f \<circ> h)"
and Q: "\<And>f. \<lbrakk>continuous_on s f; f ` s \<subseteq> u; P f\<rbrakk> \<Longrightarrow> Q(f \<circ> k)"
and Qeq: "\<And>h k. (\<And>x. x \<in> t \<Longrightarrow> h x = k x) \<Longrightarrow> Q h = Q k"
and hom: "\<And>f g. \<lbrakk>continuous_on s f; f ` s \<subseteq> u; P f;
continuous_on s g; g ` s \<subseteq> u; P g\<rbrakk>
\<Longrightarrow> homotopic_with_canon P s u f g"
and contf: "continuous_on t f" and imf: "f ` t \<subseteq> u" and Qf: "Q f"
and contg: "continuous_on t g" and img: "g ` t \<subseteq> u" and Qg: "Q g"
shows "homotopic_with_canon Q t u f g"
proof -
have feq: "\<And>x. x \<in> t \<Longrightarrow> (f \<circ> h \<circ> k) x = f x" using idhk imf by auto
have geq: "\<And>x. x \<in> t \<Longrightarrow> (g \<circ> h \<circ> k) x = g x" using idhk img by auto
have "continuous_on s (f \<circ> h)"
using contf conth continuous_on_compose imh by blast
moreover have "(f \<circ> h) ` s \<subseteq> u"
using imf imh by fastforce
moreover have "P (f \<circ> h)"
by (simp add: P Qf contf imf)
moreover have "continuous_on s (g \<circ> h)"
using contg continuous_on_compose continuous_on_subset conth imh by blast
moreover have "(g \<circ> h) ` s \<subseteq> u"
using img imh by fastforce
moreover have "P (g \<circ> h)"
by (simp add: P Qg contg img)
ultimately have "homotopic_with_canon P s u (f \<circ> h) (g \<circ> h)"
by (rule hom)
then have "homotopic_with_canon Q t u (f \<circ> h \<circ> k) (g \<circ> h \<circ> k)"
apply (rule homotopic_with_compose_continuous_right [OF homotopic_with_mono])
using Q by (auto simp: contk imk)
then show ?thesis
apply (rule homotopic_with_eq)
using feq apply auto[1]
using geq apply auto[1]
using Qeq topspace_euclidean_subtopology by blast
qed
lemma cohomotopically_trivial_retraction_null_gen:
assumes P: "\<And>f. \<lbrakk>continuous_on t f; f ` t \<subseteq> u; Q f\<rbrakk> \<Longrightarrow> P(f \<circ> h)"
and Q: "\<And>f. \<lbrakk>continuous_on s f; f ` s \<subseteq> u; P f\<rbrakk> \<Longrightarrow> Q(f \<circ> k)"
and Qeq: "\<And>h k. (\<And>x. x \<in> t \<Longrightarrow> h x = k x) \<Longrightarrow> Q h = Q k"
and hom: "\<And>f g. \<lbrakk>continuous_on s f; f ` s \<subseteq> u; P f\<rbrakk>
\<Longrightarrow> \<exists>c. homotopic_with_canon P s u f (\<lambda>x. c)"
and contf: "continuous_on t f" and imf: "f ` t \<subseteq> u" and Qf: "Q f"
obtains c where "homotopic_with_canon Q t u f (\<lambda>x. c)"
proof -
have feq: "\<And>x. x \<in> t \<Longrightarrow> (f \<circ> h \<circ> k) x = f x" using idhk imf by auto
have "continuous_on s (f \<circ> h)"
using contf conth continuous_on_compose imh by blast
moreover have "(f \<circ> h) ` s \<subseteq> u"
using imf imh by fastforce
moreover have "P (f \<circ> h)"
by (simp add: P Qf contf imf)
ultimately obtain c where "homotopic_with_canon P s u (f \<circ> h) (\<lambda>x. c)"
by (metis hom)
then have "homotopic_with_canon Q t u (f \<circ> h \<circ> k) ((\<lambda>x. c) \<circ> k)"
apply (rule homotopic_with_compose_continuous_right [OF homotopic_with_mono])
using Q by (auto simp: contk imk)
then show ?thesis
apply (rule_tac c = c in that)
apply (erule homotopic_with_eq)
using feq apply auto[1]
apply simp
using Qeq topspace_euclidean_subtopology by blast
qed
end
lemma simply_connected_retraction_gen:
shows "\<lbrakk>simply_connected S; continuous_on S h; h ` S = T;
continuous_on T k; k ` T \<subseteq> S; \<And>y. y \<in> T \<Longrightarrow> h(k y) = y\<rbrakk>
\<Longrightarrow> simply_connected T"
apply (simp add: simply_connected_def path_def path_image_def homotopic_loops_def, clarify)
apply (rule Retracts.homotopically_trivial_retraction_gen
[of S h _ k _ "\<lambda>p. pathfinish p = pathstart p" "\<lambda>p. pathfinish p = pathstart p"])
apply (simp_all add: Retracts_def pathfinish_def pathstart_def)
done
lemma homeomorphic_simply_connected:
"\<lbrakk>S homeomorphic T; simply_connected S\<rbrakk> \<Longrightarrow> simply_connected T"
by (auto simp: homeomorphic_def homeomorphism_def intro: simply_connected_retraction_gen)
lemma homeomorphic_simply_connected_eq:
"S homeomorphic T \<Longrightarrow> (simply_connected S \<longleftrightarrow> simply_connected T)"
by (metis homeomorphic_simply_connected homeomorphic_sym)
subsection\<open>Homotopy equivalence\<close>
subsection\<open>Homotopy equivalence of topological spaces.\<close>
definition\<^marker>\<open>tag important\<close> homotopy_equivalent_space
(infix "homotopy'_equivalent'_space" 50)
where "X homotopy_equivalent_space Y \<equiv>
(\<exists>f g. continuous_map X Y f \<and>
continuous_map Y X g \<and>
homotopic_with (\<lambda>x. True) X X (g \<circ> f) id \<and>
homotopic_with (\<lambda>x. True) Y Y (f \<circ> g) id)"
lemma homeomorphic_imp_homotopy_equivalent_space:
"X homeomorphic_space Y \<Longrightarrow> X homotopy_equivalent_space Y"
unfolding homeomorphic_space_def homotopy_equivalent_space_def
apply (erule ex_forward)+
by (simp add: homotopic_with_equal homotopic_with_sym homeomorphic_maps_def)
lemma homotopy_equivalent_space_refl:
"X homotopy_equivalent_space X"
by (simp add: homeomorphic_imp_homotopy_equivalent_space homeomorphic_space_refl)
lemma homotopy_equivalent_space_sym:
"X homotopy_equivalent_space Y \<longleftrightarrow> Y homotopy_equivalent_space X"
by (meson homotopy_equivalent_space_def)
lemma homotopy_eqv_trans [trans]:
assumes 1: "X homotopy_equivalent_space Y" and 2: "Y homotopy_equivalent_space U"
shows "X homotopy_equivalent_space U"
proof -
obtain f1 g1 where f1: "continuous_map X Y f1"
and g1: "continuous_map Y X g1"
and hom1: "homotopic_with (\<lambda>x. True) X X (g1 \<circ> f1) id"
"homotopic_with (\<lambda>x. True) Y Y (f1 \<circ> g1) id"
using 1 by (auto simp: homotopy_equivalent_space_def)
obtain f2 g2 where f2: "continuous_map Y U f2"
and g2: "continuous_map U Y g2"
and hom2: "homotopic_with (\<lambda>x. True) Y Y (g2 \<circ> f2) id"
"homotopic_with (\<lambda>x. True) U U (f2 \<circ> g2) id"
using 2 by (auto simp: homotopy_equivalent_space_def)
have "homotopic_with (\<lambda>f. True) X Y (g2 \<circ> f2 \<circ> f1) (id \<circ> f1)"
using f1 hom2(1) homotopic_compose_continuous_map_right by blast
then have "homotopic_with (\<lambda>f. True) X Y (g2 \<circ> (f2 \<circ> f1)) (id \<circ> f1)"
by (simp add: o_assoc)
then have "homotopic_with (\<lambda>x. True) X X
(g1 \<circ> (g2 \<circ> (f2 \<circ> f1))) (g1 \<circ> (id \<circ> f1))"
by (simp add: g1 homotopic_compose_continuous_map_left)
moreover have "homotopic_with (\<lambda>x. True) X X (g1 \<circ> id \<circ> f1) id"
using hom1 by simp
ultimately have SS: "homotopic_with (\<lambda>x. True) X X (g1 \<circ> g2 \<circ> (f2 \<circ> f1)) id"
apply (simp add: o_assoc)
apply (blast intro: homotopic_with_trans)
done
have "homotopic_with (\<lambda>f. True) U Y (f1 \<circ> g1 \<circ> g2) (id \<circ> g2)"
using g2 hom1(2) homotopic_with_compose_continuous_map_right by fastforce
then have "homotopic_with (\<lambda>f. True) U Y (f1 \<circ> (g1 \<circ> g2)) (id \<circ> g2)"
by (simp add: o_assoc)
then have "homotopic_with (\<lambda>x. True) U U
(f2 \<circ> (f1 \<circ> (g1 \<circ> g2))) (f2 \<circ> (id \<circ> g2))"
by (simp add: f2 homotopic_with_compose_continuous_map_left)
moreover have "homotopic_with (\<lambda>x. True) U U (f2 \<circ> id \<circ> g2) id"
using hom2 by simp
ultimately have UU: "homotopic_with (\<lambda>x. True) U U (f2 \<circ> f1 \<circ> (g1 \<circ> g2)) id"
apply (simp add: o_assoc)
apply (blast intro: homotopic_with_trans)
done
show ?thesis
unfolding homotopy_equivalent_space_def
by (blast intro: f1 f2 g1 g2 continuous_map_compose SS UU)
qed
lemma deformation_retraction_imp_homotopy_equivalent_space:
"\<lbrakk>homotopic_with (\<lambda>x. True) X X (s \<circ> r) id; retraction_maps X Y r s\<rbrakk>
\<Longrightarrow> X homotopy_equivalent_space Y"
unfolding homotopy_equivalent_space_def retraction_maps_def
apply (rule_tac x=r in exI)
apply (rule_tac x=s in exI)
apply (auto simp: homotopic_with_equal continuous_map_compose)
done
lemma deformation_retract_imp_homotopy_equivalent_space:
"\<lbrakk>homotopic_with (\<lambda>x. True) X X r id; retraction_maps X Y r id\<rbrakk>
\<Longrightarrow> X homotopy_equivalent_space Y"
using deformation_retraction_imp_homotopy_equivalent_space by force
lemma deformation_retract_of_space:
"S \<subseteq> topspace X \<and>
(\<exists>r. homotopic_with (\<lambda>x. True) X X id r \<and> retraction_maps X (subtopology X S) r id) \<longleftrightarrow>
S retract_of_space X \<and> (\<exists>f. homotopic_with (\<lambda>x. True) X X id f \<and> f ` (topspace X) \<subseteq> S)"
proof (cases "S \<subseteq> topspace X")
case True
moreover have "(\<exists>r. homotopic_with (\<lambda>x. True) X X id r \<and> retraction_maps X (subtopology X S) r id)
\<longleftrightarrow> (S retract_of_space X \<and> (\<exists>f. homotopic_with (\<lambda>x. True) X X id f \<and> f ` topspace X \<subseteq> S))"
unfolding retract_of_space_def
proof safe
fix f r
assume f: "homotopic_with (\<lambda>x. True) X X id f"
and fS: "f ` topspace X \<subseteq> S"
and r: "continuous_map X (subtopology X S) r"
and req: "\<forall>x\<in>S. r x = x"
show "\<exists>r. homotopic_with (\<lambda>x. True) X X id r \<and> retraction_maps X (subtopology X S) r id"
proof (intro exI conjI)
have "homotopic_with (\<lambda>x. True) X X f r"
proof (rule homotopic_with_eq)
show "homotopic_with (\<lambda>x. True) X X (r \<circ> f) (r \<circ> id)"
using homotopic_with_symD continuous_map_into_fulltopology f homotopic_compose_continuous_map_left r by blast
show "f x = (r \<circ> f) x" if "x \<in> topspace X" for x
using that fS req by auto
qed auto
then show "homotopic_with (\<lambda>x. True) X X id r"
by (rule homotopic_with_trans [OF f])
next
show "retraction_maps X (subtopology X S) r id"
by (simp add: r req retraction_maps_def)
qed
qed (use True in \<open>auto simp: retraction_maps_def topspace_subtopology_subset continuous_map_in_subtopology\<close>)
ultimately show ?thesis by simp
qed (auto simp: retract_of_space_def retraction_maps_def)
subsection\<open>Contractible spaces\<close>
text\<open>The definition (which agrees with "contractible" on subsets of Euclidean space)
is a little cryptic because we don't in fact assume that the constant "a" is in the space.
This forces the convention that the empty space / set is contractible, avoiding some special cases. \<close>
definition contractible_space where
"contractible_space X \<equiv> \<exists>a. homotopic_with (\<lambda>x. True) X X id (\<lambda>x. a)"
lemma contractible_space_top_of_set [simp]:"contractible_space (top_of_set S) \<longleftrightarrow> contractible S"
by (auto simp: contractible_space_def contractible_def)
lemma contractible_space_empty:
"topspace X = {} \<Longrightarrow> contractible_space X"
apply (simp add: contractible_space_def homotopic_with_def)
apply (rule_tac x=undefined in exI)
apply (rule_tac x="\<lambda>(t,x). if t = 0 then x else undefined" in exI)
apply (auto simp: continuous_map_on_empty)
done
lemma contractible_space_singleton:
"topspace X = {a} \<Longrightarrow> contractible_space X"
apply (simp add: contractible_space_def homotopic_with_def)
apply (rule_tac x=a in exI)
apply (rule_tac x="\<lambda>(t,x). if t = 0 then x else a" in exI)
apply (auto intro: continuous_map_eq [where f = "\<lambda>z. a"])
done
lemma contractible_space_subset_singleton:
"topspace X \<subseteq> {a} \<Longrightarrow> contractible_space X"
by (meson contractible_space_empty contractible_space_singleton subset_singletonD)
lemma contractible_space_subtopology_singleton:
"contractible_space(subtopology X {a})"
by (meson contractible_space_subset_singleton insert_subset path_connectedin_singleton path_connectedin_subtopology subsetI)
lemma contractible_space:
"contractible_space X \<longleftrightarrow>
topspace X = {} \<or>
(\<exists>a \<in> topspace X. homotopic_with (\<lambda>x. True) X X id (\<lambda>x. a))"
proof (cases "topspace X = {}")
case False
then show ?thesis
apply (auto simp: contractible_space_def)
using homotopic_with_imp_continuous_maps by fastforce
qed (simp add: contractible_space_empty)
lemma contractible_imp_path_connected_space:
assumes "contractible_space X" shows "path_connected_space X"
proof (cases "topspace X = {}")
case False
have *: "path_connected_space X"
if "a \<in> topspace X" and conth: "continuous_map (prod_topology (top_of_set {0..1}) X) X h"
and h: "\<forall>x. h (0, x) = x" "\<forall>x. h (1, x) = a"
for a and h :: "real \<times> 'a \<Rightarrow> 'a"
proof -
have "path_component_of X b a" if "b \<in> topspace X" for b
using that unfolding path_component_of_def
apply (rule_tac x="h \<circ> (\<lambda>x. (x,b))" in exI)
apply (simp add: h pathin_def)
apply (rule continuous_map_compose [OF _ conth])
apply (auto simp: continuous_map_pairwise intro!: continuous_intros continuous_map_compose continuous_map_id [unfolded id_def])
done
then show ?thesis
by (metis path_component_of_equiv path_connected_space_iff_path_component)
qed
show ?thesis
using assms False by (auto simp: contractible_space homotopic_with_def *)
qed (simp add: path_connected_space_topspace_empty)
lemma contractible_imp_connected_space:
"contractible_space X \<Longrightarrow> connected_space X"
by (simp add: contractible_imp_path_connected_space path_connected_imp_connected_space)
lemma contractible_space_alt:
"contractible_space X \<longleftrightarrow> (\<forall>a \<in> topspace X. homotopic_with (\<lambda>x. True) X X id (\<lambda>x. a))" (is "?lhs = ?rhs")
proof
assume X: ?lhs
then obtain a where a: "homotopic_with (\<lambda>x. True) X X id (\<lambda>x. a)"
by (auto simp: contractible_space_def)
show ?rhs
proof
show "homotopic_with (\<lambda>x. True) X X id (\<lambda>x. b)" if "b \<in> topspace X" for b
apply (rule homotopic_with_trans [OF a])
using homotopic_constant_maps path_connected_space_imp_path_component_of
by (metis (full_types) X a continuous_map_const contractible_imp_path_connected_space homotopic_with_imp_continuous_maps that)
qed
next
assume R: ?rhs
then show ?lhs
apply (simp add: contractible_space_def)
by (metis equals0I homotopic_on_emptyI)
qed
lemma compose_const [simp]: "f \<circ> (\<lambda>x. a) = (\<lambda>x. f a)" "(\<lambda>x. a) \<circ> g = (\<lambda>x. a)"
by (simp_all add: o_def)
lemma nullhomotopic_through_contractible_space:
assumes f: "continuous_map X Y f" and g: "continuous_map Y Z g" and Y: "contractible_space Y"
obtains c where "homotopic_with (\<lambda>h. True) X Z (g \<circ> f) (\<lambda>x. c)"
proof -
obtain b where b: "homotopic_with (\<lambda>x. True) Y Y id (\<lambda>x. b)"
using Y by (auto simp: contractible_space_def)
show thesis
using homotopic_compose_continuous_map_right
[OF homotopic_compose_continuous_map_left [OF b g] f]
by (simp add: that)
qed
lemma nullhomotopic_into_contractible_space:
assumes f: "continuous_map X Y f" and Y: "contractible_space Y"
obtains c where "homotopic_with (\<lambda>h. True) X Y f (\<lambda>x. c)"
using nullhomotopic_through_contractible_space [OF f _ Y]
by (metis continuous_map_id id_comp)
lemma nullhomotopic_from_contractible_space:
assumes f: "continuous_map X Y f" and X: "contractible_space X"
obtains c where "homotopic_with (\<lambda>h. True) X Y f (\<lambda>x. c)"
using nullhomotopic_through_contractible_space [OF _ f X]
by (metis comp_id continuous_map_id)
lemma homotopy_dominated_contractibility:
assumes f: "continuous_map X Y f" and g: "continuous_map Y X g"
and hom: "homotopic_with (\<lambda>x. True) Y Y (f \<circ> g) id" and X: "contractible_space X"
shows "contractible_space Y"
proof -
obtain c where c: "homotopic_with (\<lambda>h. True) X Y f (\<lambda>x. c)"
using nullhomotopic_from_contractible_space [OF f X] .
have "homotopic_with (\<lambda>x. True) Y Y (f \<circ> g) (\<lambda>x. c)"
using homotopic_compose_continuous_map_right [OF c g] by fastforce
then have "homotopic_with (\<lambda>x. True) Y Y id (\<lambda>x. c)"
using homotopic_with_trans [OF _ hom] homotopic_with_symD by blast
then show ?thesis
unfolding contractible_space_def ..
qed
lemma homotopy_equivalent_space_contractibility:
"X homotopy_equivalent_space Y \<Longrightarrow> (contractible_space X \<longleftrightarrow> contractible_space Y)"
unfolding homotopy_equivalent_space_def
by (blast intro: homotopy_dominated_contractibility)
lemma homeomorphic_space_contractibility:
"X homeomorphic_space Y
\<Longrightarrow> (contractible_space X \<longleftrightarrow> contractible_space Y)"
by (simp add: homeomorphic_imp_homotopy_equivalent_space homotopy_equivalent_space_contractibility)
lemma contractible_eq_homotopy_equivalent_singleton_subtopology:
"contractible_space X \<longleftrightarrow>
topspace X = {} \<or> (\<exists>a \<in> topspace X. X homotopy_equivalent_space (subtopology X {a}))"(is "?lhs = ?rhs")
proof (cases "topspace X = {}")
case False
show ?thesis
proof
assume ?lhs
then obtain a where a: "homotopic_with (\<lambda>x. True) X X id (\<lambda>x. a)"
by (auto simp: contractible_space_def)
then have "a \<in> topspace X"
by (metis False continuous_map_const homotopic_with_imp_continuous_maps)
then have "X homotopy_equivalent_space subtopology X {a}"
unfolding homotopy_equivalent_space_def
apply (rule_tac x="\<lambda>x. a" in exI)
apply (rule_tac x=id in exI)
apply (auto simp: homotopic_with_sym topspace_subtopology_subset a)
using connectedin_absolute connectedin_sing contractible_space_alt contractible_space_subtopology_singleton by fastforce
with \<open>a \<in> topspace X\<close> show ?rhs
by blast
next
assume ?rhs
then show ?lhs
by (meson False contractible_space_subtopology_singleton homotopy_equivalent_space_contractibility)
qed
qed (simp add: contractible_space_empty)
lemma contractible_space_retraction_map_image:
assumes "retraction_map X Y f" and X: "contractible_space X"
shows "contractible_space Y"
proof -
obtain g where f: "continuous_map X Y f" and g: "continuous_map Y X g" and fg: "\<forall>y \<in> topspace Y. f(g y) = y"
using assms by (auto simp: retraction_map_def retraction_maps_def)
obtain a where a: "homotopic_with (\<lambda>x. True) X X id (\<lambda>x. a)"
using X by (auto simp: contractible_space_def)
have "homotopic_with (\<lambda>x. True) Y Y id (\<lambda>x. f a)"
proof (rule homotopic_with_eq)
show "homotopic_with (\<lambda>x. True) Y Y (f \<circ> id \<circ> g) (f \<circ> (\<lambda>x. a) \<circ> g)"
using f g a homotopic_compose_continuous_map_left homotopic_compose_continuous_map_right by metis
qed (use fg in auto)
then show ?thesis
unfolding contractible_space_def by blast
qed
lemma contractible_space_prod_topology:
"contractible_space(prod_topology X Y) \<longleftrightarrow>
topspace X = {} \<or> topspace Y = {} \<or> contractible_space X \<and> contractible_space Y"
proof (cases "topspace X = {} \<or> topspace Y = {}")
case True
then have "topspace (prod_topology X Y) = {}"
by simp
then show ?thesis
by (auto simp: contractible_space_empty)
next
case False
have "contractible_space(prod_topology X Y) \<longleftrightarrow> contractible_space X \<and> contractible_space Y"
proof safe
assume XY: "contractible_space (prod_topology X Y)"
with False have "retraction_map (prod_topology X Y) X fst"
by (auto simp: contractible_space False retraction_map_fst)
then show "contractible_space X"
by (rule contractible_space_retraction_map_image [OF _ XY])
have "retraction_map (prod_topology X Y) Y snd"
using False XY by (auto simp: contractible_space False retraction_map_snd)
then show "contractible_space Y"
by (rule contractible_space_retraction_map_image [OF _ XY])
next
assume "contractible_space X" and "contractible_space Y"
with False obtain a b
where "a \<in> topspace X" and a: "homotopic_with (\<lambda>x. True) X X id (\<lambda>x. a)"
and "b \<in> topspace Y" and b: "homotopic_with (\<lambda>x. True) Y Y id (\<lambda>x. b)"
by (auto simp: contractible_space)
with False show "contractible_space (prod_topology X Y)"
apply (simp add: contractible_space)
apply (rule_tac x=a in bexI)
apply (rule_tac x=b in bexI)
using homotopic_with_prod_topology [OF a b]
apply (metis (no_types, lifting) case_prod_Pair case_prod_beta' eq_id_iff)
apply auto
done
qed
with False show ?thesis
by auto
qed
lemma contractible_space_product_topology:
"contractible_space(product_topology X I) \<longleftrightarrow>
topspace (product_topology X I) = {} \<or> (\<forall>i \<in> I. contractible_space(X i))"
proof (cases "topspace (product_topology X I) = {}")
case False
have 1: "contractible_space (X i)"
if XI: "contractible_space (product_topology X I)" and "i \<in> I"
for i
proof (rule contractible_space_retraction_map_image [OF _ XI])
show "retraction_map (product_topology X I) (X i) (\<lambda>x. x i)"
using False by (simp add: retraction_map_product_projection \<open>i \<in> I\<close>)
qed
have 2: "contractible_space (product_topology X I)"
if "x \<in> topspace (product_topology X I)" and cs: "\<forall>i\<in>I. contractible_space (X i)"
for x :: "'a \<Rightarrow> 'b"
proof -
obtain f where f: "\<And>i. i\<in>I \<Longrightarrow> homotopic_with (\<lambda>x. True) (X i) (X i) id (\<lambda>x. f i)"
using cs unfolding contractible_space_def by metis
have "homotopic_with (\<lambda>x. True)
(product_topology X I) (product_topology X I) id (\<lambda>x. restrict f I)"
by (rule homotopic_with_eq [OF homotopic_with_product_topology [OF f]]) (auto)
then show ?thesis
by (auto simp: contractible_space_def)
qed
show ?thesis
using False 1 2 by blast
qed (simp add: contractible_space_empty)
lemma contractible_space_subtopology_euclideanreal [simp]:
"contractible_space(subtopology euclideanreal S) \<longleftrightarrow> is_interval S"
(is "?lhs = ?rhs")
proof
assume ?lhs
then have "path_connectedin (subtopology euclideanreal S) S"
using contractible_imp_path_connected_space path_connectedin_topspace path_connectedin_absolute
by (simp add: contractible_imp_path_connected)
then show ?rhs
by (simp add: is_interval_path_connected_1)
next
assume ?rhs
then have "convex S"
by (simp add: is_interval_convex_1)
show ?lhs
proof (cases "S = {}")
case False
then obtain z where "z \<in> S"
by blast
show ?thesis
unfolding contractible_space_def homotopic_with_def
proof (intro exI conjI allI)
show "continuous_map (prod_topology (top_of_set {0..1}) (top_of_set S)) (top_of_set S)
(\<lambda>(t,x). (1 - t) * x + t * z)"
apply (auto simp: case_prod_unfold)
apply (intro continuous_intros)
using \<open>z \<in> S\<close> apply (auto intro: convexD [OF \<open>convex S\<close>, simplified])
done
qed auto
qed (simp add: contractible_space_empty)
qed
corollary contractible_space_euclideanreal: "contractible_space euclideanreal"
proof -
have "contractible_space (subtopology euclideanreal UNIV)"
using contractible_space_subtopology_euclideanreal by blast
then show ?thesis
by simp
qed
abbreviation\<^marker>\<open>tag important\<close> homotopy_eqv :: "'a::topological_space set \<Rightarrow> 'b::topological_space set \<Rightarrow> bool"
(infix "homotopy'_eqv" 50)
where "S homotopy_eqv T \<equiv> top_of_set S homotopy_equivalent_space top_of_set T"
lemma homeomorphic_imp_homotopy_eqv: "S homeomorphic T \<Longrightarrow> S homotopy_eqv T"
unfolding homeomorphic_def homeomorphism_def homotopy_equivalent_space_def
apply (erule ex_forward)+
apply auto
apply (metis homotopic_with_id2 topspace_euclidean_subtopology)
by (metis (full_types) homotopic_with_id2 imageE topspace_euclidean_subtopology)
lemma homotopy_eqv_inj_linear_image:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "linear f" "inj f"
shows "(f ` S) homotopy_eqv S"
apply (rule homeomorphic_imp_homotopy_eqv)
using assms homeomorphic_sym linear_homeomorphic_image by auto
lemma homotopy_eqv_translation:
fixes S :: "'a::real_normed_vector set"
shows "(+) a ` S homotopy_eqv S"
apply (rule homeomorphic_imp_homotopy_eqv)
using homeomorphic_translation homeomorphic_sym by blast
lemma homotopy_eqv_homotopic_triviality_imp:
fixes S :: "'a::real_normed_vector set"
and T :: "'b::real_normed_vector set"
and U :: "'c::real_normed_vector set"
assumes "S homotopy_eqv T"
and f: "continuous_on U f" "f ` U \<subseteq> T"
and g: "continuous_on U g" "g ` U \<subseteq> T"
and homUS: "\<And>f g. \<lbrakk>continuous_on U f; f ` U \<subseteq> S;
continuous_on U g; g ` U \<subseteq> S\<rbrakk>
\<Longrightarrow> homotopic_with_canon (\<lambda>x. True) U S f g"
shows "homotopic_with_canon (\<lambda>x. True) U T f g"
proof -
obtain h k where h: "continuous_on S h" "h ` S \<subseteq> T"
and k: "continuous_on T k" "k ` T \<subseteq> S"
and hom: "homotopic_with_canon (\<lambda>x. True) S S (k \<circ> h) id"
"homotopic_with_canon (\<lambda>x. True) T T (h \<circ> k) id"
using assms by (auto simp: homotopy_equivalent_space_def)
have "homotopic_with_canon (\<lambda>f. True) U S (k \<circ> f) (k \<circ> g)"
apply (rule homUS)
using f g k
apply (safe intro!: continuous_on_compose h k f elim!: continuous_on_subset)
apply (force simp: o_def)+
done
then have "homotopic_with_canon (\<lambda>x. True) U T (h \<circ> (k \<circ> f)) (h \<circ> (k \<circ> g))"
apply (rule homotopic_with_compose_continuous_left)
apply (simp_all add: h)
done
moreover have "homotopic_with_canon (\<lambda>x. True) U T (h \<circ> k \<circ> f) (id \<circ> f)"
apply (rule homotopic_with_compose_continuous_right [where X=T and Y=T])
apply (auto simp: hom f)
done
moreover have "homotopic_with_canon (\<lambda>x. True) U T (h \<circ> k \<circ> g) (id \<circ> g)"
apply (rule homotopic_with_compose_continuous_right [where X=T and Y=T])
apply (auto simp: hom g)
done
ultimately show "homotopic_with_canon (\<lambda>x. True) U T f g"
apply (simp add: o_assoc)
using homotopic_with_trans homotopic_with_sym by blast
qed
lemma homotopy_eqv_homotopic_triviality:
fixes S :: "'a::real_normed_vector set"
and T :: "'b::real_normed_vector set"
and U :: "'c::real_normed_vector set"
assumes "S homotopy_eqv T"
shows "(\<forall>f g. continuous_on U f \<and> f ` U \<subseteq> S \<and>
continuous_on U g \<and> g ` U \<subseteq> S
\<longrightarrow> homotopic_with_canon (\<lambda>x. True) U S f g) \<longleftrightarrow>
(\<forall>f g. continuous_on U f \<and> f ` U \<subseteq> T \<and>
continuous_on U g \<and> g ` U \<subseteq> T
\<longrightarrow> homotopic_with_canon (\<lambda>x. True) U T f g)"
(is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
by (metis assms homotopy_eqv_homotopic_triviality_imp)
next
assume ?rhs
moreover
have "T homotopy_eqv S"
using assms homotopy_equivalent_space_sym by blast
ultimately show ?lhs
by (blast intro: homotopy_eqv_homotopic_triviality_imp)
qed
lemma homotopy_eqv_cohomotopic_triviality_null_imp:
fixes S :: "'a::real_normed_vector set"
and T :: "'b::real_normed_vector set"
and U :: "'c::real_normed_vector set"
assumes "S homotopy_eqv T"
and f: "continuous_on T f" "f ` T \<subseteq> U"
and homSU: "\<And>f. \<lbrakk>continuous_on S f; f ` S \<subseteq> U\<rbrakk>
\<Longrightarrow> \<exists>c. homotopic_with_canon (\<lambda>x. True) S U f (\<lambda>x. c)"
obtains c where "homotopic_with_canon (\<lambda>x. True) T U f (\<lambda>x. c)"
proof -
obtain h k where h: "continuous_on S h" "h ` S \<subseteq> T"
and k: "continuous_on T k" "k ` T \<subseteq> S"
and hom: "homotopic_with_canon (\<lambda>x. True) S S (k \<circ> h) id"
"homotopic_with_canon (\<lambda>x. True) T T (h \<circ> k) id"
using assms by (auto simp: homotopy_equivalent_space_def)
obtain c where "homotopic_with_canon (\<lambda>x. True) S U (f \<circ> h) (\<lambda>x. c)"
apply (rule exE [OF homSU [of "f \<circ> h"]])
apply (intro continuous_on_compose h)
using h f apply (force elim!: continuous_on_subset)+
done
then have "homotopic_with_canon (\<lambda>x. True) T U ((f \<circ> h) \<circ> k) ((\<lambda>x. c) \<circ> k)"
apply (rule homotopic_with_compose_continuous_right [where X=S])
using k by auto
moreover have "homotopic_with_canon (\<lambda>x. True) T U (f \<circ> id) (f \<circ> (h \<circ> k))"
apply (rule homotopic_with_compose_continuous_left [where Y=T])
apply (simp add: hom homotopic_with_symD)
using f apply auto
done
ultimately show ?thesis
apply (rule_tac c=c in that)
apply (simp add: o_def)
using homotopic_with_trans by blast
qed
lemma homotopy_eqv_cohomotopic_triviality_null:
fixes S :: "'a::real_normed_vector set"
and T :: "'b::real_normed_vector set"
and U :: "'c::real_normed_vector set"
assumes "S homotopy_eqv T"
shows "(\<forall>f. continuous_on S f \<and> f ` S \<subseteq> U
\<longrightarrow> (\<exists>c. homotopic_with_canon (\<lambda>x. True) S U f (\<lambda>x. c))) \<longleftrightarrow>
(\<forall>f. continuous_on T f \<and> f ` T \<subseteq> U
\<longrightarrow> (\<exists>c. homotopic_with_canon (\<lambda>x. True) T U f (\<lambda>x. c)))"
apply (rule iffI)
apply (metis assms homotopy_eqv_cohomotopic_triviality_null_imp)
by (metis assms homotopy_eqv_cohomotopic_triviality_null_imp homotopy_equivalent_space_sym)
lemma homotopy_eqv_homotopic_triviality_null_imp:
fixes S :: "'a::real_normed_vector set"
and T :: "'b::real_normed_vector set"
and U :: "'c::real_normed_vector set"
assumes "S homotopy_eqv T"
and f: "continuous_on U f" "f ` U \<subseteq> T"
and homSU: "\<And>f. \<lbrakk>continuous_on U f; f ` U \<subseteq> S\<rbrakk>
\<Longrightarrow> \<exists>c. homotopic_with_canon (\<lambda>x. True) U S f (\<lambda>x. c)"
shows "\<exists>c. homotopic_with_canon (\<lambda>x. True) U T f (\<lambda>x. c)"
proof -
obtain h k where h: "continuous_on S h" "h ` S \<subseteq> T"
and k: "continuous_on T k" "k ` T \<subseteq> S"
and hom: "homotopic_with_canon (\<lambda>x. True) S S (k \<circ> h) id"
"homotopic_with_canon (\<lambda>x. True) T T (h \<circ> k) id"
using assms by (auto simp: homotopy_equivalent_space_def)
obtain c::'a where "homotopic_with_canon (\<lambda>x. True) U S (k \<circ> f) (\<lambda>x. c)"
apply (rule exE [OF homSU [of "k \<circ> f"]])
apply (intro continuous_on_compose h)
using k f apply (force elim!: continuous_on_subset)+
done
then have "homotopic_with_canon (\<lambda>x. True) U T (h \<circ> (k \<circ> f)) (h \<circ> (\<lambda>x. c))"
apply (rule homotopic_with_compose_continuous_left [where Y=S])
using h by auto
moreover have "homotopic_with_canon (\<lambda>x. True) U T (id \<circ> f) ((h \<circ> k) \<circ> f)"
apply (rule homotopic_with_compose_continuous_right [where X=T])
apply (simp add: hom homotopic_with_symD)
using f apply auto
done
ultimately show ?thesis
using homotopic_with_trans by (fastforce simp add: o_def)
qed
lemma homotopy_eqv_homotopic_triviality_null:
fixes S :: "'a::real_normed_vector set"
and T :: "'b::real_normed_vector set"
and U :: "'c::real_normed_vector set"
assumes "S homotopy_eqv T"
shows "(\<forall>f. continuous_on U f \<and> f ` U \<subseteq> S
\<longrightarrow> (\<exists>c. homotopic_with_canon (\<lambda>x. True) U S f (\<lambda>x. c))) \<longleftrightarrow>
(\<forall>f. continuous_on U f \<and> f ` U \<subseteq> T
\<longrightarrow> (\<exists>c. homotopic_with_canon (\<lambda>x. True) U T f (\<lambda>x. c)))"
apply (rule iffI)
apply (metis assms homotopy_eqv_homotopic_triviality_null_imp)
by (metis assms homotopy_eqv_homotopic_triviality_null_imp homotopy_equivalent_space_sym)
lemma homotopy_eqv_contractible_sets:
fixes S :: "'a::real_normed_vector set"
and T :: "'b::real_normed_vector set"
assumes "contractible S" "contractible T" "S = {} \<longleftrightarrow> T = {}"
shows "S homotopy_eqv T"
proof (cases "S = {}")
case True with assms show ?thesis
by (simp add: homeomorphic_imp_homotopy_eqv)
next
case False
with assms obtain a b where "a \<in> S" "b \<in> T"
by auto
then show ?thesis
unfolding homotopy_equivalent_space_def
apply (rule_tac x="\<lambda>x. b" in exI)
apply (rule_tac x="\<lambda>x. a" in exI)
apply (intro assms conjI continuous_on_id' homotopic_into_contractible)
apply (auto simp: o_def)
done
qed
lemma homotopy_eqv_empty1 [simp]:
fixes S :: "'a::real_normed_vector set"
shows "S homotopy_eqv ({}::'b::real_normed_vector set) \<longleftrightarrow> S = {}"
apply (rule iffI)
apply (metis Abstract_Topology.continuous_map_subtopology_eu emptyE equals0I homotopy_equivalent_space_def image_subset_iff)
by (simp add: homotopy_eqv_contractible_sets)
lemma homotopy_eqv_empty2 [simp]:
fixes S :: "'a::real_normed_vector set"
shows "({}::'b::real_normed_vector set) homotopy_eqv S \<longleftrightarrow> S = {}"
using homotopy_equivalent_space_sym homotopy_eqv_empty1 by blast
lemma homotopy_eqv_contractibility:
fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set"
shows "S homotopy_eqv T \<Longrightarrow> (contractible S \<longleftrightarrow> contractible T)"
by (meson contractible_space_top_of_set homotopy_equivalent_space_contractibility)
lemma homotopy_eqv_sing:
fixes S :: "'a::real_normed_vector set" and a :: "'b::real_normed_vector"
shows "S homotopy_eqv {a} \<longleftrightarrow> S \<noteq> {} \<and> contractible S"
proof (cases "S = {}")
case True then show ?thesis
by simp
next
case False then show ?thesis
by (metis contractible_sing empty_not_insert homotopy_eqv_contractibility homotopy_eqv_contractible_sets)
qed
lemma homeomorphic_contractible_eq:
fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set"
shows "S homeomorphic T \<Longrightarrow> (contractible S \<longleftrightarrow> contractible T)"
by (simp add: homeomorphic_imp_homotopy_eqv homotopy_eqv_contractibility)
lemma homeomorphic_contractible:
fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set"
shows "\<lbrakk>contractible S; S homeomorphic T\<rbrakk> \<Longrightarrow> contractible T"
by (metis homeomorphic_contractible_eq)
subsection\<^marker>\<open>tag unimportant\<close>\<open>Misc other results\<close>
lemma bounded_connected_Compl_real:
fixes S :: "real set"
assumes "bounded S" and conn: "connected(- S)"
shows "S = {}"
proof -
obtain a b where "S \<subseteq> box a b"
by (meson assms bounded_subset_box_symmetric)
then have "a \<notin> S" "b \<notin> S"
by auto
then have "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> x \<in> - S"
by (meson Compl_iff conn connected_iff_interval)
then show ?thesis
using \<open>S \<subseteq> box a b\<close> by auto
qed
corollary bounded_path_connected_Compl_real:
fixes S :: "real set"
assumes "bounded S" "path_connected(- S)" shows "S = {}"
by (simp add: assms bounded_connected_Compl_real path_connected_imp_connected)
lemma bounded_connected_Compl_1:
fixes S :: "'a::{euclidean_space} set"
assumes "bounded S" and conn: "connected(- S)" and 1: "DIM('a) = 1"
shows "S = {}"
proof -
have "DIM('a) = DIM(real)"
by (simp add: "1")
then obtain f::"'a \<Rightarrow> real" and g
where "linear f" "\<And>x. norm(f x) = norm x" "\<And>x. g(f x) = x" "\<And>y. f(g y) = y"
by (rule isomorphisms_UNIV_UNIV) blast
with \<open>bounded S\<close> have "bounded (f ` S)"
using bounded_linear_image linear_linear by blast
have "connected (f ` (-S))"
using connected_linear_image assms \<open>linear f\<close> by blast
moreover have "f ` (-S) = - (f ` S)"
apply (rule bij_image_Compl_eq)
apply (auto simp: bij_def)
apply (metis \<open>\<And>x. g (f x) = x\<close> injI)
by (metis UNIV_I \<open>\<And>y. f (g y) = y\<close> image_iff)
finally have "connected (- (f ` S))"
by simp
then have "f ` S = {}"
using \<open>bounded (f ` S)\<close> bounded_connected_Compl_real by blast
then show ?thesis
by blast
qed
subsection\<^marker>\<open>tag unimportant\<close>\<open>Some Uncountable Sets\<close>
lemma uncountable_closed_segment:
fixes a :: "'a::real_normed_vector"
assumes "a \<noteq> b" shows "uncountable (closed_segment a b)"
unfolding path_image_linepath [symmetric] path_image_def
using inj_on_linepath [OF assms] uncountable_closed_interval [of 0 1]
countable_image_inj_on by auto
lemma uncountable_open_segment:
fixes a :: "'a::real_normed_vector"
assumes "a \<noteq> b" shows "uncountable (open_segment a b)"
by (simp add: assms open_segment_def uncountable_closed_segment uncountable_minus_countable)
lemma uncountable_convex:
fixes a :: "'a::real_normed_vector"
assumes "convex S" "a \<in> S" "b \<in> S" "a \<noteq> b"
shows "uncountable S"
proof -
have "uncountable (closed_segment a b)"
by (simp add: uncountable_closed_segment assms)
then show ?thesis
by (meson assms convex_contains_segment countable_subset)
qed
lemma uncountable_ball:
fixes a :: "'a::euclidean_space"
assumes "r > 0"
shows "uncountable (ball a r)"
proof -
have "uncountable (open_segment a (a + r *\<^sub>R (SOME i. i \<in> Basis)))"
by (metis Basis_zero SOME_Basis add_cancel_right_right assms less_le scale_eq_0_iff uncountable_open_segment)
moreover have "open_segment a (a + r *\<^sub>R (SOME i. i \<in> Basis)) \<subseteq> ball a r"
using assms by (auto simp: in_segment algebra_simps dist_norm SOME_Basis)
ultimately show ?thesis
by (metis countable_subset)
qed
lemma ball_minus_countable_nonempty:
assumes "countable (A :: 'a :: euclidean_space set)" "r > 0"
shows "ball z r - A \<noteq> {}"
proof
assume *: "ball z r - A = {}"
have "uncountable (ball z r - A)"
by (intro uncountable_minus_countable assms uncountable_ball)
thus False by (subst (asm) *) auto
qed
lemma uncountable_cball:
fixes a :: "'a::euclidean_space"
assumes "r > 0"
shows "uncountable (cball a r)"
using assms countable_subset uncountable_ball by auto
lemma pairwise_disjnt_countable:
fixes \<N> :: "nat set set"
assumes "pairwise disjnt \<N>"
shows "countable \<N>"
proof -
have "inj_on (\<lambda>X. SOME n. n \<in> X) (\<N> - {{}})"
apply (clarsimp simp add: inj_on_def)
by (metis assms disjnt_insert2 insert_absorb pairwise_def subsetI subset_empty tfl_some)
then show ?thesis
by (metis countable_Diff_eq countable_def)
qed
lemma pairwise_disjnt_countable_Union:
assumes "countable (\<Union>\<N>)" and pwd: "pairwise disjnt \<N>"
shows "countable \<N>"
proof -
obtain f :: "_ \<Rightarrow> nat" where f: "inj_on f (\<Union>\<N>)"
using assms by blast
then have "pairwise disjnt (\<Union> X \<in> \<N>. {f ` X})"
using assms by (force simp: pairwise_def disjnt_inj_on_iff [OF f])
then have "countable (\<Union> X \<in> \<N>. {f ` X})"
using pairwise_disjnt_countable by blast
then show ?thesis
by (meson pwd countable_image_inj_on disjoint_image f inj_on_image pairwise_disjnt_countable)
qed
lemma connected_uncountable:
fixes S :: "'a::metric_space set"
assumes "connected S" "a \<in> S" "b \<in> S" "a \<noteq> b" shows "uncountable S"
proof -
have "continuous_on S (dist a)"
by (intro continuous_intros)
then have "connected (dist a ` S)"
by (metis connected_continuous_image \<open>connected S\<close>)
then have "closed_segment 0 (dist a b) \<subseteq> (dist a ` S)"
by (simp add: assms closed_segment_subset is_interval_connected_1 is_interval_convex)
then have "uncountable (dist a ` S)"
by (metis \<open>a \<noteq> b\<close> countable_subset dist_eq_0_iff uncountable_closed_segment)
then show ?thesis
by blast
qed
lemma path_connected_uncountable:
fixes S :: "'a::metric_space set"
assumes "path_connected S" "a \<in> S" "b \<in> S" "a \<noteq> b" shows "uncountable S"
using path_connected_imp_connected assms connected_uncountable by metis
lemma connected_finite_iff_sing:
fixes S :: "'a::metric_space set"
assumes "connected S"
shows "finite S \<longleftrightarrow> S = {} \<or> (\<exists>a. S = {a})" (is "_ = ?rhs")
proof -
have "uncountable S" if "\<not> ?rhs"
using connected_uncountable assms that by blast
then show ?thesis
using uncountable_infinite by auto
qed
lemma connected_card_eq_iff_nontrivial:
fixes S :: "'a::metric_space set"
shows "connected S \<Longrightarrow> uncountable S \<longleftrightarrow> \<not>(\<exists>a. S \<subseteq> {a})"
apply (auto simp: countable_finite finite_subset)
by (metis connected_uncountable is_singletonI' is_singleton_the_elem subset_singleton_iff)
lemma simple_path_image_uncountable:
fixes g :: "real \<Rightarrow> 'a::metric_space"
assumes "simple_path g"
shows "uncountable (path_image g)"
proof -
have "g 0 \<in> path_image g" "g (1/2) \<in> path_image g"
by (simp_all add: path_defs)
moreover have "g 0 \<noteq> g (1/2)"
using assms by (fastforce simp add: simple_path_def)
ultimately show ?thesis
apply (simp add: assms connected_card_eq_iff_nontrivial connected_simple_path_image)
by blast
qed
lemma arc_image_uncountable:
fixes g :: "real \<Rightarrow> 'a::metric_space"
assumes "arc g"
shows "uncountable (path_image g)"
by (simp add: arc_imp_simple_path assms simple_path_image_uncountable)
subsection\<^marker>\<open>tag unimportant\<close>\<open> Some simple positive connection theorems\<close>
proposition path_connected_convex_diff_countable:
fixes U :: "'a::euclidean_space set"
assumes "convex U" "\<not> collinear U" "countable S"
shows "path_connected(U - S)"
proof (clarsimp simp add: path_connected_def)
fix a b
assume "a \<in> U" "a \<notin> S" "b \<in> U" "b \<notin> S"
let ?m = "midpoint a b"
show "\<exists>g. path g \<and> path_image g \<subseteq> U - S \<and> pathstart g = a \<and> pathfinish g = b"
proof (cases "a = b")
case True
then show ?thesis
by (metis DiffI \<open>a \<in> U\<close> \<open>a \<notin> S\<close> path_component_def path_component_refl)
next
case False
then have "a \<noteq> ?m" "b \<noteq> ?m"
using midpoint_eq_endpoint by fastforce+
have "?m \<in> U"
using \<open>a \<in> U\<close> \<open>b \<in> U\<close> \<open>convex U\<close> convex_contains_segment by force
obtain c where "c \<in> U" and nc_abc: "\<not> collinear {a,b,c}"
by (metis False \<open>a \<in> U\<close> \<open>b \<in> U\<close> \<open>\<not> collinear U\<close> collinear_triples insert_absorb)
have ncoll_mca: "\<not> collinear {?m,c,a}"
by (metis (full_types) \<open>a \<noteq> ?m\<close> collinear_3_trans collinear_midpoint insert_commute nc_abc)
have ncoll_mcb: "\<not> collinear {?m,c,b}"
by (metis (full_types) \<open>b \<noteq> ?m\<close> collinear_3_trans collinear_midpoint insert_commute nc_abc)
have "c \<noteq> ?m"
by (metis collinear_midpoint insert_commute nc_abc)
then have "closed_segment ?m c \<subseteq> U"
by (simp add: \<open>c \<in> U\<close> \<open>?m \<in> U\<close> \<open>convex U\<close> closed_segment_subset)
then obtain z where z: "z \<in> closed_segment ?m c"
and disjS: "(closed_segment a z \<union> closed_segment z b) \<inter> S = {}"
proof -
have False if "closed_segment ?m c \<subseteq> {z. (closed_segment a z \<union> closed_segment z b) \<inter> S \<noteq> {}}"
proof -
have closb: "closed_segment ?m c \<subseteq>
{z \<in> closed_segment ?m c. closed_segment a z \<inter> S \<noteq> {}} \<union> {z \<in> closed_segment ?m c. closed_segment z b \<inter> S \<noteq> {}}"
using that by blast
have *: "countable {z \<in> closed_segment ?m c. closed_segment z u \<inter> S \<noteq> {}}"
if "u \<in> U" "u \<notin> S" and ncoll: "\<not> collinear {?m, c, u}" for u
proof -
have **: False if x1: "x1 \<in> closed_segment ?m c" and x2: "x2 \<in> closed_segment ?m c"
and "x1 \<noteq> x2" "x1 \<noteq> u"
and w: "w \<in> closed_segment x1 u" "w \<in> closed_segment x2 u"
and "w \<in> S" for x1 x2 w
proof -
have "x1 \<in> affine hull {?m,c}" "x2 \<in> affine hull {?m,c}"
using segment_as_ball x1 x2 by auto
then have coll_x1: "collinear {x1, ?m, c}" and coll_x2: "collinear {?m, c, x2}"
by (simp_all add: affine_hull_3_imp_collinear) (metis affine_hull_3_imp_collinear insert_commute)
have "\<not> collinear {x1, u, x2}"
proof
assume "collinear {x1, u, x2}"
then have "collinear {?m, c, u}"
by (metis (full_types) \<open>c \<noteq> ?m\<close> coll_x1 coll_x2 collinear_3_trans insert_commute ncoll \<open>x1 \<noteq> x2\<close>)
with ncoll show False ..
qed
then have "closed_segment x1 u \<inter> closed_segment u x2 = {u}"
by (blast intro!: Int_closed_segment)
then have "w = u"
using closed_segment_commute w by auto
show ?thesis
using \<open>u \<notin> S\<close> \<open>w = u\<close> that(7) by auto
qed
then have disj: "disjoint ((\<Union>z\<in>closed_segment ?m c. {closed_segment z u \<inter> S}))"
by (fastforce simp: pairwise_def disjnt_def)
have cou: "countable ((\<Union>z \<in> closed_segment ?m c. {closed_segment z u \<inter> S}) - {{}})"
apply (rule pairwise_disjnt_countable_Union [OF _ pairwise_subset [OF disj]])
apply (rule countable_subset [OF _ \<open>countable S\<close>], auto)
done
define f where "f \<equiv> \<lambda>X. (THE z. z \<in> closed_segment ?m c \<and> X = closed_segment z u \<inter> S)"
show ?thesis
proof (rule countable_subset [OF _ countable_image [OF cou, where f=f]], clarify)
fix x
assume x: "x \<in> closed_segment ?m c" "closed_segment x u \<inter> S \<noteq> {}"
show "x \<in> f ` ((\<Union>z\<in>closed_segment ?m c. {closed_segment z u \<inter> S}) - {{}})"
proof (rule_tac x="closed_segment x u \<inter> S" in image_eqI)
show "x = f (closed_segment x u \<inter> S)"
unfolding f_def
apply (rule the_equality [symmetric])
using x apply (auto simp: dest: **)
done
qed (use x in auto)
qed
qed
have "uncountable (closed_segment ?m c)"
by (metis \<open>c \<noteq> ?m\<close> uncountable_closed_segment)
then show False
using closb * [OF \<open>a \<in> U\<close> \<open>a \<notin> S\<close> ncoll_mca] * [OF \<open>b \<in> U\<close> \<open>b \<notin> S\<close> ncoll_mcb]
apply (simp add: closed_segment_commute)
by (simp add: countable_subset)
qed
then show ?thesis
by (force intro: that)
qed
show ?thesis
proof (intro exI conjI)
have "path_image (linepath a z +++ linepath z b) \<subseteq> U"
by (metis \<open>a \<in> U\<close> \<open>b \<in> U\<close> \<open>closed_segment ?m c \<subseteq> U\<close> z \<open>convex U\<close> closed_segment_subset contra_subsetD path_image_linepath subset_path_image_join)
with disjS show "path_image (linepath a z +++ linepath z b) \<subseteq> U - S"
by (force simp: path_image_join)
qed auto
qed
qed
corollary connected_convex_diff_countable:
fixes U :: "'a::euclidean_space set"
assumes "convex U" "\<not> collinear U" "countable S"
shows "connected(U - S)"
by (simp add: assms path_connected_convex_diff_countable path_connected_imp_connected)
lemma path_connected_punctured_convex:
assumes "convex S" and aff: "aff_dim S \<noteq> 1"
shows "path_connected(S - {a})"
proof -
consider "aff_dim S = -1" | "aff_dim S = 0" | "aff_dim S \<ge> 2"
using assms aff_dim_geq [of S] by linarith
then show ?thesis
proof cases
assume "aff_dim S = -1"
then show ?thesis
by (metis aff_dim_empty empty_Diff path_connected_empty)
next
assume "aff_dim S = 0"
then show ?thesis
by (metis aff_dim_eq_0 Diff_cancel Diff_empty Diff_insert0 convex_empty convex_imp_path_connected path_connected_singleton singletonD)
next
assume ge2: "aff_dim S \<ge> 2"
then have "\<not> collinear S"
proof (clarsimp simp add: collinear_affine_hull)
fix u v
assume "S \<subseteq> affine hull {u, v}"
then have "aff_dim S \<le> aff_dim {u, v}"
by (metis (no_types) aff_dim_affine_hull aff_dim_subset)
with ge2 show False
by (metis (no_types) aff_dim_2 antisym aff not_numeral_le_zero one_le_numeral order_trans)
qed
then show ?thesis
apply (rule path_connected_convex_diff_countable [OF \<open>convex S\<close>])
by simp
qed
qed
lemma connected_punctured_convex:
shows "\<lbrakk>convex S; aff_dim S \<noteq> 1\<rbrakk> \<Longrightarrow> connected(S - {a})"
using path_connected_imp_connected path_connected_punctured_convex by blast
lemma path_connected_complement_countable:
fixes S :: "'a::euclidean_space set"
assumes "2 \<le> DIM('a)" "countable S"
shows "path_connected(- S)"
proof -
have "path_connected(UNIV - S)"
apply (rule path_connected_convex_diff_countable)
using assms by (auto simp: collinear_aff_dim [of "UNIV :: 'a set"])
then show ?thesis
by (simp add: Compl_eq_Diff_UNIV)
qed
proposition path_connected_openin_diff_countable:
fixes S :: "'a::euclidean_space set"
assumes "connected S" and ope: "openin (top_of_set (affine hull S)) S"
and "\<not> collinear S" "countable T"
shows "path_connected(S - T)"
proof (clarsimp simp add: path_connected_component)
fix x y
assume xy: "x \<in> S" "x \<notin> T" "y \<in> S" "y \<notin> T"
show "path_component (S - T) x y"
proof (rule connected_equivalence_relation_gen [OF \<open>connected S\<close>, where P = "\<lambda>x. x \<notin> T"])
show "\<exists>z. z \<in> U \<and> z \<notin> T" if opeU: "openin (top_of_set S) U" and "x \<in> U" for U x
proof -
have "openin (top_of_set (affine hull S)) U"
using opeU ope openin_trans by blast
with \<open>x \<in> U\<close> obtain r where Usub: "U \<subseteq> affine hull S" and "r > 0"
and subU: "ball x r \<inter> affine hull S \<subseteq> U"
by (auto simp: openin_contains_ball)
with \<open>x \<in> U\<close> have x: "x \<in> ball x r \<inter> affine hull S"
by auto
have "\<not> S \<subseteq> {x}"
using \<open>\<not> collinear S\<close> collinear_subset by blast
then obtain x' where "x' \<noteq> x" "x' \<in> S"
by blast
obtain y where y: "y \<noteq> x" "y \<in> ball x r \<inter> affine hull S"
proof
show "x + (r / 2 / norm(x' - x)) *\<^sub>R (x' - x) \<noteq> x"
using \<open>x' \<noteq> x\<close> \<open>r > 0\<close> by auto
show "x + (r / 2 / norm (x' - x)) *\<^sub>R (x' - x) \<in> ball x r \<inter> affine hull S"
using \<open>x' \<noteq> x\<close> \<open>r > 0\<close> \<open>x' \<in> S\<close> x
by (simp add: dist_norm mem_affine_3_minus hull_inc)
qed
have "convex (ball x r \<inter> affine hull S)"
by (simp add: affine_imp_convex convex_Int)
with x y subU have "uncountable U"
by (meson countable_subset uncountable_convex)
then have "\<not> U \<subseteq> T"
using \<open>countable T\<close> countable_subset by blast
then show ?thesis by blast
qed
show "\<exists>U. openin (top_of_set S) U \<and> x \<in> U \<and>
(\<forall>x\<in>U. \<forall>y\<in>U. x \<notin> T \<and> y \<notin> T \<longrightarrow> path_component (S - T) x y)"
if "x \<in> S" for x
proof -
obtain r where Ssub: "S \<subseteq> affine hull S" and "r > 0"
and subS: "ball x r \<inter> affine hull S \<subseteq> S"
using ope \<open>x \<in> S\<close> by (auto simp: openin_contains_ball)
then have conv: "convex (ball x r \<inter> affine hull S)"
by (simp add: affine_imp_convex convex_Int)
have "\<not> aff_dim (affine hull S) \<le> 1"
using \<open>\<not> collinear S\<close> collinear_aff_dim by auto
then have "\<not> collinear (ball x r \<inter> affine hull S)"
apply (simp add: collinear_aff_dim)
by (metis (no_types, hide_lams) aff_dim_convex_Int_open IntI open_ball \<open>0 < r\<close> aff_dim_affine_hull affine_affine_hull affine_imp_convex centre_in_ball empty_iff hull_subset inf_commute subsetCE that)
then have *: "path_connected ((ball x r \<inter> affine hull S) - T)"
by (rule path_connected_convex_diff_countable [OF conv _ \<open>countable T\<close>])
have ST: "ball x r \<inter> affine hull S - T \<subseteq> S - T"
using subS by auto
show ?thesis
proof (intro exI conjI)
show "x \<in> ball x r \<inter> affine hull S"
using \<open>x \<in> S\<close> \<open>r > 0\<close> by (simp add: hull_inc)
have "openin (top_of_set (affine hull S)) (ball x r \<inter> affine hull S)"
by (subst inf.commute) (simp add: openin_Int_open)
then show "openin (top_of_set S) (ball x r \<inter> affine hull S)"
by (rule openin_subset_trans [OF _ subS Ssub])
qed (use * path_component_trans in \<open>auto simp: path_connected_component path_component_of_subset [OF ST]\<close>)
qed
qed (use xy path_component_trans in auto)
qed
corollary connected_openin_diff_countable:
fixes S :: "'a::euclidean_space set"
assumes "connected S" and ope: "openin (top_of_set (affine hull S)) S"
and "\<not> collinear S" "countable T"
shows "connected(S - T)"
by (metis path_connected_imp_connected path_connected_openin_diff_countable [OF assms])
corollary path_connected_open_diff_countable:
fixes S :: "'a::euclidean_space set"
assumes "2 \<le> DIM('a)" "open S" "connected S" "countable T"
shows "path_connected(S - T)"
proof (cases "S = {}")
case True
then show ?thesis
by (simp)
next
case False
show ?thesis
proof (rule path_connected_openin_diff_countable)
show "openin (top_of_set (affine hull S)) S"
by (simp add: assms hull_subset open_subset)
show "\<not> collinear S"
using assms False by (simp add: collinear_aff_dim aff_dim_open)
qed (simp_all add: assms)
qed
corollary connected_open_diff_countable:
fixes S :: "'a::euclidean_space set"
assumes "2 \<le> DIM('a)" "open S" "connected S" "countable T"
shows "connected(S - T)"
by (simp add: assms path_connected_imp_connected path_connected_open_diff_countable)
subsection\<^marker>\<open>tag unimportant\<close> \<open>Self-homeomorphisms shuffling points about\<close>
subsubsection\<^marker>\<open>tag unimportant\<close>\<open>The theorem \<open>homeomorphism_moving_points_exists\<close>\<close>
lemma homeomorphism_moving_point_1:
fixes a :: "'a::euclidean_space"
assumes "affine T" "a \<in> T" and u: "u \<in> ball a r \<inter> T"
obtains f g where "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f g"
"f a = u" "\<And>x. x \<in> sphere a r \<Longrightarrow> f x = x"
proof -
have nou: "norm (u - a) < r" and "u \<in> T"
using u by (auto simp: dist_norm norm_minus_commute)
then have "0 < r"
by (metis DiffD1 Diff_Diff_Int ball_eq_empty centre_in_ball not_le u)
define f where "f \<equiv> \<lambda>x. (1 - norm(x - a) / r) *\<^sub>R (u - a) + x"
have *: "False" if eq: "x + (norm y / r) *\<^sub>R u = y + (norm x / r) *\<^sub>R u"
and nou: "norm u < r" and yx: "norm y < norm x" for x y and u::'a
proof -
have "x = y + (norm x / r - (norm y / r)) *\<^sub>R u"
using eq by (simp add: algebra_simps)
then have "norm x = norm (y + ((norm x - norm y) / r) *\<^sub>R u)"
by (metis diff_divide_distrib)
also have "\<dots> \<le> norm y + norm(((norm x - norm y) / r) *\<^sub>R u)"
using norm_triangle_ineq by blast
also have "\<dots> = norm y + (norm x - norm y) * (norm u / r)"
using yx \<open>r > 0\<close>
by (simp add: field_split_simps)
also have "\<dots> < norm y + (norm x - norm y) * 1"
apply (subst add_less_cancel_left)
apply (rule mult_strict_left_mono)
using nou \<open>0 < r\<close> yx
apply (simp_all add: field_simps)
done
also have "\<dots> = norm x"
by simp
finally show False by simp
qed
have "inj f"
unfolding f_def
proof (clarsimp simp: inj_on_def)
fix x y
assume "(1 - norm (x - a) / r) *\<^sub>R (u - a) + x =
(1 - norm (y - a) / r) *\<^sub>R (u - a) + y"
then have eq: "(x - a) + (norm (y - a) / r) *\<^sub>R (u - a) = (y - a) + (norm (x - a) / r) *\<^sub>R (u - a)"
by (auto simp: algebra_simps)
show "x=y"
proof (cases "norm (x - a) = norm (y - a)")
case True
then show ?thesis
using eq by auto
next
case False
then consider "norm (x - a) < norm (y - a)" | "norm (x - a) > norm (y - a)"
by linarith
then have "False"
proof cases
case 1 show False
using * [OF _ nou 1] eq by simp
next
case 2 with * [OF eq nou] show False
by auto
qed
then show "x=y" ..
qed
qed
then have inj_onf: "inj_on f (cball a r \<inter> T)"
using inj_on_Int by fastforce
have contf: "continuous_on (cball a r \<inter> T) f"
unfolding f_def using \<open>0 < r\<close> by (intro continuous_intros) blast
have fim: "f ` (cball a r \<inter> T) = cball a r \<inter> T"
proof
have *: "norm (y + (1 - norm y / r) *\<^sub>R u) \<le> r" if "norm y \<le> r" "norm u < r" for y u::'a
proof -
have "norm (y + (1 - norm y / r) *\<^sub>R u) \<le> norm y + norm((1 - norm y / r) *\<^sub>R u)"
using norm_triangle_ineq by blast
also have "\<dots> = norm y + abs(1 - norm y / r) * norm u"
by simp
also have "\<dots> \<le> r"
proof -
have "(r - norm u) * (r - norm y) \<ge> 0"
using that by auto
then have "r * norm u + r * norm y \<le> r * r + norm u * norm y"
by (simp add: algebra_simps)
then show ?thesis
using that \<open>0 < r\<close> by (simp add: abs_if field_simps)
qed
finally show ?thesis .
qed
have "f ` (cball a r) \<subseteq> cball a r"
apply (clarsimp simp add: dist_norm norm_minus_commute f_def)
using * by (metis diff_add_eq diff_diff_add diff_diff_eq2 norm_minus_commute nou)
moreover have "f ` T \<subseteq> T"
unfolding f_def using \<open>affine T\<close> \<open>a \<in> T\<close> \<open>u \<in> T\<close>
by (force simp: add.commute mem_affine_3_minus)
ultimately show "f ` (cball a r \<inter> T) \<subseteq> cball a r \<inter> T"
by blast
next
show "cball a r \<inter> T \<subseteq> f ` (cball a r \<inter> T)"
proof (clarsimp simp add: dist_norm norm_minus_commute)
fix x
assume x: "norm (x - a) \<le> r" and "x \<in> T"
have "\<exists>v \<in> {0..1}. ((1 - v) * r - norm ((x - a) - v *\<^sub>R (u - a))) \<bullet> 1 = 0"
by (rule ivt_decreasing_component_on_1) (auto simp: x continuous_intros)
then obtain v where "0 \<le> v" "v \<le> 1"
and v: "(1 - v) * r = norm ((x - a) - v *\<^sub>R (u - a))"
by auto
then have n: "norm (a - (x - v *\<^sub>R (u - a))) = r - r * v"
by (simp add: field_simps norm_minus_commute)
show "x \<in> f ` (cball a r \<inter> T)"
proof (rule image_eqI)
show "x = f (x - v *\<^sub>R (u - a))"
using \<open>r > 0\<close> v by (simp add: f_def) (simp add: field_simps)
have "x - v *\<^sub>R (u - a) \<in> cball a r"
using \<open>r > 0\<close>\<open>0 \<le> v\<close>
by (simp add: dist_norm n)
moreover have "x - v *\<^sub>R (u - a) \<in> T"
by (simp add: f_def \<open>u \<in> T\<close> \<open>x \<in> T\<close> assms mem_affine_3_minus2)
ultimately show "x - v *\<^sub>R (u - a) \<in> cball a r \<inter> T"
by blast
qed
qed
qed
have "\<exists>g. homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f g"
apply (rule homeomorphism_compact [OF _ contf fim inj_onf])
apply (simp add: affine_closed compact_Int_closed \<open>affine T\<close>)
done
then show ?thesis
apply (rule exE)
apply (erule_tac f=f in that)
using \<open>r > 0\<close>
apply (simp_all add: f_def dist_norm norm_minus_commute)
done
qed
corollary\<^marker>\<open>tag unimportant\<close> homeomorphism_moving_point_2:
fixes a :: "'a::euclidean_space"
assumes "affine T" "a \<in> T" and u: "u \<in> ball a r \<inter> T" and v: "v \<in> ball a r \<inter> T"
obtains f g where "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f g"
"f u = v" "\<And>x. \<lbrakk>x \<in> sphere a r; x \<in> T\<rbrakk> \<Longrightarrow> f x = x"
proof -
have "0 < r"
by (metis DiffD1 Diff_Diff_Int ball_eq_empty centre_in_ball not_le u)
obtain f1 g1 where hom1: "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f1 g1"
and "f1 a = u" and f1: "\<And>x. x \<in> sphere a r \<Longrightarrow> f1 x = x"
using homeomorphism_moving_point_1 [OF \<open>affine T\<close> \<open>a \<in> T\<close> u] by blast
obtain f2 g2 where hom2: "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f2 g2"
and "f2 a = v" and f2: "\<And>x. x \<in> sphere a r \<Longrightarrow> f2 x = x"
using homeomorphism_moving_point_1 [OF \<open>affine T\<close> \<open>a \<in> T\<close> v] by blast
show ?thesis
proof
show "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) (f2 \<circ> g1) (f1 \<circ> g2)"
by (metis homeomorphism_compose homeomorphism_symD hom1 hom2)
have "g1 u = a"
using \<open>0 < r\<close> \<open>f1 a = u\<close> assms hom1 homeomorphism_apply1 by fastforce
then show "(f2 \<circ> g1) u = v"
by (simp add: \<open>f2 a = v\<close>)
show "\<And>x. \<lbrakk>x \<in> sphere a r; x \<in> T\<rbrakk> \<Longrightarrow> (f2 \<circ> g1) x = x"
using f1 f2 hom1 homeomorphism_apply1 by fastforce
qed
qed
corollary\<^marker>\<open>tag unimportant\<close> homeomorphism_moving_point_3:
fixes a :: "'a::euclidean_space"
assumes "affine T" "a \<in> T" and ST: "ball a r \<inter> T \<subseteq> S" "S \<subseteq> T"
and u: "u \<in> ball a r \<inter> T" and v: "v \<in> ball a r \<inter> T"
obtains f g where "homeomorphism S S f g"
"f u = v" "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> ball a r \<inter> T"
proof -
obtain f g where hom: "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f g"
and "f u = v" and fid: "\<And>x. \<lbrakk>x \<in> sphere a r; x \<in> T\<rbrakk> \<Longrightarrow> f x = x"
using homeomorphism_moving_point_2 [OF \<open>affine T\<close> \<open>a \<in> T\<close> u v] by blast
have gid: "\<And>x. \<lbrakk>x \<in> sphere a r; x \<in> T\<rbrakk> \<Longrightarrow> g x = x"
using fid hom homeomorphism_apply1 by fastforce
define ff where "ff \<equiv> \<lambda>x. if x \<in> ball a r \<inter> T then f x else x"
define gg where "gg \<equiv> \<lambda>x. if x \<in> ball a r \<inter> T then g x else x"
show ?thesis
proof
show "homeomorphism S S ff gg"
proof (rule homeomorphismI)
have "continuous_on ((cball a r \<inter> T) \<union> (T - ball a r)) ff"
apply (simp add: ff_def)
apply (rule continuous_on_cases)
using homeomorphism_cont1 [OF hom]
apply (auto simp: affine_closed \<open>affine T\<close> fid)
done
then show "continuous_on S ff"
apply (rule continuous_on_subset)
using ST by auto
have "continuous_on ((cball a r \<inter> T) \<union> (T - ball a r)) gg"
apply (simp add: gg_def)
apply (rule continuous_on_cases)
using homeomorphism_cont2 [OF hom]
apply (auto simp: affine_closed \<open>affine T\<close> gid)
done
then show "continuous_on S gg"
apply (rule continuous_on_subset)
using ST by auto
show "ff ` S \<subseteq> S"
proof (clarsimp simp add: ff_def)
fix x
assume "x \<in> S" and x: "dist a x < r" and "x \<in> T"
then have "f x \<in> cball a r \<inter> T"
using homeomorphism_image1 [OF hom] by force
then show "f x \<in> S"
using ST(1) \<open>x \<in> T\<close> gid hom homeomorphism_def x by fastforce
qed
show "gg ` S \<subseteq> S"
proof (clarsimp simp add: gg_def)
fix x
assume "x \<in> S" and x: "dist a x < r" and "x \<in> T"
then have "g x \<in> cball a r \<inter> T"
using homeomorphism_image2 [OF hom] by force
then have "g x \<in> ball a r"
using homeomorphism_apply2 [OF hom]
by (metis Diff_Diff_Int Diff_iff \<open>x \<in> T\<close> cball_def fid le_less mem_Collect_eq mem_ball mem_sphere x)
then show "g x \<in> S"
using ST(1) \<open>g x \<in> cball a r \<inter> T\<close> by force
qed
show "\<And>x. x \<in> S \<Longrightarrow> gg (ff x) = x"
apply (simp add: ff_def gg_def)
using homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom]
apply auto
apply (metis Int_iff homeomorphism_apply1 [OF hom] fid image_eqI less_eq_real_def mem_cball mem_sphere)
done
show "\<And>x. x \<in> S \<Longrightarrow> ff (gg x) = x"
apply (simp add: ff_def gg_def)
using homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom]
apply auto
apply (metis Int_iff fid image_eqI less_eq_real_def mem_cball mem_sphere)
done
qed
show "ff u = v"
using u by (auto simp: ff_def \<open>f u = v\<close>)
show "{x. \<not> (ff x = x \<and> gg x = x)} \<subseteq> ball a r \<inter> T"
by (auto simp: ff_def gg_def)
qed
qed
proposition\<^marker>\<open>tag unimportant\<close> homeomorphism_moving_point:
fixes a :: "'a::euclidean_space"
assumes ope: "openin (top_of_set (affine hull S)) S"
and "S \<subseteq> T"
and TS: "T \<subseteq> affine hull S"
and S: "connected S" "a \<in> S" "b \<in> S"
obtains f g where "homeomorphism T T f g" "f a = b"
"{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S"
"bounded {x. \<not> (f x = x \<and> g x = x)}"
proof -
have 1: "\<exists>h k. homeomorphism T T h k \<and> h (f d) = d \<and>
{x. \<not> (h x = x \<and> k x = x)} \<subseteq> S \<and> bounded {x. \<not> (h x = x \<and> k x = x)}"
if "d \<in> S" "f d \<in> S" and homfg: "homeomorphism T T f g"
and S: "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S"
and bo: "bounded {x. \<not> (f x = x \<and> g x = x)}" for d f g
proof (intro exI conjI)
show homgf: "homeomorphism T T g f"
by (metis homeomorphism_symD homfg)
then show "g (f d) = d"
by (meson \<open>S \<subseteq> T\<close> homeomorphism_def subsetD \<open>d \<in> S\<close>)
show "{x. \<not> (g x = x \<and> f x = x)} \<subseteq> S"
using S by blast
show "bounded {x. \<not> (g x = x \<and> f x = x)}"
using bo by (simp add: conj_commute)
qed
have 2: "\<exists>f g. homeomorphism T T f g \<and> f x = f2 (f1 x) \<and>
{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> g x = x)}"
if "x \<in> S" "f1 x \<in> S" "f2 (f1 x) \<in> S"
and hom: "homeomorphism T T f1 g1" "homeomorphism T T f2 g2"
and sub: "{x. \<not> (f1 x = x \<and> g1 x = x)} \<subseteq> S" "{x. \<not> (f2 x = x \<and> g2 x = x)} \<subseteq> S"
and bo: "bounded {x. \<not> (f1 x = x \<and> g1 x = x)}" "bounded {x. \<not> (f2 x = x \<and> g2 x = x)}"
for x f1 f2 g1 g2
proof (intro exI conjI)
show homgf: "homeomorphism T T (f2 \<circ> f1) (g1 \<circ> g2)"
by (metis homeomorphism_compose hom)
then show "(f2 \<circ> f1) x = f2 (f1 x)"
by force
show "{x. \<not> ((f2 \<circ> f1) x = x \<and> (g1 \<circ> g2) x = x)} \<subseteq> S"
using sub by force
have "bounded ({x. \<not>(f1 x = x \<and> g1 x = x)} \<union> {x. \<not>(f2 x = x \<and> g2 x = x)})"
using bo by simp
then show "bounded {x. \<not> ((f2 \<circ> f1) x = x \<and> (g1 \<circ> g2) x = x)}"
by (rule bounded_subset) auto
qed
have 3: "\<exists>U. openin (top_of_set S) U \<and>
d \<in> U \<and>
(\<forall>x\<in>U.
\<exists>f g. homeomorphism T T f g \<and> f d = x \<and>
{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and>
bounded {x. \<not> (f x = x \<and> g x = x)})"
if "d \<in> S" for d
proof -
obtain r where "r > 0" and r: "ball d r \<inter> affine hull S \<subseteq> S"
by (metis \<open>d \<in> S\<close> ope openin_contains_ball)
have *: "\<exists>f g. homeomorphism T T f g \<and> f d = e \<and>
{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and>
bounded {x. \<not> (f x = x \<and> g x = x)}" if "e \<in> S" "e \<in> ball d r" for e
apply (rule homeomorphism_moving_point_3 [of "affine hull S" d r T d e])
using r \<open>S \<subseteq> T\<close> TS that
apply (auto simp: \<open>d \<in> S\<close> \<open>0 < r\<close> hull_inc)
using bounded_subset by blast
show ?thesis
apply (rule_tac x="S \<inter> ball d r" in exI)
apply (intro conjI)
apply (simp add: openin_open_Int)
apply (simp add: \<open>0 < r\<close> that)
apply (blast intro: *)
done
qed
have "\<exists>f g. homeomorphism T T f g \<and> f a = b \<and>
{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> g x = x)}"
apply (rule connected_equivalence_relation [OF S], safe)
apply (blast intro: 1 2 3)+
done
then show ?thesis
using that by auto
qed
lemma homeomorphism_moving_points_exists_gen:
assumes K: "finite K" "\<And>i. i \<in> K \<Longrightarrow> x i \<in> S \<and> y i \<in> S"
"pairwise (\<lambda>i j. (x i \<noteq> x j) \<and> (y i \<noteq> y j)) K"
and "2 \<le> aff_dim S"
and ope: "openin (top_of_set (affine hull S)) S"
and "S \<subseteq> T" "T \<subseteq> affine hull S" "connected S"
shows "\<exists>f g. homeomorphism T T f g \<and> (\<forall>i \<in> K. f(x i) = y i) \<and>
{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> g x = x)}"
using assms
proof (induction K)
case empty
then show ?case
by (force simp: homeomorphism_ident)
next
case (insert i K)
then have xney: "\<And>j. \<lbrakk>j \<in> K; j \<noteq> i\<rbrakk> \<Longrightarrow> x i \<noteq> x j \<and> y i \<noteq> y j"
and pw: "pairwise (\<lambda>i j. x i \<noteq> x j \<and> y i \<noteq> y j) K"
and "x i \<in> S" "y i \<in> S"
and xyS: "\<And>i. i \<in> K \<Longrightarrow> x i \<in> S \<and> y i \<in> S"
by (simp_all add: pairwise_insert)
obtain f g where homfg: "homeomorphism T T f g" and feq: "\<And>i. i \<in> K \<Longrightarrow> f(x i) = y i"
and fg_sub: "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S"
and bo_fg: "bounded {x. \<not> (f x = x \<and> g x = x)}"
using insert.IH [OF xyS pw] insert.prems by (blast intro: that)
then have "\<exists>f g. homeomorphism T T f g \<and> (\<forall>i \<in> K. f(x i) = y i) \<and>
{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> g x = x)}"
using insert by blast
have aff_eq: "affine hull (S - y ` K) = affine hull S"
apply (rule affine_hull_Diff)
apply (auto simp: insert)
using \<open>y i \<in> S\<close> insert.hyps(2) xney xyS by fastforce
have f_in_S: "f x \<in> S" if "x \<in> S" for x
using homfg fg_sub homeomorphism_apply1 \<open>S \<subseteq> T\<close>
proof -
have "(f (f x) \<noteq> f x \<or> g (f x) \<noteq> f x) \<or> f x \<in> S"
by (metis \<open>S \<subseteq> T\<close> homfg subsetD homeomorphism_apply1 that)
then show ?thesis
using fg_sub by force
qed
obtain h k where homhk: "homeomorphism T T h k" and heq: "h (f (x i)) = y i"
and hk_sub: "{x. \<not> (h x = x \<and> k x = x)} \<subseteq> S - y ` K"
and bo_hk: "bounded {x. \<not> (h x = x \<and> k x = x)}"
proof (rule homeomorphism_moving_point [of "S - y`K" T "f(x i)" "y i"])
show "openin (top_of_set (affine hull (S - y ` K))) (S - y ` K)"
by (simp add: aff_eq openin_diff finite_imp_closedin image_subset_iff hull_inc insert xyS)
show "S - y ` K \<subseteq> T"
using \<open>S \<subseteq> T\<close> by auto
show "T \<subseteq> affine hull (S - y ` K)"
using insert by (simp add: aff_eq)
show "connected (S - y ` K)"
proof (rule connected_openin_diff_countable [OF \<open>connected S\<close> ope])
show "\<not> collinear S"
using collinear_aff_dim \<open>2 \<le> aff_dim S\<close> by force
show "countable (y ` K)"
using countable_finite insert.hyps(1) by blast
qed
show "f (x i) \<in> S - y ` K"
apply (auto simp: f_in_S \<open>x i \<in> S\<close>)
by (metis feq homfg \<open>x i \<in> S\<close> homeomorphism_def \<open>S \<subseteq> T\<close> \<open>i \<notin> K\<close> subsetCE xney xyS)
show "y i \<in> S - y ` K"
using insert.hyps xney by (auto simp: \<open>y i \<in> S\<close>)
qed blast
show ?case
proof (intro exI conjI)
show "homeomorphism T T (h \<circ> f) (g \<circ> k)"
using homfg homhk homeomorphism_compose by blast
show "\<forall>i \<in> insert i K. (h \<circ> f) (x i) = y i"
using feq hk_sub by (auto simp: heq)
show "{x. \<not> ((h \<circ> f) x = x \<and> (g \<circ> k) x = x)} \<subseteq> S"
using fg_sub hk_sub by force
have "bounded ({x. \<not>(f x = x \<and> g x = x)} \<union> {x. \<not>(h x = x \<and> k x = x)})"
using bo_fg bo_hk bounded_Un by blast
then show "bounded {x. \<not> ((h \<circ> f) x = x \<and> (g \<circ> k) x = x)}"
by (rule bounded_subset) auto
qed
qed
proposition\<^marker>\<open>tag unimportant\<close> homeomorphism_moving_points_exists:
fixes S :: "'a::euclidean_space set"
assumes 2: "2 \<le> DIM('a)" "open S" "connected S" "S \<subseteq> T" "finite K"
and KS: "\<And>i. i \<in> K \<Longrightarrow> x i \<in> S \<and> y i \<in> S"
and pw: "pairwise (\<lambda>i j. (x i \<noteq> x j) \<and> (y i \<noteq> y j)) K"
and S: "S \<subseteq> T" "T \<subseteq> affine hull S" "connected S"
obtains f g where "homeomorphism T T f g" "\<And>i. i \<in> K \<Longrightarrow> f(x i) = y i"
"{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S" "bounded {x. (\<not> (f x = x \<and> g x = x))}"
proof (cases "S = {}")
case True
then show ?thesis
using KS homeomorphism_ident that by fastforce
next
case False
then have affS: "affine hull S = UNIV"
by (simp add: affine_hull_open \<open>open S\<close>)
then have ope: "openin (top_of_set (affine hull S)) S"
using \<open>open S\<close> open_openin by auto
have "2 \<le> DIM('a)" by (rule 2)
also have "\<dots> = aff_dim (UNIV :: 'a set)"
by simp
also have "\<dots> \<le> aff_dim S"
by (metis aff_dim_UNIV aff_dim_affine_hull aff_dim_le_DIM affS)
finally have "2 \<le> aff_dim S"
by linarith
then show ?thesis
using homeomorphism_moving_points_exists_gen [OF \<open>finite K\<close> KS pw _ ope S] that by fastforce
qed
subsubsection\<^marker>\<open>tag unimportant\<close>\<open>The theorem \<open>homeomorphism_grouping_points_exists\<close>\<close>
lemma homeomorphism_grouping_point_1:
fixes a::real and c::real
assumes "a < b" "c < d"
obtains f g where "homeomorphism (cbox a b) (cbox c d) f g" "f a = c" "f b = d"
proof -
define f where "f \<equiv> \<lambda>x. ((d - c) / (b - a)) * x + (c - a * ((d - c) / (b - a)))"
have "\<exists>g. homeomorphism (cbox a b) (cbox c d) f g"
proof (rule homeomorphism_compact)
show "continuous_on (cbox a b) f"
apply (simp add: f_def)
apply (intro continuous_intros)
using assms by auto
have "f ` {a..b} = {c..d}"
unfolding f_def image_affinity_atLeastAtMost
using assms sum_sqs_eq by (auto simp: field_split_simps)
then show "f ` cbox a b = cbox c d"
by auto
show "inj_on f (cbox a b)"
unfolding f_def inj_on_def using assms by auto
qed auto
then obtain g where "homeomorphism (cbox a b) (cbox c d) f g" ..
then show ?thesis
proof
show "f a = c"
by (simp add: f_def)
show "f b = d"
using assms sum_sqs_eq [of a b] by (auto simp: f_def field_split_simps)
qed
qed
lemma homeomorphism_grouping_point_2:
fixes a::real and w::real
assumes hom_ab: "homeomorphism (cbox a b) (cbox u v) f1 g1"
and hom_bc: "homeomorphism (cbox b c) (cbox v w) f2 g2"
and "b \<in> cbox a c" "v \<in> cbox u w"
and eq: "f1 a = u" "f1 b = v" "f2 b = v" "f2 c = w"
obtains f g where "homeomorphism (cbox a c) (cbox u w) f g" "f a = u" "f c = w"
"\<And>x. x \<in> cbox a b \<Longrightarrow> f x = f1 x" "\<And>x. x \<in> cbox b c \<Longrightarrow> f x = f2 x"
proof -
have le: "a \<le> b" "b \<le> c" "u \<le> v" "v \<le> w"
using assms by simp_all
then have ac: "cbox a c = cbox a b \<union> cbox b c" and uw: "cbox u w = cbox u v \<union> cbox v w"
by auto
define f where "f \<equiv> \<lambda>x. if x \<le> b then f1 x else f2 x"
have "\<exists>g. homeomorphism (cbox a c) (cbox u w) f g"
proof (rule homeomorphism_compact)
have cf1: "continuous_on (cbox a b) f1"
using hom_ab homeomorphism_cont1 by blast
have cf2: "continuous_on (cbox b c) f2"
using hom_bc homeomorphism_cont1 by blast
show "continuous_on (cbox a c) f"
apply (simp add: f_def)
apply (rule continuous_on_cases_le [OF continuous_on_subset [OF cf1] continuous_on_subset [OF cf2]])
using le eq apply (force)+
done
have "f ` cbox a b = f1 ` cbox a b" "f ` cbox b c = f2 ` cbox b c"
unfolding f_def using eq by force+
then show "f ` cbox a c = cbox u w"
apply (simp only: ac uw image_Un)
by (metis hom_ab hom_bc homeomorphism_def)
have neq12: "f1 x \<noteq> f2 y" if x: "a \<le> x" "x \<le> b" and y: "b < y" "y \<le> c" for x y
proof -
have "f1 x \<in> cbox u v"
by (metis hom_ab homeomorphism_def image_eqI mem_box_real(2) x)
moreover have "f2 y \<in> cbox v w"
by (metis (full_types) hom_bc homeomorphism_def image_subset_iff mem_box_real(2) not_le not_less_iff_gr_or_eq order_refl y)
moreover have "f2 y \<noteq> f2 b"
by (metis cancel_comm_monoid_add_class.diff_cancel diff_gt_0_iff_gt hom_bc homeomorphism_def le(2) less_imp_le less_numeral_extra(3) mem_box_real(2) order_refl y)
ultimately show ?thesis
using le eq by simp
qed
have "inj_on f1 (cbox a b)"
by (metis (full_types) hom_ab homeomorphism_def inj_onI)
moreover have "inj_on f2 (cbox b c)"
by (metis (full_types) hom_bc homeomorphism_def inj_onI)
ultimately show "inj_on f (cbox a c)"
apply (simp (no_asm) add: inj_on_def)
apply (simp add: f_def inj_on_eq_iff)
using neq12 apply force
done
qed auto
then obtain g where "homeomorphism (cbox a c) (cbox u w) f g" ..
then show ?thesis
apply (rule that)
using eq le by (auto simp: f_def)
qed
lemma homeomorphism_grouping_point_3:
fixes a::real
assumes cbox_sub: "cbox c d \<subseteq> box a b" "cbox u v \<subseteq> box a b"
and box_ne: "box c d \<noteq> {}" "box u v \<noteq> {}"
obtains f g where "homeomorphism (cbox a b) (cbox a b) f g" "f a = a" "f b = b"
"\<And>x. x \<in> cbox c d \<Longrightarrow> f x \<in> cbox u v"
proof -
have less: "a < c" "a < u" "d < b" "v < b" "c < d" "u < v" "cbox c d \<noteq> {}"
using assms
by (simp_all add: cbox_sub subset_eq)
obtain f1 g1 where 1: "homeomorphism (cbox a c) (cbox a u) f1 g1"
and f1_eq: "f1 a = a" "f1 c = u"
using homeomorphism_grouping_point_1 [OF \<open>a < c\<close> \<open>a < u\<close>] .
obtain f2 g2 where 2: "homeomorphism (cbox c d) (cbox u v) f2 g2"
and f2_eq: "f2 c = u" "f2 d = v"
using homeomorphism_grouping_point_1 [OF \<open>c < d\<close> \<open>u < v\<close>] .
obtain f3 g3 where 3: "homeomorphism (cbox d b) (cbox v b) f3 g3"
and f3_eq: "f3 d = v" "f3 b = b"
using homeomorphism_grouping_point_1 [OF \<open>d < b\<close> \<open>v < b\<close>] .
obtain f4 g4 where 4: "homeomorphism (cbox a d) (cbox a v) f4 g4" and "f4 a = a" "f4 d = v"
and f4_eq: "\<And>x. x \<in> cbox a c \<Longrightarrow> f4 x = f1 x" "\<And>x. x \<in> cbox c d \<Longrightarrow> f4 x = f2 x"
using homeomorphism_grouping_point_2 [OF 1 2] less by (auto simp: f1_eq f2_eq)
obtain f g where fg: "homeomorphism (cbox a b) (cbox a b) f g" "f a = a" "f b = b"
and f_eq: "\<And>x. x \<in> cbox a d \<Longrightarrow> f x = f4 x" "\<And>x. x \<in> cbox d b \<Longrightarrow> f x = f3 x"
using homeomorphism_grouping_point_2 [OF 4 3] less by (auto simp: f4_eq f3_eq f2_eq f1_eq)
show ?thesis
apply (rule that [OF fg])
using f4_eq f_eq homeomorphism_image1 [OF 2]
apply simp
by (metis atLeastAtMost_iff box_real(1) box_real(2) cbox_sub(1) greaterThanLessThan_iff imageI less_eq_real_def subset_eq)
qed
lemma homeomorphism_grouping_point_4:
fixes T :: "real set"
assumes "open U" "open S" "connected S" "U \<noteq> {}" "finite K" "K \<subseteq> S" "U \<subseteq> S" "S \<subseteq> T"
obtains f g where "homeomorphism T T f g"
"\<And>x. x \<in> K \<Longrightarrow> f x \<in> U" "{x. (\<not> (f x = x \<and> g x = x))} \<subseteq> S"
"bounded {x. (\<not> (f x = x \<and> g x = x))}"
proof -
obtain c d where "box c d \<noteq> {}" "cbox c d \<subseteq> U"
proof -
obtain u where "u \<in> U"
using \<open>U \<noteq> {}\<close> by blast
then obtain e where "e > 0" "cball u e \<subseteq> U"
using \<open>open U\<close> open_contains_cball by blast
then show ?thesis
by (rule_tac c=u and d="u+e" in that) (auto simp: dist_norm subset_iff)
qed
have "compact K"
by (simp add: \<open>finite K\<close> finite_imp_compact)
obtain a b where "box a b \<noteq> {}" "K \<subseteq> cbox a b" "cbox a b \<subseteq> S"
proof (cases "K = {}")
case True then show ?thesis
using \<open>box c d \<noteq> {}\<close> \<open>cbox c d \<subseteq> U\<close> \<open>U \<subseteq> S\<close> that by blast
next
case False
then obtain a b where "a \<in> K" "b \<in> K"
and a: "\<And>x. x \<in> K \<Longrightarrow> a \<le> x" and b: "\<And>x. x \<in> K \<Longrightarrow> x \<le> b"
using compact_attains_inf compact_attains_sup by (metis \<open>compact K\<close>)+
obtain e where "e > 0" "cball b e \<subseteq> S"
using \<open>open S\<close> open_contains_cball
by (metis \<open>b \<in> K\<close> \<open>K \<subseteq> S\<close> subsetD)
show ?thesis
proof
show "box a (b + e) \<noteq> {}"
using \<open>0 < e\<close> \<open>b \<in> K\<close> a by force
show "K \<subseteq> cbox a (b + e)"
using \<open>0 < e\<close> a b by fastforce
have "a \<in> S"
using \<open>a \<in> K\<close> assms(6) by blast
have "b + e \<in> S"
using \<open>0 < e\<close> \<open>cball b e \<subseteq> S\<close> by (force simp: dist_norm)
show "cbox a (b + e) \<subseteq> S"
using \<open>a \<in> S\<close> \<open>b + e \<in> S\<close> \<open>connected S\<close> connected_contains_Icc by auto
qed
qed
obtain w z where "cbox w z \<subseteq> S" and sub_wz: "cbox a b \<union> cbox c d \<subseteq> box w z"
proof -
have "a \<in> S" "b \<in> S"
using \<open>box a b \<noteq> {}\<close> \<open>cbox a b \<subseteq> S\<close> by auto
moreover have "c \<in> S" "d \<in> S"
using \<open>box c d \<noteq> {}\<close> \<open>cbox c d \<subseteq> U\<close> \<open>U \<subseteq> S\<close> by force+
ultimately have "min a c \<in> S" "max b d \<in> S"
by linarith+
then obtain e1 e2 where "e1 > 0" "cball (min a c) e1 \<subseteq> S" "e2 > 0" "cball (max b d) e2 \<subseteq> S"
using \<open>open S\<close> open_contains_cball by metis
then have *: "min a c - e1 \<in> S" "max b d + e2 \<in> S"
by (auto simp: dist_norm)
show ?thesis
proof
show "cbox (min a c - e1) (max b d+ e2) \<subseteq> S"
using * \<open>connected S\<close> connected_contains_Icc by auto
show "cbox a b \<union> cbox c d \<subseteq> box (min a c - e1) (max b d + e2)"
using \<open>0 < e1\<close> \<open>0 < e2\<close> by auto
qed
qed
then
obtain f g where hom: "homeomorphism (cbox w z) (cbox w z) f g"
and "f w = w" "f z = z"
and fin: "\<And>x. x \<in> cbox a b \<Longrightarrow> f x \<in> cbox c d"
using homeomorphism_grouping_point_3 [of a b w z c d]
using \<open>box a b \<noteq> {}\<close> \<open>box c d \<noteq> {}\<close> by blast
have contfg: "continuous_on (cbox w z) f" "continuous_on (cbox w z) g"
using hom homeomorphism_def by blast+
define f' where "f' \<equiv> \<lambda>x. if x \<in> cbox w z then f x else x"
define g' where "g' \<equiv> \<lambda>x. if x \<in> cbox w z then g x else x"
show ?thesis
proof
have T: "cbox w z \<union> (T - box w z) = T"
using \<open>cbox w z \<subseteq> S\<close> \<open>S \<subseteq> T\<close> by auto
show "homeomorphism T T f' g'"
proof
have clo: "closedin (top_of_set (cbox w z \<union> (T - box w z))) (T - box w z)"
by (metis Diff_Diff_Int Diff_subset T closedin_def open_box openin_open_Int topspace_euclidean_subtopology)
have "continuous_on (cbox w z \<union> (T - box w z)) f'" "continuous_on (cbox w z \<union> (T - box w z)) g'"
unfolding f'_def g'_def
apply (safe intro!: continuous_on_cases_local contfg continuous_on_id clo)
apply (simp_all add: closed_subset)
using \<open>f w = w\<close> \<open>f z = z\<close> apply force
by (metis \<open>f w = w\<close> \<open>f z = z\<close> hom homeomorphism_def less_eq_real_def mem_box_real(2))
then show "continuous_on T f'" "continuous_on T g'"
by (simp_all only: T)
show "f' ` T \<subseteq> T"
unfolding f'_def
by clarsimp (metis \<open>cbox w z \<subseteq> S\<close> \<open>S \<subseteq> T\<close> subsetD hom homeomorphism_def imageI mem_box_real(2))
show "g' ` T \<subseteq> T"
unfolding g'_def
by clarsimp (metis \<open>cbox w z \<subseteq> S\<close> \<open>S \<subseteq> T\<close> subsetD hom homeomorphism_def imageI mem_box_real(2))
show "\<And>x. x \<in> T \<Longrightarrow> g' (f' x) = x"
unfolding f'_def g'_def
using homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] by fastforce
show "\<And>y. y \<in> T \<Longrightarrow> f' (g' y) = y"
unfolding f'_def g'_def
using homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] by fastforce
qed
show "\<And>x. x \<in> K \<Longrightarrow> f' x \<in> U"
using fin sub_wz \<open>K \<subseteq> cbox a b\<close> \<open>cbox c d \<subseteq> U\<close> by (force simp: f'_def)
show "{x. \<not> (f' x = x \<and> g' x = x)} \<subseteq> S"
using \<open>cbox w z \<subseteq> S\<close> by (auto simp: f'_def g'_def)
show "bounded {x. \<not> (f' x = x \<and> g' x = x)}"
apply (rule bounded_subset [of "cbox w z"])
using bounded_cbox apply blast
apply (auto simp: f'_def g'_def)
done
qed
qed
proposition\<^marker>\<open>tag unimportant\<close> homeomorphism_grouping_points_exists:
fixes S :: "'a::euclidean_space set"
assumes "open U" "open S" "connected S" "U \<noteq> {}" "finite K" "K \<subseteq> S" "U \<subseteq> S" "S \<subseteq> T"
obtains f g where "homeomorphism T T f g" "{x. (\<not> (f x = x \<and> g x = x))} \<subseteq> S"
"bounded {x. (\<not> (f x = x \<and> g x = x))}" "\<And>x. x \<in> K \<Longrightarrow> f x \<in> U"
proof (cases "2 \<le> DIM('a)")
case True
have TS: "T \<subseteq> affine hull S"
using affine_hull_open assms by blast
have "infinite U"
using \<open>open U\<close> \<open>U \<noteq> {}\<close> finite_imp_not_open by blast
then obtain P where "P \<subseteq> U" "finite P" "card K = card P"
using infinite_arbitrarily_large by metis
then obtain \<gamma> where \<gamma>: "bij_betw \<gamma> K P"
using \<open>finite K\<close> finite_same_card_bij by blast
obtain f g where "homeomorphism T T f g" "\<And>i. i \<in> K \<Longrightarrow> f (id i) = \<gamma> i" "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S" "bounded {x. \<not> (f x = x \<and> g x = x)}"
proof (rule homeomorphism_moving_points_exists [OF True \<open>open S\<close> \<open>connected S\<close> \<open>S \<subseteq> T\<close> \<open>finite K\<close>])
show "\<And>i. i \<in> K \<Longrightarrow> id i \<in> S \<and> \<gamma> i \<in> S"
using \<open>P \<subseteq> U\<close> \<open>bij_betw \<gamma> K P\<close> \<open>K \<subseteq> S\<close> \<open>U \<subseteq> S\<close> bij_betwE by blast
show "pairwise (\<lambda>i j. id i \<noteq> id j \<and> \<gamma> i \<noteq> \<gamma> j) K"
using \<gamma> by (auto simp: pairwise_def bij_betw_def inj_on_def)
qed (use affine_hull_open assms that in auto)
then show ?thesis
using \<gamma> \<open>P \<subseteq> U\<close> bij_betwE by (fastforce simp add: intro!: that)
next
case False
with DIM_positive have "DIM('a) = 1"
by (simp add: dual_order.antisym)
then obtain h::"'a \<Rightarrow>real" and j
where "linear h" "linear j"
and noh: "\<And>x. norm(h x) = norm x" and noj: "\<And>y. norm(j y) = norm y"
and hj: "\<And>x. j(h x) = x" "\<And>y. h(j y) = y"
and ranh: "surj h"
using isomorphisms_UNIV_UNIV
by (metis (mono_tags, hide_lams) DIM_real UNIV_eq_I range_eqI)
obtain f g where hom: "homeomorphism (h ` T) (h ` T) f g"
and f: "\<And>x. x \<in> h ` K \<Longrightarrow> f x \<in> h ` U"
and sub: "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> h ` S"
and bou: "bounded {x. \<not> (f x = x \<and> g x = x)}"
apply (rule homeomorphism_grouping_point_4 [of "h ` U" "h ` S" "h ` K" "h ` T"])
by (simp_all add: assms image_mono \<open>linear h\<close> open_surjective_linear_image connected_linear_image ranh)
have jf: "j (f (h x)) = x \<longleftrightarrow> f (h x) = h x" for x
by (metis hj)
have jg: "j (g (h x)) = x \<longleftrightarrow> g (h x) = h x" for x
by (metis hj)
have cont_hj: "continuous_on X h" "continuous_on Y j" for X Y
by (simp_all add: \<open>linear h\<close> \<open>linear j\<close> linear_linear linear_continuous_on)
show ?thesis
proof
show "homeomorphism T T (j \<circ> f \<circ> h) (j \<circ> g \<circ> h)"
proof
show "continuous_on T (j \<circ> f \<circ> h)" "continuous_on T (j \<circ> g \<circ> h)"
using hom homeomorphism_def
by (blast intro: continuous_on_compose cont_hj)+
show "(j \<circ> f \<circ> h) ` T \<subseteq> T" "(j \<circ> g \<circ> h) ` T \<subseteq> T"
by auto (metis (mono_tags, hide_lams) hj(1) hom homeomorphism_def imageE imageI)+
show "\<And>x. x \<in> T \<Longrightarrow> (j \<circ> g \<circ> h) ((j \<circ> f \<circ> h) x) = x"
using hj hom homeomorphism_apply1 by fastforce
show "\<And>y. y \<in> T \<Longrightarrow> (j \<circ> f \<circ> h) ((j \<circ> g \<circ> h) y) = y"
using hj hom homeomorphism_apply2 by fastforce
qed
show "{x. \<not> ((j \<circ> f \<circ> h) x = x \<and> (j \<circ> g \<circ> h) x = x)} \<subseteq> S"
apply (clarsimp simp: jf jg hj)
using sub hj
apply (drule_tac c="h x" in subsetD, force)
by (metis imageE)
have "bounded (j ` {x. (\<not> (f x = x \<and> g x = x))})"
by (rule bounded_linear_image [OF bou]) (use \<open>linear j\<close> linear_conv_bounded_linear in auto)
moreover
have *: "{x. \<not>((j \<circ> f \<circ> h) x = x \<and> (j \<circ> g \<circ> h) x = x)} = j ` {x. (\<not> (f x = x \<and> g x = x))}"
using hj by (auto simp: jf jg image_iff, metis+)
ultimately show "bounded {x. \<not> ((j \<circ> f \<circ> h) x = x \<and> (j \<circ> g \<circ> h) x = x)}"
by metis
show "\<And>x. x \<in> K \<Longrightarrow> (j \<circ> f \<circ> h) x \<in> U"
using f hj by fastforce
qed
qed
proposition\<^marker>\<open>tag unimportant\<close> homeomorphism_grouping_points_exists_gen:
fixes S :: "'a::euclidean_space set"
assumes opeU: "openin (top_of_set S) U"
and opeS: "openin (top_of_set (affine hull S)) S"
and "U \<noteq> {}" "finite K" "K \<subseteq> S" and S: "S \<subseteq> T" "T \<subseteq> affine hull S" "connected S"
obtains f g where "homeomorphism T T f g" "{x. (\<not> (f x = x \<and> g x = x))} \<subseteq> S"
"bounded {x. (\<not> (f x = x \<and> g x = x))}" "\<And>x. x \<in> K \<Longrightarrow> f x \<in> U"
proof (cases "2 \<le> aff_dim S")
case True
have opeU': "openin (top_of_set (affine hull S)) U"
using opeS opeU openin_trans by blast
obtain u where "u \<in> U" "u \<in> S"
using \<open>U \<noteq> {}\<close> opeU openin_imp_subset by fastforce+
have "infinite U"
apply (rule infinite_openin [OF opeU \<open>u \<in> U\<close>])
apply (rule connected_imp_perfect_aff_dim [OF \<open>connected S\<close> _ \<open>u \<in> S\<close>])
using True apply simp
done
then obtain P where "P \<subseteq> U" "finite P" "card K = card P"
using infinite_arbitrarily_large by metis
then obtain \<gamma> where \<gamma>: "bij_betw \<gamma> K P"
using \<open>finite K\<close> finite_same_card_bij by blast
have "\<exists>f g. homeomorphism T T f g \<and> (\<forall>i \<in> K. f(id i) = \<gamma> i) \<and>
{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> g x = x)}"
proof (rule homeomorphism_moving_points_exists_gen [OF \<open>finite K\<close> _ _ True opeS S])
show "\<And>i. i \<in> K \<Longrightarrow> id i \<in> S \<and> \<gamma> i \<in> S"
by (metis id_apply opeU openin_contains_cball subsetCE \<open>P \<subseteq> U\<close> \<open>bij_betw \<gamma> K P\<close> \<open>K \<subseteq> S\<close> bij_betwE)
show "pairwise (\<lambda>i j. id i \<noteq> id j \<and> \<gamma> i \<noteq> \<gamma> j) K"
using \<gamma> by (auto simp: pairwise_def bij_betw_def inj_on_def)
qed
then show ?thesis
using \<gamma> \<open>P \<subseteq> U\<close> bij_betwE by (fastforce simp add: intro!: that)
next
case False
with aff_dim_geq [of S] consider "aff_dim S = -1" | "aff_dim S = 0" | "aff_dim S = 1" by linarith
then show ?thesis
proof cases
assume "aff_dim S = -1"
then have "S = {}"
using aff_dim_empty by blast
then have "False"
using \<open>U \<noteq> {}\<close> \<open>K \<subseteq> S\<close> openin_imp_subset [OF opeU] by blast
then show ?thesis ..
next
assume "aff_dim S = 0"
then obtain a where "S = {a}"
using aff_dim_eq_0 by blast
then have "K \<subseteq> U"
using \<open>U \<noteq> {}\<close> \<open>K \<subseteq> S\<close> openin_imp_subset [OF opeU] by blast
show ?thesis
apply (rule that [of id id])
using \<open>K \<subseteq> U\<close> by (auto intro: homeomorphismI)
next
assume "aff_dim S = 1"
then have "affine hull S homeomorphic (UNIV :: real set)"
by (auto simp: homeomorphic_affine_sets)
then obtain h::"'a\<Rightarrow>real" and j where homhj: "homeomorphism (affine hull S) UNIV h j"
using homeomorphic_def by blast
then have h: "\<And>x. x \<in> affine hull S \<Longrightarrow> j(h(x)) = x" and j: "\<And>y. j y \<in> affine hull S \<and> h(j y) = y"
by (auto simp: homeomorphism_def)
have connh: "connected (h ` S)"
by (meson Topological_Spaces.connected_continuous_image \<open>connected S\<close> homeomorphism_cont1 homeomorphism_of_subsets homhj hull_subset top_greatest)
have hUS: "h ` U \<subseteq> h ` S"
by (meson homeomorphism_imp_open_map homeomorphism_of_subsets homhj hull_subset opeS opeU open_UNIV openin_open_eq)
have opn: "openin (top_of_set (affine hull S)) U \<Longrightarrow> open (h ` U)" for U
using homeomorphism_imp_open_map [OF homhj] by simp
have "open (h ` U)" "open (h ` S)"
by (auto intro: opeS opeU openin_trans opn)
then obtain f g where hom: "homeomorphism (h ` T) (h ` T) f g"
and f: "\<And>x. x \<in> h ` K \<Longrightarrow> f x \<in> h ` U"
and sub: "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> h ` S"
and bou: "bounded {x. \<not> (f x = x \<and> g x = x)}"
apply (rule homeomorphism_grouping_points_exists [of "h ` U" "h ` S" "h ` K" "h ` T"])
using assms by (auto simp: connh hUS)
have jf: "\<And>x. x \<in> affine hull S \<Longrightarrow> j (f (h x)) = x \<longleftrightarrow> f (h x) = h x"
by (metis h j)
have jg: "\<And>x. x \<in> affine hull S \<Longrightarrow> j (g (h x)) = x \<longleftrightarrow> g (h x) = h x"
by (metis h j)
have cont_hj: "continuous_on T h" "continuous_on Y j" for Y
apply (rule continuous_on_subset [OF _ \<open>T \<subseteq> affine hull S\<close>])
using homeomorphism_def homhj apply blast
by (meson continuous_on_subset homeomorphism_def homhj top_greatest)
define f' where "f' \<equiv> \<lambda>x. if x \<in> affine hull S then (j \<circ> f \<circ> h) x else x"
define g' where "g' \<equiv> \<lambda>x. if x \<in> affine hull S then (j \<circ> g \<circ> h) x else x"
show ?thesis
proof
show "homeomorphism T T f' g'"
proof
have "continuous_on T (j \<circ> f \<circ> h)"
apply (intro continuous_on_compose cont_hj)
using hom homeomorphism_def by blast
then show "continuous_on T f'"
apply (rule continuous_on_eq)
using \<open>T \<subseteq> affine hull S\<close> f'_def by auto
have "continuous_on T (j \<circ> g \<circ> h)"
apply (intro continuous_on_compose cont_hj)
using hom homeomorphism_def by blast
then show "continuous_on T g'"
apply (rule continuous_on_eq)
using \<open>T \<subseteq> affine hull S\<close> g'_def by auto
show "f' ` T \<subseteq> T"
proof (clarsimp simp: f'_def)
fix x assume "x \<in> T"
then have "f (h x) \<in> h ` T"
by (metis (no_types) hom homeomorphism_def image_subset_iff subset_refl)
then show "j (f (h x)) \<in> T"
using \<open>T \<subseteq> affine hull S\<close> h by auto
qed
show "g' ` T \<subseteq> T"
proof (clarsimp simp: g'_def)
fix x assume "x \<in> T"
then have "g (h x) \<in> h ` T"
by (metis (no_types) hom homeomorphism_def image_subset_iff subset_refl)
then show "j (g (h x)) \<in> T"
using \<open>T \<subseteq> affine hull S\<close> h by auto
qed
show "\<And>x. x \<in> T \<Longrightarrow> g' (f' x) = x"
using h j hom homeomorphism_apply1 by (fastforce simp add: f'_def g'_def)
show "\<And>y. y \<in> T \<Longrightarrow> f' (g' y) = y"
using h j hom homeomorphism_apply2 by (fastforce simp add: f'_def g'_def)
qed
next
show "{x. \<not> (f' x = x \<and> g' x = x)} \<subseteq> S"
apply (clarsimp simp: f'_def g'_def jf jg)
apply (rule imageE [OF subsetD [OF sub]], force)
by (metis h hull_inc)
next
have "compact (j ` closure {x. \<not> (f x = x \<and> g x = x)})"
using bou by (auto simp: compact_continuous_image cont_hj)
then have "bounded (j ` {x. \<not> (f x = x \<and> g x = x)})"
by (rule bounded_closure_image [OF compact_imp_bounded])
moreover
have *: "{x \<in> affine hull S. j (f (h x)) \<noteq> x \<or> j (g (h x)) \<noteq> x} = j ` {x. (\<not> (f x = x \<and> g x = x))}"
using h j by (auto simp: image_iff; metis)
ultimately have "bounded {x \<in> affine hull S. j (f (h x)) \<noteq> x \<or> j (g (h x)) \<noteq> x}"
by metis
then show "bounded {x. \<not> (f' x = x \<and> g' x = x)}"
by (simp add: f'_def g'_def Collect_mono bounded_subset)
next
show "f' x \<in> U" if "x \<in> K" for x
proof -
have "U \<subseteq> S"
using opeU openin_imp_subset by blast
then have "j (f (h x)) \<in> U"
using f h hull_subset that by fastforce
then show "f' x \<in> U"
using \<open>K \<subseteq> S\<close> S f'_def that by auto
qed
qed
qed
qed
subsection\<open>Nullhomotopic mappings\<close>
text\<open> A mapping out of a sphere is nullhomotopic iff it extends to the ball.
This even works out in the degenerate cases when the radius is \<open>\<le>\<close> 0, and
we also don't need to explicitly assume continuity since it's already implicit
in both sides of the equivalence.\<close>
lemma nullhomotopic_from_lemma:
assumes contg: "continuous_on (cball a r - {a}) g"
and fa: "\<And>e. 0 < e
\<Longrightarrow> \<exists>d. 0 < d \<and> (\<forall>x. x \<noteq> a \<and> norm(x - a) < d \<longrightarrow> norm(g x - f a) < e)"
and r: "\<And>x. x \<in> cball a r \<and> x \<noteq> a \<Longrightarrow> f x = g x"
shows "continuous_on (cball a r) f"
proof (clarsimp simp: continuous_on_eq_continuous_within Ball_def)
fix x
assume x: "dist a x \<le> r"
show "continuous (at x within cball a r) f"
proof (cases "x=a")
case True
then show ?thesis
by (metis continuous_within_eps_delta fa dist_norm dist_self r)
next
case False
show ?thesis
proof (rule continuous_transform_within [where f=g and d = "norm(x-a)"])
have "\<exists>d>0. \<forall>x'\<in>cball a r.
dist x' x < d \<longrightarrow> dist (g x') (g x) < e" if "e>0" for e
proof -
obtain d where "d > 0"
and d: "\<And>x'. \<lbrakk>dist x' a \<le> r; x' \<noteq> a; dist x' x < d\<rbrakk> \<Longrightarrow>
dist (g x') (g x) < e"
using contg False x \<open>e>0\<close>
unfolding continuous_on_iff by (fastforce simp add: dist_commute intro: that)
show ?thesis
using \<open>d > 0\<close> \<open>x \<noteq> a\<close>
by (rule_tac x="min d (norm(x - a))" in exI)
(auto simp: dist_commute dist_norm [symmetric] intro!: d)
qed
then show "continuous (at x within cball a r) g"
using contg False by (auto simp: continuous_within_eps_delta)
show "0 < norm (x - a)"
using False by force
show "x \<in> cball a r"
by (simp add: x)
show "\<And>x'. \<lbrakk>x' \<in> cball a r; dist x' x < norm (x - a)\<rbrakk>
\<Longrightarrow> g x' = f x'"
by (metis dist_commute dist_norm less_le r)
qed
qed
qed
proposition nullhomotopic_from_sphere_extension:
fixes f :: "'M::euclidean_space \<Rightarrow> 'a::real_normed_vector"
shows "(\<exists>c. homotopic_with_canon (\<lambda>x. True) (sphere a r) S f (\<lambda>x. c)) \<longleftrightarrow>
(\<exists>g. continuous_on (cball a r) g \<and> g ` (cball a r) \<subseteq> S \<and>
(\<forall>x \<in> sphere a r. g x = f x))"
(is "?lhs = ?rhs")
proof (cases r "0::real" rule: linorder_cases)
case less
then show ?thesis
by (simp add: homotopic_on_emptyI)
next
case equal
then show ?thesis
apply (auto simp: homotopic_with)
apply (rule_tac x="\<lambda>x. h (0, a)" in exI)
apply (fastforce simp add:)
using continuous_on_const by blast
next
case greater
let ?P = "continuous_on {x. norm(x - a) = r} f \<and> f ` {x. norm(x - a) = r} \<subseteq> S"
have ?P if ?lhs using that
proof
fix c
assume c: "homotopic_with_canon (\<lambda>x. True) (sphere a r) S f (\<lambda>x. c)"
then have contf: "continuous_on (sphere a r) f"
by (metis homotopic_with_imp_continuous)
moreover have fim: "f ` sphere a r \<subseteq> S"
by (meson continuous_map_subtopology_eu c homotopic_with_imp_continuous_maps)
show ?P
using contf fim by (auto simp: sphere_def dist_norm norm_minus_commute)
qed
moreover have ?P if ?rhs using that
proof
fix g
assume g: "continuous_on (cball a r) g \<and> g ` cball a r \<subseteq> S \<and> (\<forall>xa\<in>sphere a r. g xa = f xa)"
then
show ?P
apply (safe elim!: continuous_on_eq [OF continuous_on_subset])
apply (auto simp: dist_norm norm_minus_commute)
by (metis dist_norm image_subset_iff mem_sphere norm_minus_commute sphere_cball subsetCE)
qed
moreover have ?thesis if ?P
proof
assume ?lhs
then obtain c where "homotopic_with_canon (\<lambda>x. True) (sphere a r) S (\<lambda>x. c) f"
using homotopic_with_sym by blast
then obtain h where conth: "continuous_on ({0..1::real} \<times> sphere a r) h"
and him: "h ` ({0..1} \<times> sphere a r) \<subseteq> S"
and h: "\<And>x. h(0, x) = c" "\<And>x. h(1, x) = f x"
by (auto simp: homotopic_with_def)
obtain b1::'M where "b1 \<in> Basis"
using SOME_Basis by auto
have "c \<in> S"
apply (rule him [THEN subsetD])
apply (rule_tac x = "(0, a + r *\<^sub>R b1)" in image_eqI)
using h greater \<open>b1 \<in> Basis\<close>
apply (auto simp: dist_norm)
done
have uconth: "uniformly_continuous_on ({0..1::real} \<times> (sphere a r)) h"
by (force intro: compact_Times conth compact_uniformly_continuous)
let ?g = "\<lambda>x. h (norm (x - a)/r,
a + (if x = a then r *\<^sub>R b1 else (r / norm(x - a)) *\<^sub>R (x - a)))"
let ?g' = "\<lambda>x. h (norm (x - a)/r, a + (r / norm(x - a)) *\<^sub>R (x - a))"
show ?rhs
proof (intro exI conjI)
have "continuous_on (cball a r - {a}) ?g'"
apply (rule continuous_on_compose2 [OF conth])
apply (intro continuous_intros)
using greater apply (auto simp: dist_norm norm_minus_commute)
done
then show "continuous_on (cball a r) ?g"
proof (rule nullhomotopic_from_lemma)
show "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> norm (?g' x - ?g a) < e" if "0 < e" for e
proof -
obtain d where "0 < d"
and d: "\<And>x x'. \<lbrakk>x \<in> {0..1} \<times> sphere a r; x' \<in> {0..1} \<times> sphere a r; dist x' x < d\<rbrakk>
\<Longrightarrow> dist (h x') (h x) < e"
using uniformly_continuous_onE [OF uconth \<open>0 < e\<close>] by auto
have *: "norm (h (norm (x - a) / r,
a + (r / norm (x - a)) *\<^sub>R (x - a)) - h (0, a + r *\<^sub>R b1)) < e"
if "x \<noteq> a" "norm (x - a) < r" "norm (x - a) < d * r" for x
proof -
have "norm (h (norm (x - a) / r, a + (r / norm (x - a)) *\<^sub>R (x - a)) - h (0, a + r *\<^sub>R b1)) =
norm (h (norm (x - a) / r, a + (r / norm (x - a)) *\<^sub>R (x - a)) - h (0, a + (r / norm (x - a)) *\<^sub>R (x - a)))"
by (simp add: h)
also have "\<dots> < e"
apply (rule d [unfolded dist_norm])
using greater \<open>0 < d\<close> \<open>b1 \<in> Basis\<close> that
by (simp_all add: dist_norm) (simp add: field_simps)
finally show ?thesis .
qed
show ?thesis
apply (rule_tac x = "min r (d * r)" in exI)
using greater \<open>0 < d\<close> by (auto simp: *)
qed
show "\<And>x. x \<in> cball a r \<and> x \<noteq> a \<Longrightarrow> ?g x = ?g' x"
by auto
qed
next
show "?g ` cball a r \<subseteq> S"
using greater him \<open>c \<in> S\<close>
by (force simp: h dist_norm norm_minus_commute)
next
show "\<forall>x\<in>sphere a r. ?g x = f x"
using greater by (auto simp: h dist_norm norm_minus_commute)
qed
next
assume ?rhs
then obtain g where contg: "continuous_on (cball a r) g"
and gim: "g ` cball a r \<subseteq> S"
and gf: "\<forall>x \<in> sphere a r. g x = f x"
by auto
let ?h = "\<lambda>y. g (a + (fst y) *\<^sub>R (snd y - a))"
have "continuous_on ({0..1} \<times> sphere a r) ?h"
proof (rule continuous_on_compose2 [OF contg])
show "continuous_on ({0..1} \<times> sphere a r) (\<lambda>x. a + fst x *\<^sub>R (snd x - a))"
by (intro continuous_intros)
qed (auto simp: dist_norm norm_minus_commute mult_left_le_one_le)
moreover
have "?h ` ({0..1} \<times> sphere a r) \<subseteq> S"
by (auto simp: dist_norm norm_minus_commute mult_left_le_one_le gim [THEN subsetD])
moreover
have "\<forall>x\<in>sphere a r. ?h (0, x) = g a" "\<forall>x\<in>sphere a r. ?h (1, x) = f x"
by (auto simp: dist_norm norm_minus_commute mult_left_le_one_le gf)
ultimately have "homotopic_with_canon (\<lambda>x. True) (sphere a r) S (\<lambda>x. g a) f"
by (auto simp: homotopic_with)
then show ?lhs
using homotopic_with_symD by blast
qed
ultimately
show ?thesis by meson
qed
end
\ No newline at end of file
diff --git a/src/HOL/Analysis/Improper_Integral.thy b/src/HOL/Analysis/Improper_Integral.thy
--- a/src/HOL/Analysis/Improper_Integral.thy
+++ b/src/HOL/Analysis/Improper_Integral.thy
@@ -1,2268 +1,2268 @@
(* Title: HOL/Analysis/Improper_Integral.thy
Author: LC Paulson (ported from HOL Light)
*)
section \<open>Continuity of the indefinite integral; improper integral theorem\<close>
theory "Improper_Integral"
imports Equivalence_Lebesgue_Henstock_Integration
begin
subsection \<open>Equiintegrability\<close>
text\<open>The definition here only really makes sense for an elementary set.
We just use compact intervals in applications below.\<close>
definition\<^marker>\<open>tag important\<close> equiintegrable_on (infixr "equiintegrable'_on" 46)
where "F equiintegrable_on I \<equiv>
(\<forall>f \<in> F. f integrable_on I) \<and>
(\<forall>e > 0. \<exists>\<gamma>. gauge \<gamma> \<and>
(\<forall>f \<D>. f \<in> F \<and> \<D> tagged_division_of I \<and> \<gamma> fine \<D>
\<longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) < e))"
lemma equiintegrable_on_integrable:
"\<lbrakk>F equiintegrable_on I; f \<in> F\<rbrakk> \<Longrightarrow> f integrable_on I"
using equiintegrable_on_def by metis
lemma equiintegrable_on_sing [simp]:
"{f} equiintegrable_on cbox a b \<longleftrightarrow> f integrable_on cbox a b"
by (simp add: equiintegrable_on_def has_integral_integral has_integral integrable_on_def)
lemma equiintegrable_on_subset: "\<lbrakk>F equiintegrable_on I; G \<subseteq> F\<rbrakk> \<Longrightarrow> G equiintegrable_on I"
unfolding equiintegrable_on_def Ball_def
by (erule conj_forward imp_forward all_forward ex_forward | blast)+
lemma equiintegrable_on_Un:
assumes "F equiintegrable_on I" "G equiintegrable_on I"
shows "(F \<union> G) equiintegrable_on I"
unfolding equiintegrable_on_def
proof (intro conjI impI allI)
show "\<forall>f\<in>F \<union> G. f integrable_on I"
using assms unfolding equiintegrable_on_def by blast
show "\<exists>\<gamma>. gauge \<gamma> \<and>
(\<forall>f \<D>. f \<in> F \<union> G \<and>
\<D> tagged_division_of I \<and> \<gamma> fine \<D> \<longrightarrow>
norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon>)"
if "\<epsilon> > 0" for \<epsilon>
proof -
obtain \<gamma>1 where "gauge \<gamma>1"
and \<gamma>1: "\<And>f \<D>. f \<in> F \<and> \<D> tagged_division_of I \<and> \<gamma>1 fine \<D>
\<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon>"
using assms \<open>\<epsilon> > 0\<close> unfolding equiintegrable_on_def by auto
obtain \<gamma>2 where "gauge \<gamma>2"
and \<gamma>2: "\<And>f \<D>. f \<in> G \<and> \<D> tagged_division_of I \<and> \<gamma>2 fine \<D>
\<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon>"
using assms \<open>\<epsilon> > 0\<close> unfolding equiintegrable_on_def by auto
have "gauge (\<lambda>x. \<gamma>1 x \<inter> \<gamma>2 x)"
using \<open>gauge \<gamma>1\<close> \<open>gauge \<gamma>2\<close> by blast
moreover have "\<forall>f \<D>. f \<in> F \<union> G \<and> \<D> tagged_division_of I \<and> (\<lambda>x. \<gamma>1 x \<inter> \<gamma>2 x) fine \<D> \<longrightarrow>
norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon>"
using \<gamma>1 \<gamma>2 by (auto simp: fine_Int)
ultimately show ?thesis
by (intro exI conjI) assumption+
qed
qed
lemma equiintegrable_on_insert:
assumes "f integrable_on cbox a b" "F equiintegrable_on cbox a b"
shows "(insert f F) equiintegrable_on cbox a b"
by (metis assms equiintegrable_on_Un equiintegrable_on_sing insert_is_Un)
lemma equiintegrable_cmul:
assumes F: "F equiintegrable_on I"
shows "(\<Union>c \<in> {-k..k}. \<Union>f \<in> F. {(\<lambda>x. c *\<^sub>R f x)}) equiintegrable_on I"
unfolding equiintegrable_on_def
proof (intro conjI impI allI ballI)
show "f integrable_on I"
if "f \<in> (\<Union>c\<in>{- k..k}. \<Union>f\<in>F. {\<lambda>x. c *\<^sub>R f x})"
for f :: "'a \<Rightarrow> 'b"
using that assms equiintegrable_on_integrable integrable_cmul by blast
show "\<exists>\<gamma>. gauge \<gamma> \<and> (\<forall>f \<D>. f \<in> (\<Union>c\<in>{- k..k}. \<Union>f\<in>F. {\<lambda>x. c *\<^sub>R f x}) \<and> \<D> tagged_division_of I
\<and> \<gamma> fine \<D> \<longrightarrow> norm ((\<Sum>(x, K)\<in>\<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon>)"
if "\<epsilon> > 0" for \<epsilon>
proof -
obtain \<gamma> where "gauge \<gamma>"
and \<gamma>: "\<And>f \<D>. \<lbrakk>f \<in> F; \<D> tagged_division_of I; \<gamma> fine \<D>\<rbrakk>
\<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon> / (\<bar>k\<bar> + 1)"
using assms \<open>\<epsilon> > 0\<close> unfolding equiintegrable_on_def
by (metis add.commute add.right_neutral add_strict_mono divide_pos_pos norm_eq_zero real_norm_def zero_less_norm_iff zero_less_one)
moreover have "norm ((\<Sum>(x, K)\<in>\<D>. content K *\<^sub>R c *\<^sub>R (f x)) - integral I (\<lambda>x. c *\<^sub>R f x)) < \<epsilon>"
if c: "c \<in> {- k..k}"
and "f \<in> F" "\<D> tagged_division_of I" "\<gamma> fine \<D>"
for \<D> c f
proof -
have "norm ((\<Sum>x\<in>\<D>. case x of (x, K) \<Rightarrow> content K *\<^sub>R c *\<^sub>R f x) - integral I (\<lambda>x. c *\<^sub>R f x))
= \<bar>c\<bar> * norm ((\<Sum>x\<in>\<D>. case x of (x, K) \<Rightarrow> content K *\<^sub>R f x) - integral I f)"
by (simp add: algebra_simps scale_sum_right case_prod_unfold flip: norm_scaleR)
also have "\<dots> \<le> (\<bar>k\<bar> + 1) * norm ((\<Sum>x\<in>\<D>. case x of (x, K) \<Rightarrow> content K *\<^sub>R f x) - integral I f)"
using c by (auto simp: mult_right_mono)
also have "\<dots> < (\<bar>k\<bar> + 1) * (\<epsilon> / (\<bar>k\<bar> + 1))"
by (rule mult_strict_left_mono) (use \<gamma> less_eq_real_def that in auto)
also have "\<dots> = \<epsilon>"
by auto
finally show ?thesis .
qed
ultimately show ?thesis
by (rule_tac x="\<gamma>" in exI) auto
qed
qed
lemma equiintegrable_add:
assumes F: "F equiintegrable_on I" and G: "G equiintegrable_on I"
shows "(\<Union>f \<in> F. \<Union>g \<in> G. {(\<lambda>x. f x + g x)}) equiintegrable_on I"
unfolding equiintegrable_on_def
proof (intro conjI impI allI ballI)
show "f integrable_on I"
if "f \<in> (\<Union>f\<in>F. \<Union>g\<in>G. {\<lambda>x. f x + g x})" for f
using that equiintegrable_on_integrable assms by (auto intro: integrable_add)
show "\<exists>\<gamma>. gauge \<gamma> \<and> (\<forall>f \<D>. f \<in> (\<Union>f\<in>F. \<Union>g\<in>G. {\<lambda>x. f x + g x}) \<and> \<D> tagged_division_of I
\<and> \<gamma> fine \<D> \<longrightarrow> norm ((\<Sum>(x, K)\<in>\<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon>)"
if "\<epsilon> > 0" for \<epsilon>
proof -
obtain \<gamma>1 where "gauge \<gamma>1"
and \<gamma>1: "\<And>f \<D>. \<lbrakk>f \<in> F; \<D> tagged_division_of I; \<gamma>1 fine \<D>\<rbrakk>
\<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon>/2"
using assms \<open>\<epsilon> > 0\<close> unfolding equiintegrable_on_def by (meson half_gt_zero_iff)
obtain \<gamma>2 where "gauge \<gamma>2"
and \<gamma>2: "\<And>g \<D>. \<lbrakk>g \<in> G; \<D> tagged_division_of I; \<gamma>2 fine \<D>\<rbrakk>
\<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g x) - integral I g) < \<epsilon>/2"
using assms \<open>\<epsilon> > 0\<close> unfolding equiintegrable_on_def by (meson half_gt_zero_iff)
have "gauge (\<lambda>x. \<gamma>1 x \<inter> \<gamma>2 x)"
using \<open>gauge \<gamma>1\<close> \<open>gauge \<gamma>2\<close> by blast
moreover have "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R h x) - integral I h) < \<epsilon>"
if h: "h \<in> (\<Union>f\<in>F. \<Union>g\<in>G. {\<lambda>x. f x + g x})"
and \<D>: "\<D> tagged_division_of I" and fine: "(\<lambda>x. \<gamma>1 x \<inter> \<gamma>2 x) fine \<D>"
for h \<D>
proof -
obtain f g where "f \<in> F" "g \<in> G" and heq: "h = (\<lambda>x. f x + g x)"
using h by blast
then have int: "f integrable_on I" "g integrable_on I"
using F G equiintegrable_on_def by blast+
have "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R h x) - integral I h)
= norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x + content K *\<^sub>R g x) - (integral I f + integral I g))"
by (simp add: heq algebra_simps integral_add int)
also have "\<dots> = norm (((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f + (\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g x) - integral I g))"
by (simp add: sum.distrib algebra_simps case_prod_unfold)
also have "\<dots> \<le> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) + norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g x) - integral I g)"
by (metis (mono_tags) add_diff_eq norm_triangle_ineq)
also have "\<dots> < \<epsilon>/2 + \<epsilon>/2"
using \<gamma>1 [OF \<open>f \<in> F\<close> \<D>] \<gamma>2 [OF \<open>g \<in> G\<close> \<D>] fine by (simp add: fine_Int)
finally show ?thesis by simp
qed
ultimately show ?thesis
by meson
qed
qed
lemma equiintegrable_minus:
assumes "F equiintegrable_on I"
shows "(\<Union>f \<in> F. {(\<lambda>x. - f x)}) equiintegrable_on I"
by (force intro: equiintegrable_on_subset [OF equiintegrable_cmul [OF assms, of 1]])
lemma equiintegrable_diff:
assumes F: "F equiintegrable_on I" and G: "G equiintegrable_on I"
shows "(\<Union>f \<in> F. \<Union>g \<in> G. {(\<lambda>x. f x - g x)}) equiintegrable_on I"
by (rule equiintegrable_on_subset [OF equiintegrable_add [OF F equiintegrable_minus [OF G]]]) auto
lemma equiintegrable_sum:
fixes F :: "('a::euclidean_space \<Rightarrow> 'b::euclidean_space) set"
assumes "F equiintegrable_on cbox a b"
shows "(\<Union>I \<in> Collect finite. \<Union>c \<in> {c. (\<forall>i \<in> I. c i \<ge> 0) \<and> sum c I = 1}.
\<Union>f \<in> I \<rightarrow> F. {(\<lambda>x. sum (\<lambda>i::'j. c i *\<^sub>R f i x) I)}) equiintegrable_on cbox a b"
(is "?G equiintegrable_on _")
unfolding equiintegrable_on_def
proof (intro conjI impI allI ballI)
show "f integrable_on cbox a b" if "f \<in> ?G" for f
using that assms by (auto simp: equiintegrable_on_def intro!: integrable_sum integrable_cmul)
show "\<exists>\<gamma>. gauge \<gamma>
\<and> (\<forall>g \<D>. g \<in> ?G \<and> \<D> tagged_division_of cbox a b \<and> \<gamma> fine \<D>
\<longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g x) - integral (cbox a b) g) < \<epsilon>)"
if "\<epsilon> > 0" for \<epsilon>
proof -
obtain \<gamma> where "gauge \<gamma>"
and \<gamma>: "\<And>f \<D>. \<lbrakk>f \<in> F; \<D> tagged_division_of cbox a b; \<gamma> fine \<D>\<rbrakk>
\<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral (cbox a b) f) < \<epsilon> / 2"
using assms \<open>\<epsilon> > 0\<close> unfolding equiintegrable_on_def by (meson half_gt_zero_iff)
moreover have "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g x) - integral (cbox a b) g) < \<epsilon>"
if g: "g \<in> ?G"
and \<D>: "\<D> tagged_division_of cbox a b"
and fine: "\<gamma> fine \<D>"
for \<D> g
proof -
obtain I c f where "finite I" and 0: "\<And>i::'j. i \<in> I \<Longrightarrow> 0 \<le> c i"
and 1: "sum c I = 1" and f: "f \<in> I \<rightarrow> F" and geq: "g = (\<lambda>x. \<Sum>i\<in>I. c i *\<^sub>R f i x)"
using g by auto
have fi_int: "f i integrable_on cbox a b" if "i \<in> I" for i
by (metis Pi_iff assms equiintegrable_on_def f that)
have *: "integral (cbox a b) (\<lambda>x. c i *\<^sub>R f i x) = (\<Sum>(x, K)\<in>\<D>. integral K (\<lambda>x. c i *\<^sub>R f i x))"
if "i \<in> I" for i
proof -
have "f i integrable_on cbox a b"
by (metis Pi_iff assms equiintegrable_on_def f that)
then show ?thesis
by (intro \<D> integrable_cmul integral_combine_tagged_division_topdown)
qed
have "finite \<D>"
using \<D> by blast
have swap: "(\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R (\<Sum>i\<in>I. c i *\<^sub>R f i x))
= (\<Sum>i\<in>I. c i *\<^sub>R (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f i x))"
by (simp add: scale_sum_right case_prod_unfold algebra_simps) (rule sum.swap)
have "norm ((\<Sum>(x, K)\<in>\<D>. content K *\<^sub>R g x) - integral (cbox a b) g)
= norm ((\<Sum>i\<in>I. c i *\<^sub>R ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f i x) - integral (cbox a b) (f i))))"
unfolding geq swap
by (simp add: scaleR_right.sum algebra_simps integral_sum fi_int integrable_cmul \<open>finite I\<close> sum_subtractf flip: sum_diff)
also have "\<dots> \<le> (\<Sum>i\<in>I. c i * \<epsilon> / 2)"
proof (rule sum_norm_le)
show "norm (c i *\<^sub>R ((\<Sum>(xa, K)\<in>\<D>. content K *\<^sub>R f i xa) - integral (cbox a b) (f i))) \<le> c i * \<epsilon> / 2"
if "i \<in> I" for i
proof -
have "norm ((\<Sum>(x, K)\<in>\<D>. content K *\<^sub>R f i x) - integral (cbox a b) (f i)) \<le> \<epsilon>/2"
using \<gamma> [OF _ \<D> fine, of "f i"] funcset_mem [OF f] that by auto
then show ?thesis
using that by (auto simp: 0 mult.assoc intro: mult_left_mono)
qed
qed
also have "\<dots> < \<epsilon>"
using 1 \<open>\<epsilon> > 0\<close> by (simp add: flip: sum_divide_distrib sum_distrib_right)
finally show ?thesis .
qed
ultimately show ?thesis
by (rule_tac x="\<gamma>" in exI) auto
qed
qed
corollary equiintegrable_sum_real:
fixes F :: "(real \<Rightarrow> 'b::euclidean_space) set"
assumes "F equiintegrable_on {a..b}"
shows "(\<Union>I \<in> Collect finite. \<Union>c \<in> {c. (\<forall>i \<in> I. c i \<ge> 0) \<and> sum c I = 1}.
\<Union>f \<in> I \<rightarrow> F. {(\<lambda>x. sum (\<lambda>i. c i *\<^sub>R f i x) I)})
equiintegrable_on {a..b}"
using equiintegrable_sum [of F a b] assms by auto
text\<open> Basic combining theorems for the interval of integration.\<close>
lemma equiintegrable_on_null [simp]:
"content(cbox a b) = 0 \<Longrightarrow> F equiintegrable_on cbox a b"
apply (auto simp: equiintegrable_on_def)
by (metis gauge_trivial norm_eq_zero sum_content_null)
text\<open> Main limit theorem for an equiintegrable sequence.\<close>
theorem equiintegrable_limit:
fixes g :: "'a :: euclidean_space \<Rightarrow> 'b :: banach"
assumes feq: "range f equiintegrable_on cbox a b"
and to_g: "\<And>x. x \<in> cbox a b \<Longrightarrow> (\<lambda>n. f n x) \<longlonglongrightarrow> g x"
shows "g integrable_on cbox a b \<and> (\<lambda>n. integral (cbox a b) (f n)) \<longlonglongrightarrow> integral (cbox a b) g"
proof -
have "Cauchy (\<lambda>n. integral(cbox a b) (f n))"
proof (clarsimp simp add: Cauchy_def)
fix e::real
assume "0 < e"
then have e3: "0 < e/3"
by simp
then obtain \<gamma> where "gauge \<gamma>"
and \<gamma>: "\<And>n \<D>. \<lbrakk>\<D> tagged_division_of cbox a b; \<gamma> fine \<D>\<rbrakk>
\<Longrightarrow> norm((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f n x) - integral (cbox a b) (f n)) < e/3"
using feq unfolding equiintegrable_on_def
by (meson image_eqI iso_tuple_UNIV_I)
obtain \<D> where \<D>: "\<D> tagged_division_of (cbox a b)" and "\<gamma> fine \<D>" "finite \<D>"
by (meson \<open>gauge \<gamma>\<close> fine_division_exists tagged_division_of_finite)
with \<gamma> have \<delta>T: "\<And>n. dist ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f n x)) (integral (cbox a b) (f n)) < e/3"
by (force simp: dist_norm)
have "(\<lambda>n. \<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f n x) \<longlonglongrightarrow> (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R g x)"
using \<D> to_g by (auto intro!: tendsto_sum tendsto_scaleR)
then have "Cauchy (\<lambda>n. \<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f n x)"
by (meson convergent_eq_Cauchy)
with e3 obtain M where
M: "\<And>m n. \<lbrakk>m\<ge>M; n\<ge>M\<rbrakk> \<Longrightarrow> dist (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f m x) (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f n x)
< e/3"
unfolding Cauchy_def by blast
have "\<And>m n. \<lbrakk>m\<ge>M; n\<ge>M;
dist (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f m x) (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f n x) < e/3\<rbrakk>
\<Longrightarrow> dist (integral (cbox a b) (f m)) (integral (cbox a b) (f n)) < e"
by (metis \<delta>T dist_commute dist_triangle_third [OF _ _ \<delta>T])
then show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (integral (cbox a b) (f m)) (integral (cbox a b) (f n)) < e"
using M by auto
qed
then obtain L where L: "(\<lambda>n. integral (cbox a b) (f n)) \<longlonglongrightarrow> L"
by (meson convergent_eq_Cauchy)
have "(g has_integral L) (cbox a b)"
proof (clarsimp simp: has_integral)
fix e::real assume "0 < e"
then have e2: "0 < e/2"
by simp
then obtain \<gamma> where "gauge \<gamma>"
and \<gamma>: "\<And>n \<D>. \<lbrakk>\<D> tagged_division_of cbox a b; \<gamma> fine \<D>\<rbrakk>
\<Longrightarrow> norm((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f n x) - integral (cbox a b) (f n)) < e/2"
using feq unfolding equiintegrable_on_def
by (meson image_eqI iso_tuple_UNIV_I)
moreover
have "norm ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R g x) - L) < e"
if "\<D> tagged_division_of cbox a b" "\<gamma> fine \<D>" for \<D>
proof -
have "norm ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R g x) - L) \<le> e/2"
proof (rule Lim_norm_ubound)
show "(\<lambda>n. (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f n x) - integral (cbox a b) (f n)) \<longlonglongrightarrow> (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R g x) - L"
using to_g that L
by (intro tendsto_diff tendsto_sum) (auto simp: tag_in_interval tendsto_scaleR)
show "\<forall>\<^sub>F n in sequentially.
norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f n x) - integral (cbox a b) (f n)) \<le> e/2"
by (intro eventuallyI less_imp_le \<gamma> that)
qed auto
with \<open>0 < e\<close> show ?thesis
by linarith
qed
ultimately
show "\<exists>\<gamma>. gauge \<gamma> \<and>
(\<forall>\<D>. \<D> tagged_division_of cbox a b \<and> \<gamma> fine \<D> \<longrightarrow>
norm ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R g x) - L) < e)"
by meson
qed
with L show ?thesis
by (simp add: \<open>(\<lambda>n. integral (cbox a b) (f n)) \<longlonglongrightarrow> L\<close> has_integral_integrable_integral)
qed
lemma equiintegrable_reflect:
assumes "F equiintegrable_on cbox a b"
shows "(\<lambda>f. f \<circ> uminus) ` F equiintegrable_on cbox (-b) (-a)"
proof -
have "\<exists>\<gamma>. gauge \<gamma> \<and>
(\<forall>f \<D>. f \<in> (\<lambda>f. f \<circ> uminus) ` F \<and> \<D> tagged_division_of cbox (- b) (- a) \<and> \<gamma> fine \<D> \<longrightarrow>
norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral (cbox (- b) (- a)) f) < e)"
if "gauge \<gamma>" and
\<gamma>: "\<And>f \<D>. \<lbrakk>f \<in> F; \<D> tagged_division_of cbox a b; \<gamma> fine \<D>\<rbrakk> \<Longrightarrow>
norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral (cbox a b) f) < e" for e \<gamma>
proof (intro exI, safe)
show "gauge (\<lambda>x. uminus ` \<gamma> (-x))"
by (metis \<open>gauge \<gamma>\<close> gauge_reflect)
show "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R (f \<circ> uminus) x) - integral (cbox (- b) (- a)) (f \<circ> uminus)) < e"
if "f \<in> F" and tag: "\<D> tagged_division_of cbox (- b) (- a)"
and fine: "(\<lambda>x. uminus ` \<gamma> (- x)) fine \<D>" for f \<D>
proof -
have 1: "(\<lambda>(x,K). (- x, uminus ` K)) ` \<D> tagged_partial_division_of cbox a b"
if "\<D> tagged_partial_division_of cbox (- b) (- a)"
proof -
have "- y \<in> cbox a b"
if "\<And>x K. (x,K) \<in> \<D> \<Longrightarrow> x \<in> K \<and> K \<subseteq> cbox (- b) (- a) \<and> (\<exists>a b. K = cbox a b)"
"(x, Y) \<in> \<D>" "y \<in> Y" for x Y y
proof -
have "y \<in> uminus ` cbox a b"
using that by auto
then show "- y \<in> cbox a b"
by force
qed
with that show ?thesis
by (fastforce simp: tagged_partial_division_of_def interior_negations image_iff)
qed
have 2: "\<exists>K. (\<exists>x. (x,K) \<in> (\<lambda>(x,K). (- x, uminus ` K)) ` \<D>) \<and> x \<in> K"
if "\<Union>{K. \<exists>x. (x,K) \<in> \<D>} = cbox (- b) (- a)" "x \<in> cbox a b" for x
proof -
have xm: "x \<in> uminus ` \<Union>{A. \<exists>a. (a, A) \<in> \<D>}"
by (simp add: that)
then obtain a X where "-x \<in> X" "(a, X) \<in> \<D>"
by auto
then show ?thesis
by (metis (no_types, lifting) add.inverse_inverse image_iff pair_imageI)
qed
have 3: "\<And>x X y. \<lbrakk>\<D> tagged_partial_division_of cbox (- b) (- a); (x, X) \<in> \<D>; y \<in> X\<rbrakk> \<Longrightarrow> - y \<in> cbox a b"
by (metis (no_types, lifting) equation_minus_iff imageE subsetD tagged_partial_division_ofD(3) uminus_interval_vector)
have tag': "(\<lambda>(x,K). (- x, uminus ` K)) ` \<D> tagged_division_of cbox a b"
using tag by (auto simp: tagged_division_of_def dest: 1 2 3)
have fine': "\<gamma> fine (\<lambda>(x,K). (- x, uminus ` K)) ` \<D>"
using fine by (fastforce simp: fine_def)
have inj: "inj_on (\<lambda>(x,K). (- x, uminus ` K)) \<D>"
unfolding inj_on_def by force
have eq: "content (uminus ` I) = content I"
if I: "(x, I) \<in> \<D>" and fnz: "f (- x) \<noteq> 0" for x I
proof -
obtain a b where "I = cbox a b"
using tag I that by (force simp: tagged_division_of_def tagged_partial_division_of_def)
then show ?thesis
using content_image_affinity_cbox [of "-1" 0] by auto
qed
have "(\<Sum>(x,K) \<in> (\<lambda>(x,K). (- x, uminus ` K)) ` \<D>. content K *\<^sub>R f x) =
(\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f (- x))"
apply (simp add: sum.reindex [OF inj])
apply (auto simp: eq intro!: sum.cong)
done
then show ?thesis
using \<gamma> [OF \<open>f \<in> F\<close> tag' fine'] integral_reflect
by (metis (mono_tags, lifting) Henstock_Kurzweil_Integration.integral_cong comp_apply split_def sum.cong)
qed
qed
then show ?thesis
using assms
apply (auto simp: equiintegrable_on_def)
apply (rule integrable_eq)
by auto
qed
subsection\<open>Subinterval restrictions for equiintegrable families\<close>
text\<open>First, some technical lemmas about minimizing a "flat" part of a sum over a division.\<close>
lemma lemma0:
assumes "i \<in> Basis"
shows "content (cbox u v) / (interval_upperbound (cbox u v) \<bullet> i - interval_lowerbound (cbox u v) \<bullet> i) =
(if content (cbox u v) = 0 then 0
else \<Prod>j \<in> Basis - {i}. interval_upperbound (cbox u v) \<bullet> j - interval_lowerbound (cbox u v) \<bullet> j)"
proof (cases "content (cbox u v) = 0")
case True
then show ?thesis by simp
next
case False
then show ?thesis
using prod.subset_diff [of "{i}" Basis] assms
by (force simp: content_cbox_if divide_simps split: if_split_asm)
qed
lemma content_division_lemma1:
assumes div: "\<D> division_of S" and S: "S \<subseteq> cbox a b" and i: "i \<in> Basis"
and mt: "\<And>K. K \<in> \<D> \<Longrightarrow> content K \<noteq> 0"
and disj: "(\<forall>K \<in> \<D>. K \<inter> {x. x \<bullet> i = a \<bullet> i} \<noteq> {}) \<or> (\<forall>K \<in> \<D>. K \<inter> {x. x \<bullet> i = b \<bullet> i} \<noteq> {})"
shows "(b \<bullet> i - a \<bullet> i) * (\<Sum>K\<in>\<D>. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))
\<le> content(cbox a b)" (is "?lhs \<le> ?rhs")
proof -
have "finite \<D>"
using div by blast
define extend where
"extend \<equiv> \<lambda>K. cbox (\<Sum>j \<in> Basis. if j = i then (a \<bullet> i) *\<^sub>R i else (interval_lowerbound K \<bullet> j) *\<^sub>R j)
(\<Sum>j \<in> Basis. if j = i then (b \<bullet> i) *\<^sub>R i else (interval_upperbound K \<bullet> j) *\<^sub>R j)"
have div_subset_cbox: "\<And>K. K \<in> \<D> \<Longrightarrow> K \<subseteq> cbox a b"
using S div by auto
have "\<And>K. K \<in> \<D> \<Longrightarrow> K \<noteq> {}"
using div by blast
have extend: "extend K \<noteq> {}" "extend K \<subseteq> cbox a b" if K: "K \<in> \<D>" for K
proof -
obtain u v where K: "K = cbox u v" "K \<noteq> {}" "K \<subseteq> cbox a b"
using K cbox_division_memE [OF _ div] by (meson div_subset_cbox)
with i show "extend K \<noteq> {}" "extend K \<subseteq> cbox a b"
- apply (auto simp: extend_def subset_box box_ne_empty sum_if_inner)
+ apply (auto simp: extend_def subset_box box_ne_empty)
by fastforce
qed
have int_extend_disjoint:
"interior(extend K1) \<inter> interior(extend K2) = {}" if K: "K1 \<in> \<D>" "K2 \<in> \<D>" "K1 \<noteq> K2" for K1 K2
proof -
obtain u v where K1: "K1 = cbox u v" "K1 \<noteq> {}" "K1 \<subseteq> cbox a b"
using K cbox_division_memE [OF _ div] by (meson div_subset_cbox)
obtain w z where K2: "K2 = cbox w z" "K2 \<noteq> {}" "K2 \<subseteq> cbox a b"
using K cbox_division_memE [OF _ div] by (meson div_subset_cbox)
have cboxes: "cbox u v \<in> \<D>" "cbox w z \<in> \<D>" "cbox u v \<noteq> cbox w z"
using K1 K2 that by auto
with div have "interior (cbox u v) \<inter> interior (cbox w z) = {}"
by blast
moreover
have "\<exists>x. x \<in> box u v \<and> x \<in> box w z"
if "x \<in> interior (extend K1)" "x \<in> interior (extend K2)" for x
proof -
have "a \<bullet> i < x \<bullet> i" "x \<bullet> i < b \<bullet> i"
and ux: "\<And>k. k \<in> Basis - {i} \<Longrightarrow> u \<bullet> k < x \<bullet> k"
and xv: "\<And>k. k \<in> Basis - {i} \<Longrightarrow> x \<bullet> k < v \<bullet> k"
and wx: "\<And>k. k \<in> Basis - {i} \<Longrightarrow> w \<bullet> k < x \<bullet> k"
and xz: "\<And>k. k \<in> Basis - {i} \<Longrightarrow> x \<bullet> k < z \<bullet> k"
- using that K1 K2 i by (auto simp: extend_def box_ne_empty sum_if_inner mem_box)
+ using that K1 K2 i by (auto simp: extend_def box_ne_empty mem_box)
have "box u v \<noteq> {}" "box w z \<noteq> {}"
using cboxes interior_cbox by (auto simp: content_eq_0_interior dest: mt)
then obtain q s
where q: "\<And>k. k \<in> Basis \<Longrightarrow> w \<bullet> k < q \<bullet> k \<and> q \<bullet> k < z \<bullet> k"
and s: "\<And>k. k \<in> Basis \<Longrightarrow> u \<bullet> k < s \<bullet> k \<and> s \<bullet> k < v \<bullet> k"
by (meson all_not_in_conv mem_box(1))
show ?thesis using disj
proof
assume "\<forall>K\<in>\<D>. K \<inter> {x. x \<bullet> i = a \<bullet> i} \<noteq> {}"
then have uva: "(cbox u v) \<inter> {x. x \<bullet> i = a \<bullet> i} \<noteq> {}"
and wza: "(cbox w z) \<inter> {x. x \<bullet> i = a \<bullet> i} \<noteq> {}"
using cboxes by (auto simp: content_eq_0_interior)
then obtain r t where "r \<bullet> i = a \<bullet> i" and r: "\<And>k. k \<in> Basis \<Longrightarrow> w \<bullet> k \<le> r \<bullet> k \<and> r \<bullet> k \<le> z \<bullet> k"
and "t \<bullet> i = a \<bullet> i" and t: "\<And>k. k \<in> Basis \<Longrightarrow> u \<bullet> k \<le> t \<bullet> k \<and> t \<bullet> k \<le> v \<bullet> k"
by (fastforce simp: mem_box)
have u: "u \<bullet> i < q \<bullet> i"
using i K2(1) K2(3) \<open>t \<bullet> i = a \<bullet> i\<close> q s t [OF i] by (force simp: subset_box)
have w: "w \<bullet> i < s \<bullet> i"
using i K1(1) K1(3) \<open>r \<bullet> i = a \<bullet> i\<close> s r [OF i] by (force simp: subset_box)
let ?x = "(\<Sum>j \<in> Basis. if j = i then min (q \<bullet> i) (s \<bullet> i) *\<^sub>R i else (x \<bullet> j) *\<^sub>R j)"
show ?thesis
proof (intro exI conjI)
show "?x \<in> box u v"
using \<open>i \<in> Basis\<close> s apply (clarsimp simp: mem_box)
apply (subst sum_if_inner; simp)+
apply (fastforce simp: u ux xv)
done
show "?x \<in> box w z"
using \<open>i \<in> Basis\<close> q apply (clarsimp simp: mem_box)
apply (subst sum_if_inner; simp)+
apply (fastforce simp: w wx xz)
done
qed
next
assume "\<forall>K\<in>\<D>. K \<inter> {x. x \<bullet> i = b \<bullet> i} \<noteq> {}"
then have uva: "(cbox u v) \<inter> {x. x \<bullet> i = b \<bullet> i} \<noteq> {}"
and wza: "(cbox w z) \<inter> {x. x \<bullet> i = b \<bullet> i} \<noteq> {}"
using cboxes by (auto simp: content_eq_0_interior)
then obtain r t where "r \<bullet> i = b \<bullet> i" and r: "\<And>k. k \<in> Basis \<Longrightarrow> w \<bullet> k \<le> r \<bullet> k \<and> r \<bullet> k \<le> z \<bullet> k"
and "t \<bullet> i = b \<bullet> i" and t: "\<And>k. k \<in> Basis \<Longrightarrow> u \<bullet> k \<le> t \<bullet> k \<and> t \<bullet> k \<le> v \<bullet> k"
by (fastforce simp: mem_box)
have z: "s \<bullet> i < z \<bullet> i"
using K1(1) K1(3) \<open>r \<bullet> i = b \<bullet> i\<close> r [OF i] i s by (force simp: subset_box)
have v: "q \<bullet> i < v \<bullet> i"
using K2(1) K2(3) \<open>t \<bullet> i = b \<bullet> i\<close> t [OF i] i q by (force simp: subset_box)
let ?x = "(\<Sum>j \<in> Basis. if j = i then max (q \<bullet> i) (s \<bullet> i) *\<^sub>R i else (x \<bullet> j) *\<^sub>R j)"
show ?thesis
proof (intro exI conjI)
show "?x \<in> box u v"
using \<open>i \<in> Basis\<close> s apply (clarsimp simp: mem_box)
apply (subst sum_if_inner; simp)+
apply (fastforce simp: v ux xv)
done
show "?x \<in> box w z"
using \<open>i \<in> Basis\<close> q apply (clarsimp simp: mem_box)
apply (subst sum_if_inner; simp)+
apply (fastforce simp: z wx xz)
done
qed
qed
qed
ultimately show ?thesis by auto
qed
have "?lhs = (\<Sum>K\<in>\<D>. (b \<bullet> i - a \<bullet> i) * content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))"
by (simp add: sum_distrib_left)
also have "\<dots> = sum (content \<circ> extend) \<D>"
proof (rule sum.cong [OF refl])
fix K assume "K \<in> \<D>"
then obtain u v where K: "K = cbox u v" "cbox u v \<noteq> {}" "K \<subseteq> cbox a b"
using cbox_division_memE [OF _ div] div_subset_cbox by metis
then have uv: "u \<bullet> i < v \<bullet> i"
using mt [OF \<open>K \<in> \<D>\<close>] \<open>i \<in> Basis\<close> content_eq_0 by fastforce
have "insert i (Basis \<inter> -{i}) = Basis"
using \<open>i \<in> Basis\<close> by auto
then have "(b \<bullet> i - a \<bullet> i) * content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)
= (b \<bullet> i - a \<bullet> i) * (\<Prod>i \<in> insert i (Basis \<inter> -{i}). v \<bullet> i - u \<bullet> i) / (interval_upperbound (cbox u v) \<bullet> i - interval_lowerbound (cbox u v) \<bullet> i)"
using K box_ne_empty(1) content_cbox by fastforce
also have "... = (\<Prod>x\<in>Basis. if x = i then b \<bullet> x - a \<bullet> x
else (interval_upperbound (cbox u v) - interval_lowerbound (cbox u v)) \<bullet> x)"
using \<open>i \<in> Basis\<close> K uv by (simp add: prod.If_cases) (simp add: algebra_simps)
also have "... = (\<Prod>k\<in>Basis.
(\<Sum>j\<in>Basis. if j = i then (b \<bullet> i - a \<bullet> i) *\<^sub>R i else ((interval_upperbound (cbox u v) - interval_lowerbound (cbox u v)) \<bullet> j) *\<^sub>R j) \<bullet> k)"
using \<open>i \<in> Basis\<close> by (subst prod.cong [OF refl sum_if_inner]; simp)
also have "... = (\<Prod>k\<in>Basis.
(\<Sum>j\<in>Basis. if j = i then (b \<bullet> i) *\<^sub>R i else (interval_upperbound (cbox u v) \<bullet> j) *\<^sub>R j) \<bullet> k -
(\<Sum>j\<in>Basis. if j = i then (a \<bullet> i) *\<^sub>R i else (interval_lowerbound (cbox u v) \<bullet> j) *\<^sub>R j) \<bullet> k)"
apply (rule prod.cong [OF refl])
using \<open>i \<in> Basis\<close>
apply (subst sum_if_inner; simp add: algebra_simps)+
done
also have "... = (content \<circ> extend) K"
using \<open>i \<in> Basis\<close> K box_ne_empty
apply (simp add: extend_def)
apply (subst content_cbox, auto)
done
finally show "(b \<bullet> i - a \<bullet> i) * content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)
= (content \<circ> extend) K" .
qed
also have "... = sum content (extend ` \<D>)"
proof -
have "\<lbrakk>K1 \<in> \<D>; K2 \<in> \<D>; K1 \<noteq> K2; extend K1 = extend K2\<rbrakk> \<Longrightarrow> content (extend K1) = 0" for K1 K2
using int_extend_disjoint [of K1 K2] extend_def by (simp add: content_eq_0_interior)
then show ?thesis
by (simp add: comm_monoid_add_class.sum.reindex_nontrivial [OF \<open>finite \<D>\<close>])
qed
also have "... \<le> ?rhs"
proof (rule subadditive_content_division)
show "extend ` \<D> division_of \<Union> (extend ` \<D>)"
using int_extend_disjoint apply (auto simp: division_of_def \<open>finite \<D>\<close> extend)
using extend_def apply blast
done
show "\<Union> (extend ` \<D>) \<subseteq> cbox a b"
using extend by fastforce
qed
finally show ?thesis .
qed
proposition sum_content_area_over_thin_division:
assumes div: "\<D> division_of S" and S: "S \<subseteq> cbox a b" and i: "i \<in> Basis"
and "a \<bullet> i \<le> c" "c \<le> b \<bullet> i"
and nonmt: "\<And>K. K \<in> \<D> \<Longrightarrow> K \<inter> {x. x \<bullet> i = c} \<noteq> {}"
shows "(b \<bullet> i - a \<bullet> i) * (\<Sum>K\<in>\<D>. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))
\<le> 2 * content(cbox a b)"
proof (cases "content(cbox a b) = 0")
case True
have "(\<Sum>K\<in>\<D>. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) = 0"
using S div by (force intro!: sum.neutral content_0_subset [OF True])
then show ?thesis
by (auto simp: True)
next
case False
then have "content(cbox a b) > 0"
using zero_less_measure_iff by blast
then have "a \<bullet> i < b \<bullet> i" if "i \<in> Basis" for i
using content_pos_lt_eq that by blast
have "finite \<D>"
using div by blast
define Dlec where "Dlec \<equiv> {L \<in> (\<lambda>L. L \<inter> {x. x \<bullet> i \<le> c}) ` \<D>. content L \<noteq> 0}"
define Dgec where "Dgec \<equiv> {L \<in> (\<lambda>L. L \<inter> {x. x \<bullet> i \<ge> c}) ` \<D>. content L \<noteq> 0}"
define a' where "a' \<equiv> (\<Sum>j\<in>Basis. (if j = i then c else a \<bullet> j) *\<^sub>R j)"
define b' where "b' \<equiv> (\<Sum>j\<in>Basis. (if j = i then c else b \<bullet> j) *\<^sub>R j)"
have Dlec_cbox: "\<And>K. K \<in> Dlec \<Longrightarrow> \<exists>a b. K = cbox a b"
using interval_split [OF i] div by (fastforce simp: Dlec_def division_of_def)
then have lec_is_cbox: "\<lbrakk>content (L \<inter> {x. x \<bullet> i \<le> c}) \<noteq> 0; L \<in> \<D>\<rbrakk> \<Longrightarrow> \<exists>a b. L \<inter> {x. x \<bullet> i \<le> c} = cbox a b" for L
using Dlec_def by blast
have Dgec_cbox: "\<And>K. K \<in> Dgec \<Longrightarrow> \<exists>a b. K = cbox a b"
using interval_split [OF i] div by (fastforce simp: Dgec_def division_of_def)
then have gec_is_cbox: "\<lbrakk>content (L \<inter> {x. x \<bullet> i \<ge> c}) \<noteq> 0; L \<in> \<D>\<rbrakk> \<Longrightarrow> \<exists>a b. L \<inter> {x. x \<bullet> i \<ge> c} = cbox a b" for L
using Dgec_def by blast
have "(b' \<bullet> i - a \<bullet> i) * (\<Sum>K\<in>Dlec. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<le> content(cbox a b')"
proof (rule content_division_lemma1)
show "Dlec division_of \<Union>Dlec"
unfolding division_of_def
proof (intro conjI ballI Dlec_cbox)
show "\<And>K1 K2. \<lbrakk>K1 \<in> Dlec; K2 \<in> Dlec\<rbrakk> \<Longrightarrow> K1 \<noteq> K2 \<longrightarrow> interior K1 \<inter> interior K2 = {}"
by (clarsimp simp: Dlec_def) (use div in auto)
qed (use \<open>finite \<D>\<close> Dlec_def in auto)
show "\<Union>Dlec \<subseteq> cbox a b'"
using Dlec_def div S by (auto simp: b'_def division_of_def mem_box)
show "(\<forall>K\<in>Dlec. K \<inter> {x. x \<bullet> i = a \<bullet> i} \<noteq> {}) \<or> (\<forall>K\<in>Dlec. K \<inter> {x. x \<bullet> i = b' \<bullet> i} \<noteq> {})"
- using nonmt by (fastforce simp: Dlec_def b'_def sum_if_inner i)
+ using nonmt by (fastforce simp: Dlec_def b'_def i)
qed (use i Dlec_def in auto)
moreover
have "(\<Sum>K\<in>Dlec. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) =
(\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<le> c}))) K)"
apply (subst sum.reindex_nontrivial [OF \<open>finite \<D>\<close>, symmetric], simp)
apply (metis division_split_left_inj [OF div] lec_is_cbox content_eq_0_interior)
unfolding Dlec_def using \<open>finite \<D>\<close> apply (auto simp: sum.mono_neutral_left)
done
moreover have "(b' \<bullet> i - a \<bullet> i) = (c - a \<bullet> i)"
- by (simp add: b'_def sum_if_inner i)
+ by (simp add: b'_def i)
ultimately
have lec: "(c - a \<bullet> i) * (\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<le> c}))) K)
\<le> content(cbox a b')"
by simp
have "(b \<bullet> i - a' \<bullet> i) * (\<Sum>K\<in>Dgec. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<le> content(cbox a' b)"
proof (rule content_division_lemma1)
show "Dgec division_of \<Union>Dgec"
unfolding division_of_def
proof (intro conjI ballI Dgec_cbox)
show "\<And>K1 K2. \<lbrakk>K1 \<in> Dgec; K2 \<in> Dgec\<rbrakk> \<Longrightarrow> K1 \<noteq> K2 \<longrightarrow> interior K1 \<inter> interior K2 = {}"
by (clarsimp simp: Dgec_def) (use div in auto)
qed (use \<open>finite \<D>\<close> Dgec_def in auto)
show "\<Union>Dgec \<subseteq> cbox a' b"
using Dgec_def div S by (auto simp: a'_def division_of_def mem_box)
show "(\<forall>K\<in>Dgec. K \<inter> {x. x \<bullet> i = a' \<bullet> i} \<noteq> {}) \<or> (\<forall>K\<in>Dgec. K \<inter> {x. x \<bullet> i = b \<bullet> i} \<noteq> {})"
- using nonmt by (fastforce simp: Dgec_def a'_def sum_if_inner i)
+ using nonmt by (fastforce simp: Dgec_def a'_def i)
qed (use i Dgec_def in auto)
moreover
have "(\<Sum>K\<in>Dgec. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) =
(\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<ge> c}))) K)"
apply (subst sum.reindex_nontrivial [OF \<open>finite \<D>\<close>, symmetric], simp)
apply (metis division_split_right_inj [OF div] gec_is_cbox content_eq_0_interior)
unfolding Dgec_def using \<open>finite \<D>\<close> apply (auto simp: sum.mono_neutral_left)
done
moreover have "(b \<bullet> i - a' \<bullet> i) = (b \<bullet> i - c)"
- by (simp add: a'_def sum_if_inner i)
+ by (simp add: a'_def i)
ultimately
have gec: "(b \<bullet> i - c) * (\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<ge> c}))) K)
\<le> content(cbox a' b)"
by simp
show ?thesis
proof (cases "c = a \<bullet> i \<or> c = b \<bullet> i")
case True
then show ?thesis
proof
assume c: "c = a \<bullet> i"
then have "a' = a"
- apply (simp add: sum_if_inner i a'_def cong: if_cong)
+ apply (simp add: i a'_def cong: if_cong)
using euclidean_representation [of a] sum.cong [OF refl, of Basis "\<lambda>i. (a \<bullet> i) *\<^sub>R i"] by presburger
then have "content (cbox a' b) \<le> 2 * content (cbox a b)" by simp
moreover
have eq: "(\<Sum>K\<in>\<D>. content (K \<inter> {x. a \<bullet> i \<le> x \<bullet> i}) /
(interval_upperbound (K \<inter> {x. a \<bullet> i \<le> x \<bullet> i}) \<bullet> i - interval_lowerbound (K \<inter> {x. a \<bullet> i \<le> x \<bullet> i}) \<bullet> i))
= (\<Sum>K\<in>\<D>. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))"
(is "sum ?f _ = sum ?g _")
proof (rule sum.cong [OF refl])
fix K assume "K \<in> \<D>"
then have "a \<bullet> i \<le> x \<bullet> i" if "x \<in> K" for x
by (metis S UnionI div division_ofD(6) i mem_box(2) subsetCE that)
then have "K \<inter> {x. a \<bullet> i \<le> x \<bullet> i} = K"
by blast
then show "?f K = ?g K"
by simp
qed
ultimately show ?thesis
using gec c eq by auto
next
assume c: "c = b \<bullet> i"
then have "b' = b"
- apply (simp add: sum_if_inner i b'_def cong: if_cong)
+ apply (simp add: i b'_def cong: if_cong)
using euclidean_representation [of b] sum.cong [OF refl, of Basis "\<lambda>i. (b \<bullet> i) *\<^sub>R i"] by presburger
then have "content (cbox a b') \<le> 2 * content (cbox a b)" by simp
moreover
have eq: "(\<Sum>K\<in>\<D>. content (K \<inter> {x. x \<bullet> i \<le> b \<bullet> i}) /
(interval_upperbound (K \<inter> {x. x \<bullet> i \<le> b \<bullet> i}) \<bullet> i - interval_lowerbound (K \<inter> {x. x \<bullet> i \<le> b \<bullet> i}) \<bullet> i))
= (\<Sum>K\<in>\<D>. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))"
(is "sum ?f _ = sum ?g _")
proof (rule sum.cong [OF refl])
fix K assume "K \<in> \<D>"
then have "x \<bullet> i \<le> b \<bullet> i" if "x \<in> K" for x
by (metis S UnionI div division_ofD(6) i mem_box(2) subsetCE that)
then have "K \<inter> {x. x \<bullet> i \<le> b \<bullet> i} = K"
by blast
then show "?f K = ?g K"
by simp
qed
ultimately show ?thesis
using lec c eq by auto
qed
next
case False
have prod_if: "(\<Prod>k\<in>Basis \<inter> - {i}. f k) = (\<Prod>k\<in>Basis. f k) / f i" if "f i \<noteq> (0::real)" for f
using that mk_disjoint_insert [OF i]
apply (clarsimp simp add: field_split_simps)
- by (metis Int_insert_left_if0 finite_Basis finite_insert le_iff_inf mult.commute order_refl prod.insert subset_Compl_singleton)
+ by (metis Int_insert_left_if0 finite_Basis finite_insert le_iff_inf order_refl prod.insert subset_Compl_singleton)
have abc: "a \<bullet> i < c" "c < b \<bullet> i"
using False assms by auto
then have "(\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<le> c}))) K)
\<le> content(cbox a b') / (c - a \<bullet> i)"
"(\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<ge> c}))) K)
\<le> content(cbox a' b) / (b \<bullet> i - c)"
- using lec gec by (simp_all add: field_split_simps mult.commute)
+ using lec gec by (simp_all add: field_split_simps)
moreover
have "(\<Sum>K\<in>\<D>. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))
\<le> (\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<le> c}))) K) +
(\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<ge> c}))) K)"
(is "?lhs \<le> ?rhs")
proof -
have "?lhs \<le>
(\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<le> c}))) K +
((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<ge> c}))) K)"
(is "sum ?f _ \<le> sum ?g _")
proof (rule sum_mono)
fix K assume "K \<in> \<D>"
then obtain u v where uv: "K = cbox u v"
using div by blast
obtain u' v' where uv': "cbox u v \<inter> {x. x \<bullet> i \<le> c} = cbox u v'"
"cbox u v \<inter> {x. c \<le> x \<bullet> i} = cbox u' v"
"\<And>k. k \<in> Basis \<Longrightarrow> u' \<bullet> k = (if k = i then max (u \<bullet> i) c else u \<bullet> k)"
"\<And>k. k \<in> Basis \<Longrightarrow> v' \<bullet> k = (if k = i then min (v \<bullet> i) c else v \<bullet> k)"
using i by (auto simp: interval_split)
have *: "\<lbrakk>content (cbox u v') = 0; content (cbox u' v) = 0\<rbrakk> \<Longrightarrow> content (cbox u v) = 0"
"content (cbox u' v) \<noteq> 0 \<Longrightarrow> content (cbox u v) \<noteq> 0"
"content (cbox u v') \<noteq> 0 \<Longrightarrow> content (cbox u v) \<noteq> 0"
using i uv uv' by (auto simp: content_eq_0 le_max_iff_disj min_le_iff_disj split: if_split_asm intro: order_trans)
show "?f K \<le> ?g K"
using i uv uv' apply (clarsimp simp add: lemma0 * intro!: prod_nonneg)
by (metis content_eq_0 le_less_linear order.strict_implies_order)
qed
also have "... = ?rhs"
by (simp add: sum.distrib)
finally show ?thesis .
qed
moreover have "content (cbox a b') / (c - a \<bullet> i) = content (cbox a b) / (b \<bullet> i - a \<bullet> i)"
using i abc
apply (simp add: field_simps a'_def b'_def measure_lborel_cbox_eq inner_diff)
apply (auto simp: if_distrib if_distrib [of "\<lambda>f. f x" for x] prod.If_cases [of Basis "\<lambda>x. x = i", simplified] prod_if field_simps)
done
moreover have "content (cbox a' b) / (b \<bullet> i - c) = content (cbox a b) / (b \<bullet> i - a \<bullet> i)"
using i abc
apply (simp add: field_simps a'_def b'_def measure_lborel_cbox_eq inner_diff)
apply (auto simp: if_distrib prod.If_cases [of Basis "\<lambda>x. x = i", simplified] prod_if field_simps)
done
ultimately
have "(\<Sum>K\<in>\<D>. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))
\<le> 2 * content (cbox a b) / (b \<bullet> i - a \<bullet> i)"
by linarith
then show ?thesis
- using abc by (simp add: field_split_simps mult.commute)
+ using abc by (simp add: field_split_simps)
qed
qed
proposition bounded_equiintegral_over_thin_tagged_partial_division:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes F: "F equiintegrable_on cbox a b" and f: "f \<in> F" and "0 < \<epsilon>"
and norm_f: "\<And>h x. \<lbrakk>h \<in> F; x \<in> cbox a b\<rbrakk> \<Longrightarrow> norm(h x) \<le> norm(f x)"
obtains \<gamma> where "gauge \<gamma>"
"\<And>c i S h. \<lbrakk>c \<in> cbox a b; i \<in> Basis; S tagged_partial_division_of cbox a b;
\<gamma> fine S; h \<in> F; \<And>x K. (x,K) \<in> S \<Longrightarrow> (K \<inter> {x. x \<bullet> i = c \<bullet> i} \<noteq> {})\<rbrakk>
\<Longrightarrow> (\<Sum>(x,K) \<in> S. norm (integral K h)) < \<epsilon>"
proof (cases "content(cbox a b) = 0")
case True
show ?thesis
proof
show "gauge (\<lambda>x. ball x 1)"
by (simp add: gauge_trivial)
show "(\<Sum>(x,K) \<in> S. norm (integral K h)) < \<epsilon>"
if "S tagged_partial_division_of cbox a b" "(\<lambda>x. ball x 1) fine S" for S and h:: "'a \<Rightarrow> 'b"
proof -
have "(\<Sum>(x,K) \<in> S. norm (integral K h)) = 0"
using that True content_0_subset
by (fastforce simp: tagged_partial_division_of_def intro: sum.neutral)
with \<open>0 < \<epsilon>\<close> show ?thesis
by simp
qed
qed
next
case False
then have contab_gt0: "content(cbox a b) > 0"
by (simp add: zero_less_measure_iff)
then have a_less_b: "\<And>i. i \<in> Basis \<Longrightarrow> a\<bullet>i < b\<bullet>i"
by (auto simp: content_pos_lt_eq)
obtain \<gamma>0 where "gauge \<gamma>0"
and \<gamma>0: "\<And>S h. \<lbrakk>S tagged_partial_division_of cbox a b; \<gamma>0 fine S; h \<in> F\<rbrakk>
\<Longrightarrow> (\<Sum>(x,K) \<in> S. norm (content K *\<^sub>R h x - integral K h)) < \<epsilon>/2"
proof -
obtain \<gamma> where "gauge \<gamma>"
and \<gamma>: "\<And>f \<D>. \<lbrakk>f \<in> F; \<D> tagged_division_of cbox a b; \<gamma> fine \<D>\<rbrakk>
\<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral (cbox a b) f)
< \<epsilon>/(5 * (Suc DIM('b)))"
proof -
have e5: "\<epsilon>/(5 * (Suc DIM('b))) > 0"
using \<open>\<epsilon> > 0\<close> by auto
then show ?thesis
using F that by (auto simp: equiintegrable_on_def)
qed
show ?thesis
proof
show "gauge \<gamma>"
by (rule \<open>gauge \<gamma>\<close>)
show "(\<Sum>(x,K) \<in> S. norm (content K *\<^sub>R h x - integral K h)) < \<epsilon>/2"
if "S tagged_partial_division_of cbox a b" "\<gamma> fine S" "h \<in> F" for S h
proof -
have "(\<Sum>(x,K) \<in> S. norm (content K *\<^sub>R h x - integral K h)) \<le> 2 * real DIM('b) * (\<epsilon>/(5 * Suc DIM('b)))"
proof (rule Henstock_lemma_part2 [of h a b])
show "h integrable_on cbox a b"
using that F equiintegrable_on_def by metis
show "gauge \<gamma>"
by (rule \<open>gauge \<gamma>\<close>)
qed (use that \<open>\<epsilon> > 0\<close> \<gamma> in auto)
also have "... < \<epsilon>/2"
using \<open>\<epsilon> > 0\<close> by (simp add: divide_simps)
finally show ?thesis .
qed
qed
qed
define \<gamma> where "\<gamma> \<equiv> \<lambda>x. \<gamma>0 x \<inter>
ball x ((\<epsilon>/8 / (norm(f x) + 1)) * (INF m\<in>Basis. b \<bullet> m - a \<bullet> m) / content(cbox a b))"
have "gauge (\<lambda>x. ball x
(\<epsilon> * (INF m\<in>Basis. b \<bullet> m - a \<bullet> m) / ((8 * norm (f x) + 8) * content (cbox a b))))"
using \<open>0 < content (cbox a b)\<close> \<open>0 < \<epsilon>\<close> a_less_b
- apply (auto simp: gauge_def field_split_simps mult_less_0_iff zero_less_mult_iff add_nonneg_eq_0_iff finite_less_Inf_iff)
+ apply (auto simp: gauge_def field_split_simps add_nonneg_eq_0_iff finite_less_Inf_iff)
apply (meson add_increasing measure_nonneg mult_nonneg_nonneg norm_ge_zero not_le zero_le_numeral)
done
then have "gauge \<gamma>"
unfolding \<gamma>_def using \<open>gauge \<gamma>0\<close> gauge_Int by auto
moreover
have "(\<Sum>(x,K) \<in> S. norm (integral K h)) < \<epsilon>"
if "c \<in> cbox a b" "i \<in> Basis" and S: "S tagged_partial_division_of cbox a b"
and "\<gamma> fine S" "h \<in> F" and ne: "\<And>x K. (x,K) \<in> S \<Longrightarrow> K \<inter> {x. x \<bullet> i = c \<bullet> i} \<noteq> {}" for c i S h
proof -
have "cbox c b \<subseteq> cbox a b"
by (meson mem_box(2) order_refl subset_box(1) that(1))
have "finite S"
using S by blast
have "\<gamma>0 fine S" and fineS:
"(\<lambda>x. ball x (\<epsilon> * (INF m\<in>Basis. b \<bullet> m - a \<bullet> m) / ((8 * norm (f x) + 8) * content (cbox a b)))) fine S"
using \<open>\<gamma> fine S\<close> by (auto simp: \<gamma>_def fine_Int)
then have "(\<Sum>(x,K) \<in> S. norm (content K *\<^sub>R h x - integral K h)) < \<epsilon>/2"
by (intro \<gamma>0 that fineS)
moreover have "(\<Sum>(x,K) \<in> S. norm (integral K h) - norm (content K *\<^sub>R h x - integral K h)) \<le> \<epsilon>/2"
proof -
have "(\<Sum>(x,K) \<in> S. norm (integral K h) - norm (content K *\<^sub>R h x - integral K h))
\<le> (\<Sum>(x,K) \<in> S. norm (content K *\<^sub>R h x))"
proof (clarify intro!: sum_mono)
fix x K
assume xK: "(x,K) \<in> S"
have "norm (integral K h) - norm (content K *\<^sub>R h x - integral K h) \<le> norm (integral K h - (integral K h - content K *\<^sub>R h x))"
by (metis norm_minus_commute norm_triangle_ineq2)
also have "... \<le> norm (content K *\<^sub>R h x)"
by simp
finally show "norm (integral K h) - norm (content K *\<^sub>R h x - integral K h) \<le> norm (content K *\<^sub>R h x)" .
qed
also have "... \<le> (\<Sum>(x,K) \<in> S. \<epsilon>/4 * (b \<bullet> i - a \<bullet> i) / content (cbox a b) *
content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))"
proof (clarify intro!: sum_mono)
fix x K
assume xK: "(x,K) \<in> S"
then have x: "x \<in> cbox a b"
using S unfolding tagged_partial_division_of_def by (meson subset_iff)
let ?\<Delta> = "interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i"
show "norm (content K *\<^sub>R h x) \<le> \<epsilon>/4 * (b \<bullet> i - a \<bullet> i) / content (cbox a b) * content K / ?\<Delta>"
proof (cases "content K = 0")
case True
then show ?thesis by simp
next
case False
then have Kgt0: "content K > 0"
using zero_less_measure_iff by blast
moreover
obtain u v where uv: "K = cbox u v"
using S \<open>(x,K) \<in> S\<close> by blast
then have u_less_v: "\<And>i. i \<in> Basis \<Longrightarrow> u \<bullet> i < v \<bullet> i"
using content_pos_lt_eq uv Kgt0 by blast
then have dist_uv: "dist u v > 0"
using that by auto
ultimately have "norm (h x) \<le> (\<epsilon> * (b \<bullet> i - a \<bullet> i)) / (4 * content (cbox a b) * ?\<Delta>)"
proof -
have "dist x u < \<epsilon> * (INF m\<in>Basis. b \<bullet> m - a \<bullet> m) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2"
"dist x v < \<epsilon> * (INF m\<in>Basis. b \<bullet> m - a \<bullet> m) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2"
using fineS u_less_v uv xK
by (force simp: fine_def mem_box field_simps dest!: bspec)+
moreover have "\<epsilon> * (INF m\<in>Basis. b \<bullet> m - a \<bullet> m) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2
\<le> \<epsilon> * (b \<bullet> i - a \<bullet> i) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2"
apply (intro mult_left_mono divide_right_mono)
using \<open>i \<in> Basis\<close> \<open>0 < \<epsilon>\<close> apply (auto simp: intro!: cInf_le_finite)
done
ultimately
have "dist x u < \<epsilon> * (b \<bullet> i - a \<bullet> i) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2"
"dist x v < \<epsilon> * (b \<bullet> i - a \<bullet> i) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2"
by linarith+
then have duv: "dist u v < \<epsilon> * (b \<bullet> i - a \<bullet> i) / (4 * (norm (f x) + 1) * content (cbox a b))"
using dist_triangle_half_r by blast
have uvi: "\<bar>v \<bullet> i - u \<bullet> i\<bar> \<le> norm (v - u)"
by (metis inner_commute inner_diff_right \<open>i \<in> Basis\<close> Basis_le_norm)
have "norm (h x) \<le> norm (f x)"
using x that by (auto simp: norm_f)
also have "... < (norm (f x) + 1)"
by simp
also have "... < \<epsilon> * (b \<bullet> i - a \<bullet> i) / dist u v / (4 * content (cbox a b))"
using duv dist_uv contab_gt0
- apply (simp add: field_split_simps algebra_simps mult_less_0_iff zero_less_mult_iff split: if_split_asm)
+ apply (simp add: field_split_simps split: if_split_asm)
by (meson add_nonneg_nonneg linorder_not_le measure_nonneg mult_nonneg_nonneg norm_ge_zero zero_le_numeral)
also have "... = \<epsilon> * (b \<bullet> i - a \<bullet> i) / norm (v - u) / (4 * content (cbox a b))"
by (simp add: dist_norm norm_minus_commute)
also have "... \<le> \<epsilon> * (b \<bullet> i - a \<bullet> i) / \<bar>v \<bullet> i - u \<bullet> i\<bar> / (4 * content (cbox a b))"
apply (intro mult_right_mono divide_left_mono divide_right_mono uvi)
using \<open>0 < \<epsilon>\<close> a_less_b [OF \<open>i \<in> Basis\<close>] u_less_v [OF \<open>i \<in> Basis\<close>] contab_gt0
by (auto simp: less_eq_real_def zero_less_mult_iff that)
also have "... = \<epsilon> * (b \<bullet> i - a \<bullet> i)
/ (4 * content (cbox a b) * ?\<Delta>)"
using uv False that(2) u_less_v by fastforce
finally show ?thesis by simp
qed
with Kgt0 have "norm (content K *\<^sub>R h x) \<le> content K * ((\<epsilon>/4 * (b \<bullet> i - a \<bullet> i) / content (cbox a b)) / ?\<Delta>)"
using mult_left_mono by fastforce
also have "... = \<epsilon>/4 * (b \<bullet> i - a \<bullet> i) / content (cbox a b) *
content K / ?\<Delta>"
by (simp add: field_split_simps)
finally show ?thesis .
qed
qed
also have "... = (\<Sum>K\<in>snd ` S. \<epsilon>/4 * (b \<bullet> i - a \<bullet> i) / content (cbox a b) * content K
/ (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))"
apply (rule sum.over_tagged_division_lemma [OF tagged_partial_division_of_Union_self [OF S]])
apply (simp add: box_eq_empty(1) content_eq_0)
done
also have "... = \<epsilon>/2 * ((b \<bullet> i - a \<bullet> i) / (2 * content (cbox a b)) * (\<Sum>K\<in>snd ` S. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)))"
by (simp add: sum_distrib_left mult.assoc)
also have "... \<le> (\<epsilon>/2) * 1"
proof (rule mult_left_mono)
have "(b \<bullet> i - a \<bullet> i) * (\<Sum>K\<in>snd ` S. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))
\<le> 2 * content (cbox a b)"
proof (rule sum_content_area_over_thin_division)
show "snd ` S division_of \<Union>(snd ` S)"
by (auto intro: S tagged_partial_division_of_Union_self division_of_tagged_division)
show "\<Union>(snd ` S) \<subseteq> cbox a b"
using S by force
show "a \<bullet> i \<le> c \<bullet> i" "c \<bullet> i \<le> b \<bullet> i"
using mem_box(2) that by blast+
qed (use that in auto)
then show "(b \<bullet> i - a \<bullet> i) / (2 * content (cbox a b)) * (\<Sum>K\<in>snd ` S. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<le> 1"
by (simp add: contab_gt0)
qed (use \<open>0 < \<epsilon>\<close> in auto)
finally show ?thesis by simp
qed
then have "(\<Sum>(x,K) \<in> S. norm (integral K h)) - (\<Sum>(x,K) \<in> S. norm (content K *\<^sub>R h x - integral K h)) \<le> \<epsilon>/2"
by (simp add: Groups_Big.sum_subtractf [symmetric])
ultimately show "(\<Sum>(x,K) \<in> S. norm (integral K h)) < \<epsilon>"
by linarith
qed
ultimately show ?thesis using that by auto
qed
proposition equiintegrable_halfspace_restrictions_le:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes F: "F equiintegrable_on cbox a b" and f: "f \<in> F"
and norm_f: "\<And>h x. \<lbrakk>h \<in> F; x \<in> cbox a b\<rbrakk> \<Longrightarrow> norm(h x) \<le> norm(f x)"
shows "(\<Union>i \<in> Basis. \<Union>c. \<Union>h \<in> F. {(\<lambda>x. if x \<bullet> i \<le> c then h x else 0)})
equiintegrable_on cbox a b"
proof (cases "content(cbox a b) = 0")
case True
then show ?thesis by simp
next
case False
then have "content(cbox a b) > 0"
using zero_less_measure_iff by blast
then have "a \<bullet> i < b \<bullet> i" if "i \<in> Basis" for i
using content_pos_lt_eq that by blast
have int_F: "f integrable_on cbox a b" if "f \<in> F" for f
using F that by (simp add: equiintegrable_on_def)
let ?CI = "\<lambda>K h x. content K *\<^sub>R h x - integral K h"
show ?thesis
unfolding equiintegrable_on_def
proof (intro conjI; clarify)
show int_lec: "\<lbrakk>i \<in> Basis; h \<in> F\<rbrakk> \<Longrightarrow> (\<lambda>x. if x \<bullet> i \<le> c then h x else 0) integrable_on cbox a b" for i c h
using integrable_restrict_Int [of "{x. x \<bullet> i \<le> c}" h]
apply (auto simp: interval_split Int_commute mem_box intro!: integrable_on_subcbox int_F)
by (metis (full_types, hide_lams) min.bounded_iff)
show "\<exists>\<gamma>. gauge \<gamma> \<and>
(\<forall>f T. f \<in> (\<Union>i\<in>Basis. \<Union>c. \<Union>h\<in>F. {\<lambda>x. if x \<bullet> i \<le> c then h x else 0}) \<and>
T tagged_division_of cbox a b \<and> \<gamma> fine T \<longrightarrow>
norm ((\<Sum>(x,K) \<in> T. content K *\<^sub>R f x) - integral (cbox a b) f) < \<epsilon>)"
if "\<epsilon> > 0" for \<epsilon>
proof -
obtain \<gamma>0 where "gauge \<gamma>0" and \<gamma>0:
"\<And>c i S h. \<lbrakk>c \<in> cbox a b; i \<in> Basis; S tagged_partial_division_of cbox a b;
\<gamma>0 fine S; h \<in> F; \<And>x K. (x,K) \<in> S \<Longrightarrow> (K \<inter> {x. x \<bullet> i = c \<bullet> i} \<noteq> {})\<rbrakk>
\<Longrightarrow> (\<Sum>(x,K) \<in> S. norm (integral K h)) < \<epsilon>/12"
apply (rule bounded_equiintegral_over_thin_tagged_partial_division [OF F f, of \<open>\<epsilon>/12\<close>])
using \<open>\<epsilon> > 0\<close> by (auto simp: norm_f)
obtain \<gamma>1 where "gauge \<gamma>1"
and \<gamma>1: "\<And>h T. \<lbrakk>h \<in> F; T tagged_division_of cbox a b; \<gamma>1 fine T\<rbrakk>
\<Longrightarrow> norm ((\<Sum>(x,K) \<in> T. content K *\<^sub>R h x) - integral (cbox a b) h)
< \<epsilon>/(7 * (Suc DIM('b)))"
proof -
have e5: "\<epsilon>/(7 * (Suc DIM('b))) > 0"
using \<open>\<epsilon> > 0\<close> by auto
then show ?thesis
using F that by (auto simp: equiintegrable_on_def)
qed
have h_less3: "(\<Sum>(x,K) \<in> T. norm (?CI K h x)) < \<epsilon>/3"
if "T tagged_partial_division_of cbox a b" "\<gamma>1 fine T" "h \<in> F" for T h
proof -
have "(\<Sum>(x,K) \<in> T. norm (?CI K h x)) \<le> 2 * real DIM('b) * (\<epsilon>/(7 * Suc DIM('b)))"
proof (rule Henstock_lemma_part2 [of h a b])
show "h integrable_on cbox a b"
using that F equiintegrable_on_def by metis
show "gauge \<gamma>1"
by (rule \<open>gauge \<gamma>1\<close>)
qed (use that \<open>\<epsilon> > 0\<close> \<gamma>1 in auto)
also have "... < \<epsilon>/3"
using \<open>\<epsilon> > 0\<close> by (simp add: divide_simps)
finally show ?thesis .
qed
have *: "norm ((\<Sum>(x,K) \<in> T. content K *\<^sub>R f x) - integral (cbox a b) f) < \<epsilon>"
if f: "f = (\<lambda>x. if x \<bullet> i \<le> c then h x else 0)"
and T: "T tagged_division_of cbox a b"
and fine: "(\<lambda>x. \<gamma>0 x \<inter> \<gamma>1 x) fine T" and "i \<in> Basis" "h \<in> F" for f T i c h
proof (cases "a \<bullet> i \<le> c \<and> c \<le> b \<bullet> i")
case True
have "finite T"
using T by blast
define T' where "T' \<equiv> {(x,K) \<in> T. K \<inter> {x. x \<bullet> i \<le> c} \<noteq> {}}"
then have "T' \<subseteq> T"
by auto
then have "finite T'"
using \<open>finite T\<close> infinite_super by blast
have T'_tagged: "T' tagged_partial_division_of cbox a b"
by (meson T \<open>T' \<subseteq> T\<close> tagged_division_of_def tagged_partial_division_subset)
have fine': "\<gamma>0 fine T'" "\<gamma>1 fine T'"
using \<open>T' \<subseteq> T\<close> fine_Int fine_subset fine by blast+
have int_KK': "(\<Sum>(x,K) \<in> T. integral K f) = (\<Sum>(x,K) \<in> T'. integral K f)"
apply (rule sum.mono_neutral_right [OF \<open>finite T\<close> \<open>T' \<subseteq> T\<close>])
using f \<open>finite T\<close> \<open>T' \<subseteq> T\<close>
using integral_restrict_Int [of _ "{x. x \<bullet> i \<le> c}" h]
apply (auto simp: T'_def Int_commute)
done
have "(\<Sum>(x,K) \<in> T. content K *\<^sub>R f x) = (\<Sum>(x,K) \<in> T'. content K *\<^sub>R f x)"
apply (rule sum.mono_neutral_right [OF \<open>finite T\<close> \<open>T' \<subseteq> T\<close>])
using T f \<open>finite T\<close> \<open>T' \<subseteq> T\<close> apply (force simp: T'_def)
done
moreover have "norm ((\<Sum>(x,K) \<in> T'. content K *\<^sub>R f x) - integral (cbox a b) f) < \<epsilon>"
proof -
have *: "norm y < \<epsilon>" if "norm x < \<epsilon>/3" "norm(x - y) \<le> 2 * \<epsilon>/3" for x y::'b
proof -
have "norm y \<le> norm x + norm(x - y)"
by (metis norm_minus_commute norm_triangle_sub)
also have "\<dots> < \<epsilon>/3 + 2*\<epsilon>/3"
using that by linarith
also have "... = \<epsilon>"
by simp
finally show ?thesis .
qed
have "norm (\<Sum>(x,K) \<in> T'. ?CI K h x)
\<le> (\<Sum>(x,K) \<in> T'. norm (?CI K h x))"
by (simp add: norm_sum split_def)
also have "... < \<epsilon>/3"
by (intro h_less3 T'_tagged fine' that)
finally have "norm (\<Sum>(x,K) \<in> T'. ?CI K h x) < \<epsilon>/3" .
moreover have "integral (cbox a b) f = (\<Sum>(x,K) \<in> T. integral K f)"
using int_lec that by (auto simp: integral_combine_tagged_division_topdown)
moreover have "norm (\<Sum>(x,K) \<in> T'. ?CI K h x - ?CI K f x)
\<le> 2*\<epsilon>/3"
proof -
define T'' where "T'' \<equiv> {(x,K) \<in> T'. \<not> (K \<subseteq> {x. x \<bullet> i \<le> c})}"
then have "T'' \<subseteq> T'"
by auto
then have "finite T''"
using \<open>finite T'\<close> infinite_super by blast
have T''_tagged: "T'' tagged_partial_division_of cbox a b"
using T'_tagged \<open>T'' \<subseteq> T'\<close> tagged_partial_division_subset by blast
have fine'': "\<gamma>0 fine T''" "\<gamma>1 fine T''"
using \<open>T'' \<subseteq> T'\<close> fine' by (blast intro: fine_subset)+
have "(\<Sum>(x,K) \<in> T'. ?CI K h x - ?CI K f x)
= (\<Sum>(x,K) \<in> T''. ?CI K h x - ?CI K f x)"
proof (clarify intro!: sum.mono_neutral_right [OF \<open>finite T'\<close> \<open>T'' \<subseteq> T'\<close>])
fix x K
assume "(x,K) \<in> T'" "(x,K) \<notin> T''"
then have "x \<in> K" "x \<bullet> i \<le> c" "{x. x \<bullet> i \<le> c} \<inter> K = K"
using T''_def T'_tagged by blast+
then show "?CI K h x - ?CI K f x = 0"
using integral_restrict_Int [of _ "{x. x \<bullet> i \<le> c}" h] by (auto simp: f)
qed
moreover have "norm (\<Sum>(x,K) \<in> T''. ?CI K h x - ?CI K f x) \<le> 2*\<epsilon>/3"
proof -
define A where "A \<equiv> {(x,K) \<in> T''. x \<bullet> i \<le> c}"
define B where "B \<equiv> {(x,K) \<in> T''. x \<bullet> i > c}"
then have "A \<subseteq> T''" "B \<subseteq> T''" and disj: "A \<inter> B = {}" and T''_eq: "T'' = A \<union> B"
by (auto simp: A_def B_def)
then have "finite A" "finite B"
using \<open>finite T''\<close> by (auto intro: finite_subset)
have A_tagged: "A tagged_partial_division_of cbox a b"
using T''_tagged \<open>A \<subseteq> T''\<close> tagged_partial_division_subset by blast
have fineA: "\<gamma>0 fine A" "\<gamma>1 fine A"
using \<open>A \<subseteq> T''\<close> fine'' by (blast intro: fine_subset)+
have B_tagged: "B tagged_partial_division_of cbox a b"
using T''_tagged \<open>B \<subseteq> T''\<close> tagged_partial_division_subset by blast
have fineB: "\<gamma>0 fine B" "\<gamma>1 fine B"
using \<open>B \<subseteq> T''\<close> fine'' by (blast intro: fine_subset)+
have "norm (\<Sum>(x,K) \<in> T''. ?CI K h x - ?CI K f x)
\<le> (\<Sum>(x,K) \<in> T''. norm (?CI K h x - ?CI K f x))"
by (simp add: norm_sum split_def)
also have "... = (\<Sum>(x,K) \<in> A. norm (?CI K h x - ?CI K f x)) +
(\<Sum>(x,K) \<in> B. norm (?CI K h x - ?CI K f x))"
by (simp add: sum.union_disjoint T''_eq disj \<open>finite A\<close> \<open>finite B\<close>)
also have "... = (\<Sum>(x,K) \<in> A. norm (integral K h - integral K f)) +
(\<Sum>(x,K) \<in> B. norm (?CI K h x + integral K f))"
by (auto simp: A_def B_def f norm_minus_commute intro!: sum.cong arg_cong2 [where f= "(+)"])
also have "... \<le> (\<Sum>(x,K)\<in>A. norm (integral K h)) +
(\<Sum>(x,K)\<in>(\<lambda>(x,K). (x,K \<inter> {x. x \<bullet> i \<le> c})) ` A. norm (integral K h))
+ ((\<Sum>(x,K)\<in>B. norm (?CI K h x)) +
(\<Sum>(x,K)\<in>B. norm (integral K h)) +
(\<Sum>(x,K)\<in>(\<lambda>(x,K). (x,K \<inter> {x. c \<le> x \<bullet> i})) ` B. norm (integral K h)))"
proof (rule add_mono)
show "(\<Sum>(x,K)\<in>A. norm (integral K h - integral K f))
\<le> (\<Sum>(x,K)\<in>A. norm (integral K h)) +
(\<Sum>(x,K)\<in>(\<lambda>(x,K). (x,K \<inter> {x. x \<bullet> i \<le> c})) ` A.
norm (integral K h))"
proof (subst sum.reindex_nontrivial [OF \<open>finite A\<close>], clarsimp)
fix x K L
assume "(x,K) \<in> A" "(x,L) \<in> A"
and int_ne0: "integral (L \<inter> {x. x \<bullet> i \<le> c}) h \<noteq> 0"
and eq: "K \<inter> {x. x \<bullet> i \<le> c} = L \<inter> {x. x \<bullet> i \<le> c}"
have False if "K \<noteq> L"
proof -
obtain u v where uv: "L = cbox u v"
using T'_tagged \<open>(x, L) \<in> A\<close> \<open>A \<subseteq> T''\<close> \<open>T'' \<subseteq> T'\<close> by blast
have "A tagged_division_of \<Union>(snd ` A)"
using A_tagged tagged_partial_division_of_Union_self by auto
then have "interior (K \<inter> {x. x \<bullet> i \<le> c}) = {}"
apply (rule tagged_division_split_left_inj [OF _ \<open>(x,K) \<in> A\<close> \<open>(x,L) \<in> A\<close>])
using that eq \<open>i \<in> Basis\<close> by auto
then show False
using interval_split [OF \<open>i \<in> Basis\<close>] int_ne0 content_eq_0_interior eq uv by fastforce
qed
then show "K = L" by blast
next
show "(\<Sum>(x,K) \<in> A. norm (integral K h - integral K f))
\<le> (\<Sum>(x,K) \<in> A. norm (integral K h)) +
sum ((\<lambda>(x,K). norm (integral K h)) \<circ> (\<lambda>(x,K). (x,K \<inter> {x. x \<bullet> i \<le> c}))) A"
using integral_restrict_Int [of _ "{x. x \<bullet> i \<le> c}" h] f
by (auto simp: Int_commute A_def [symmetric] sum.distrib [symmetric] intro!: sum_mono norm_triangle_ineq4)
qed
next
show "(\<Sum>(x,K)\<in>B. norm (?CI K h x + integral K f))
\<le> (\<Sum>(x,K)\<in>B. norm (?CI K h x)) + (\<Sum>(x,K)\<in>B. norm (integral K h)) +
(\<Sum>(x,K)\<in>(\<lambda>(x,K). (x,K \<inter> {x. c \<le> x \<bullet> i})) ` B. norm (integral K h))"
proof (subst sum.reindex_nontrivial [OF \<open>finite B\<close>], clarsimp)
fix x K L
assume "(x,K) \<in> B" "(x,L) \<in> B"
and int_ne0: "integral (L \<inter> {x. c \<le> x \<bullet> i}) h \<noteq> 0"
and eq: "K \<inter> {x. c \<le> x \<bullet> i} = L \<inter> {x. c \<le> x \<bullet> i}"
have False if "K \<noteq> L"
proof -
obtain u v where uv: "L = cbox u v"
using T'_tagged \<open>(x, L) \<in> B\<close> \<open>B \<subseteq> T''\<close> \<open>T'' \<subseteq> T'\<close> by blast
have "B tagged_division_of \<Union>(snd ` B)"
using B_tagged tagged_partial_division_of_Union_self by auto
then have "interior (K \<inter> {x. c \<le> x \<bullet> i}) = {}"
apply (rule tagged_division_split_right_inj [OF _ \<open>(x,K) \<in> B\<close> \<open>(x,L) \<in> B\<close>])
using that eq \<open>i \<in> Basis\<close> by auto
then show False
using interval_split [OF \<open>i \<in> Basis\<close>] int_ne0
content_eq_0_interior eq uv by fastforce
qed
then show "K = L" by blast
next
show "(\<Sum>(x,K) \<in> B. norm (?CI K h x + integral K f))
\<le> (\<Sum>(x,K) \<in> B. norm (?CI K h x)) +
(\<Sum>(x,K) \<in> B. norm (integral K h)) + sum ((\<lambda>(x,K). norm (integral K h)) \<circ> (\<lambda>(x,K). (x,K \<inter> {x. c \<le> x \<bullet> i}))) B"
proof (clarsimp simp: B_def [symmetric] sum.distrib [symmetric] intro!: sum_mono)
fix x K
assume "(x,K) \<in> B"
have *: "i = i1 + i2 \<Longrightarrow> norm(c + i1) \<le> norm c + norm i + norm(i2)"
for i::'b and c i1 i2
by (metis add.commute add.left_commute add_diff_cancel_right' dual_order.refl norm_add_rule_thm norm_triangle_ineq4)
obtain u v where uv: "K = cbox u v"
using T'_tagged \<open>(x,K) \<in> B\<close> \<open>B \<subseteq> T''\<close> \<open>T'' \<subseteq> T'\<close> by blast
have "h integrable_on cbox a b"
by (simp add: int_F \<open>h \<in> F\<close>)
then have huv: "h integrable_on cbox u v"
apply (rule integrable_on_subcbox)
using B_tagged \<open>(x,K) \<in> B\<close> uv by blast
have "integral K h = integral K f + integral (K \<inter> {x. c \<le> x \<bullet> i}) h"
using integral_restrict_Int [of _ "{x. x \<bullet> i \<le> c}" h] f uv \<open>i \<in> Basis\<close>
by (simp add: Int_commute integral_split [OF huv \<open>i \<in> Basis\<close>])
then show "norm (?CI K h x + integral K f)
\<le> norm (?CI K h x) + norm (integral K h) + norm (integral (K \<inter> {x. c \<le> x \<bullet> i}) h)"
by (rule *)
qed
qed
qed
also have "... \<le> 2*\<epsilon>/3"
proof -
have overlap: "K \<inter> {x. x \<bullet> i = c} \<noteq> {}" if "(x,K) \<in> T''" for x K
proof -
obtain y y' where y: "y' \<in> K" "c < y' \<bullet> i" "y \<in> K" "y \<bullet> i \<le> c"
using that T''_def T'_def \<open>(x,K) \<in> T''\<close> by fastforce
obtain u v where uv: "K = cbox u v"
using T''_tagged \<open>(x,K) \<in> T''\<close> by blast
then have "connected K"
- by (simp add: is_interval_cbox is_interval_connected)
+ by (simp add: is_interval_connected)
then have "(\<exists>z \<in> K. z \<bullet> i = c)"
using y connected_ivt_component by fastforce
then show ?thesis
by fastforce
qed
have **: "\<lbrakk>x < \<epsilon>/12; y < \<epsilon>/12; z \<le> \<epsilon>/2\<rbrakk> \<Longrightarrow> x + y + z \<le> 2 * \<epsilon>/3" for x y z
by auto
show ?thesis
proof (rule **)
have cb_ab: "(\<Sum>j \<in> Basis. if j = i then c *\<^sub>R i else (a \<bullet> j) *\<^sub>R j) \<in> cbox a b"
using \<open>i \<in> Basis\<close> True \<open>\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i < b \<bullet> i\<close>
apply (clarsimp simp add: mem_box)
apply (subst sum_if_inner | force)+
done
show "(\<Sum>(x,K) \<in> A. norm (integral K h)) < \<epsilon>/12"
apply (rule \<gamma>0 [OF cb_ab \<open>i \<in> Basis\<close> A_tagged fineA(1) \<open>h \<in> F\<close>])
using \<open>i \<in> Basis\<close> \<open>A \<subseteq> T''\<close> overlap
apply (subst sum_if_inner | force)+
done
have 1: "(\<lambda>(x,K). (x,K \<inter> {x. x \<bullet> i \<le> c})) ` A tagged_partial_division_of cbox a b"
using \<open>finite A\<close> \<open>i \<in> Basis\<close>
apply (auto simp: tagged_partial_division_of_def)
using A_tagged apply (auto simp: A_def)
using interval_split(1) by blast
have 2: "\<gamma>0 fine (\<lambda>(x,K). (x,K \<inter> {x. x \<bullet> i \<le> c})) ` A"
using fineA(1) fine_def by fastforce
show "(\<Sum>(x,K) \<in> (\<lambda>(x,K). (x,K \<inter> {x. x \<bullet> i \<le> c})) ` A. norm (integral K h)) < \<epsilon>/12"
apply (rule \<gamma>0 [OF cb_ab \<open>i \<in> Basis\<close> 1 2 \<open>h \<in> F\<close>])
using \<open>i \<in> Basis\<close> apply (subst sum_if_inner | force)+
using overlap apply (auto simp: A_def)
done
have *: "\<lbrakk>x < \<epsilon>/3; y < \<epsilon>/12; z < \<epsilon>/12\<rbrakk> \<Longrightarrow> x + y + z \<le> \<epsilon>/2" for x y z
by auto
show "(\<Sum>(x,K) \<in> B. norm (?CI K h x)) +
(\<Sum>(x,K) \<in> B. norm (integral K h)) +
(\<Sum>(x,K) \<in> (\<lambda>(x,K). (x,K \<inter> {x. c \<le> x \<bullet> i})) ` B. norm (integral K h))
\<le> \<epsilon>/2"
proof (rule *)
show "(\<Sum>(x,K) \<in> B. norm (?CI K h x)) < \<epsilon>/3"
by (intro h_less3 B_tagged fineB that)
show "(\<Sum>(x,K) \<in> B. norm (integral K h)) < \<epsilon>/12"
apply (rule \<gamma>0 [OF cb_ab \<open>i \<in> Basis\<close> B_tagged fineB(1) \<open>h \<in> F\<close>])
using \<open>i \<in> Basis\<close> \<open>B \<subseteq> T''\<close> overlap by (subst sum_if_inner | force)+
have 1: "(\<lambda>(x,K). (x,K \<inter> {x. c \<le> x \<bullet> i})) ` B tagged_partial_division_of cbox a b"
using \<open>finite B\<close> \<open>i \<in> Basis\<close>
apply (auto simp: tagged_partial_division_of_def)
using B_tagged apply (auto simp: B_def)
using interval_split(2) by blast
have 2: "\<gamma>0 fine (\<lambda>(x,K). (x,K \<inter> {x. c \<le> x \<bullet> i})) ` B"
using fineB(1) fine_def by fastforce
show "(\<Sum>(x,K) \<in> (\<lambda>(x,K). (x,K \<inter> {x. c \<le> x \<bullet> i})) ` B. norm (integral K h)) < \<epsilon>/12"
apply (rule \<gamma>0 [OF cb_ab \<open>i \<in> Basis\<close> 1 2 \<open>h \<in> F\<close>])
using \<open>i \<in> Basis\<close> apply (subst sum_if_inner | force)+
using overlap apply (auto simp: B_def)
done
qed
qed
qed
finally show ?thesis .
qed
ultimately show ?thesis by metis
qed
ultimately show ?thesis
by (simp add: sum_subtractf [symmetric] int_KK' *)
qed
ultimately show ?thesis by metis
next
case False
then consider "c < a \<bullet> i" | "b \<bullet> i < c"
by auto
then show ?thesis
proof cases
case 1
then have f0: "f x = 0" if "x \<in> cbox a b" for x
using that f \<open>i \<in> Basis\<close> mem_box(2) by force
then have int_f0: "integral (cbox a b) f = 0"
by (simp add: integral_cong)
have f0_tag: "f x = 0" if "(x,K) \<in> T" for x K
using T f0 that by (force simp: tagged_division_of_def)
then have "(\<Sum>(x,K) \<in> T. content K *\<^sub>R f x) = 0"
by (metis (mono_tags, lifting) real_vector.scale_eq_0_iff split_conv sum.neutral surj_pair)
then show ?thesis
using \<open>0 < \<epsilon>\<close> by (simp add: int_f0)
next
case 2
then have fh: "f x = h x" if "x \<in> cbox a b" for x
using that f \<open>i \<in> Basis\<close> mem_box(2) by force
then have int_f: "integral (cbox a b) f = integral (cbox a b) h"
using integral_cong by blast
have fh_tag: "f x = h x" if "(x,K) \<in> T" for x K
using T fh that by (force simp: tagged_division_of_def)
then have "(\<Sum>(x,K) \<in> T. content K *\<^sub>R f x) = (\<Sum>(x,K) \<in> T. content K *\<^sub>R h x)"
by (metis (mono_tags, lifting) split_cong sum.cong)
with \<open>0 < \<epsilon>\<close> show ?thesis
apply (simp add: int_f)
apply (rule less_trans [OF \<gamma>1])
using that fine_Int apply (force simp: divide_simps)+
done
qed
qed
have "gauge (\<lambda>x. \<gamma>0 x \<inter> \<gamma>1 x)"
by (simp add: \<open>gauge \<gamma>0\<close> \<open>gauge \<gamma>1\<close> gauge_Int)
then show ?thesis
by (auto intro: *)
qed
qed
qed
corollary equiintegrable_halfspace_restrictions_ge:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes F: "F equiintegrable_on cbox a b" and f: "f \<in> F"
and norm_f: "\<And>h x. \<lbrakk>h \<in> F; x \<in> cbox a b\<rbrakk> \<Longrightarrow> norm(h x) \<le> norm(f x)"
shows "(\<Union>i \<in> Basis. \<Union>c. \<Union>h \<in> F. {(\<lambda>x. if x \<bullet> i \<ge> c then h x else 0)})
equiintegrable_on cbox a b"
proof -
have *: "(\<Union>i\<in>Basis. \<Union>c. \<Union>h\<in>(\<lambda>f. f \<circ> uminus) ` F. {\<lambda>x. if x \<bullet> i \<le> c then h x else 0})
equiintegrable_on cbox (- b) (- a)"
proof (rule equiintegrable_halfspace_restrictions_le)
show "(\<lambda>f. f \<circ> uminus) ` F equiintegrable_on cbox (- b) (- a)"
using F equiintegrable_reflect by blast
show "f \<circ> uminus \<in> (\<lambda>f. f \<circ> uminus) ` F"
using f by auto
show "\<And>h x. \<lbrakk>h \<in> (\<lambda>f. f \<circ> uminus) ` F; x \<in> cbox (- b) (- a)\<rbrakk> \<Longrightarrow> norm (h x) \<le> norm ((f \<circ> uminus) x)"
using f apply (clarsimp simp:)
by (metis add.inverse_inverse image_eqI norm_f uminus_interval_vector)
qed
have eq: "(\<lambda>f. f \<circ> uminus) `
(\<Union>i\<in>Basis. \<Union>c. \<Union>h\<in>F. {\<lambda>x. if x \<bullet> i \<le> c then (h \<circ> uminus) x else 0}) =
(\<Union>i\<in>Basis. \<Union>c. \<Union>h\<in>F. {\<lambda>x. if c \<le> x \<bullet> i then h x else 0})"
apply (auto simp: o_def cong: if_cong)
using minus_le_iff apply fastforce
apply (rule_tac x="\<lambda>x. if c \<le> (-x) \<bullet> i then h(-x) else 0" in image_eqI)
using le_minus_iff apply fastforce+
done
show ?thesis
using equiintegrable_reflect [OF *] by (auto simp: eq)
qed
corollary equiintegrable_halfspace_restrictions_lt:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes F: "F equiintegrable_on cbox a b" and f: "f \<in> F"
and norm_f: "\<And>h x. \<lbrakk>h \<in> F; x \<in> cbox a b\<rbrakk> \<Longrightarrow> norm(h x) \<le> norm(f x)"
shows "(\<Union>i \<in> Basis. \<Union>c. \<Union>h \<in> F. {(\<lambda>x. if x \<bullet> i < c then h x else 0)}) equiintegrable_on cbox a b"
(is "?G equiintegrable_on cbox a b")
proof -
have *: "(\<Union>i\<in>Basis. \<Union>c. \<Union>h\<in>F. {\<lambda>x. if c \<le> x \<bullet> i then h x else 0}) equiintegrable_on cbox a b"
using equiintegrable_halfspace_restrictions_ge [OF F f] norm_f by auto
have "(\<lambda>x. if x \<bullet> i < c then h x else 0) = (\<lambda>x. h x - (if c \<le> x \<bullet> i then h x else 0))"
if "i \<in> Basis" "h \<in> F" for i c h
using that by force
then show ?thesis
by (blast intro: equiintegrable_on_subset [OF equiintegrable_diff [OF F *]])
qed
corollary equiintegrable_halfspace_restrictions_gt:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes F: "F equiintegrable_on cbox a b" and f: "f \<in> F"
and norm_f: "\<And>h x. \<lbrakk>h \<in> F; x \<in> cbox a b\<rbrakk> \<Longrightarrow> norm(h x) \<le> norm(f x)"
shows "(\<Union>i \<in> Basis. \<Union>c. \<Union>h \<in> F. {(\<lambda>x. if x \<bullet> i > c then h x else 0)}) equiintegrable_on cbox a b"
(is "?G equiintegrable_on cbox a b")
proof -
have *: "(\<Union>i\<in>Basis. \<Union>c. \<Union>h\<in>F. {\<lambda>x. if c \<ge> x \<bullet> i then h x else 0}) equiintegrable_on cbox a b"
using equiintegrable_halfspace_restrictions_le [OF F f] norm_f by auto
have "(\<lambda>x. if x \<bullet> i > c then h x else 0) = (\<lambda>x. h x - (if c \<ge> x \<bullet> i then h x else 0))"
if "i \<in> Basis" "h \<in> F" for i c h
using that by force
then show ?thesis
by (blast intro: equiintegrable_on_subset [OF equiintegrable_diff [OF F *]])
qed
proposition equiintegrable_closed_interval_restrictions:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes f: "f integrable_on cbox a b"
shows "(\<Union>c d. {(\<lambda>x. if x \<in> cbox c d then f x else 0)}) equiintegrable_on cbox a b"
proof -
let ?g = "\<lambda>B c d x. if \<forall>i\<in>B. c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i then f x else 0"
have *: "insert f (\<Union>c d. {?g B c d}) equiintegrable_on cbox a b" if "B \<subseteq> Basis" for B
proof -
have "finite B"
using finite_Basis finite_subset \<open>B \<subseteq> Basis\<close> by blast
then show ?thesis using \<open>B \<subseteq> Basis\<close>
proof (induction B)
case empty
with f show ?case by auto
next
case (insert i B)
then have "i \<in> Basis"
by auto
have *: "norm (h x) \<le> norm (f x)"
if "h \<in> insert f (\<Union>c d. {?g B c d})" "x \<in> cbox a b" for h x
using that by auto
have "(\<Union>i\<in>Basis.
\<Union>\<xi>. \<Union>h\<in>insert f (\<Union>i\<in>Basis. \<Union>\<psi>. \<Union>h\<in>insert f (\<Union>c d. {?g B c d}). {\<lambda>x. if x \<bullet> i \<le> \<psi> then h x else 0}).
{\<lambda>x. if \<xi> \<le> x \<bullet> i then h x else 0})
equiintegrable_on cbox a b"
proof (rule equiintegrable_halfspace_restrictions_ge [where f=f])
show "insert f (\<Union>i\<in>Basis. \<Union>\<xi>. \<Union>h\<in>insert f (\<Union>c d. {?g B c d}).
{\<lambda>x. if x \<bullet> i \<le> \<xi> then h x else 0}) equiintegrable_on cbox a b"
apply (intro * f equiintegrable_on_insert equiintegrable_halfspace_restrictions_le [OF insert.IH insertI1])
using insert.prems apply auto
done
show"norm(h x) \<le> norm(f x)"
if "h \<in> insert f (\<Union>i\<in>Basis. \<Union>\<xi>. \<Union>h\<in>insert f (\<Union>c d. {?g B c d}). {\<lambda>x. if x \<bullet> i \<le> \<xi> then h x else 0})"
"x \<in> cbox a b" for h x
using that by auto
qed auto
then have "insert f (\<Union>i\<in>Basis.
\<Union>\<xi>. \<Union>h\<in>insert f (\<Union>i\<in>Basis. \<Union>\<psi>. \<Union>h\<in>insert f (\<Union>c d. {?g B c d}). {\<lambda>x. if x \<bullet> i \<le> \<psi> then h x else 0}).
{\<lambda>x. if \<xi> \<le> x \<bullet> i then h x else 0})
equiintegrable_on cbox a b"
by (blast intro: f equiintegrable_on_insert)
then show ?case
apply (rule equiintegrable_on_subset, clarify)
using \<open>i \<in> Basis\<close> apply simp
apply (drule_tac x=i in bspec, assumption)
apply (drule_tac x="c \<bullet> i" in spec, clarify)
apply (drule_tac x=i in bspec, assumption)
apply (drule_tac x="d \<bullet> i" in spec)
apply (clarsimp simp add: fun_eq_iff)
apply (drule_tac x=c in spec)
apply (drule_tac x=d in spec)
apply (simp add: split: if_split_asm)
done
qed
qed
show ?thesis
by (rule equiintegrable_on_subset [OF * [OF subset_refl]]) (auto simp: mem_box)
qed
subsection\<open>Continuity of the indefinite integral\<close>
proposition indefinite_integral_continuous:
fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
assumes int_f: "f integrable_on cbox a b"
and c: "c \<in> cbox a b" and d: "d \<in> cbox a b" "0 < \<epsilon>"
obtains \<delta> where "0 < \<delta>"
"\<And>c' d'. \<lbrakk>c' \<in> cbox a b; d' \<in> cbox a b; norm(c' - c) \<le> \<delta>; norm(d' - d) \<le> \<delta>\<rbrakk>
\<Longrightarrow> norm(integral(cbox c' d') f - integral(cbox c d) f) < \<epsilon>"
proof -
{ assume "\<exists>c' d'. c' \<in> cbox a b \<and> d' \<in> cbox a b \<and> norm(c' - c) \<le> \<delta> \<and> norm(d' - d) \<le> \<delta> \<and>
norm(integral(cbox c' d') f - integral(cbox c d) f) \<ge> \<epsilon>"
(is "\<exists>c' d'. ?\<Phi> c' d' \<delta>") if "0 < \<delta>" for \<delta>
then have "\<exists>c' d'. ?\<Phi> c' d' (1 / Suc n)" for n
by simp
then obtain u v where "\<And>n. ?\<Phi> (u n) (v n) (1 / Suc n)"
by metis
then have u: "u n \<in> cbox a b" and norm_u: "norm(u n - c) \<le> 1 / Suc n"
and v: "v n \<in> cbox a b" and norm_v: "norm(v n - d) \<le> 1 / Suc n"
and \<epsilon>: "\<epsilon> \<le> norm (integral (cbox (u n) (v n)) f - integral (cbox c d) f)" for n
by blast+
then have False
proof -
have uvn: "cbox (u n) (v n) \<subseteq> cbox a b" for n
by (meson u v mem_box(2) subset_box(1))
define S where "S \<equiv> \<Union>i \<in> Basis. {x. x \<bullet> i = c \<bullet> i} \<union> {x. x \<bullet> i = d \<bullet> i}"
have "negligible S"
unfolding S_def by force
then have int_f': "(\<lambda>x. if x \<in> S then 0 else f x) integrable_on cbox a b"
by (force intro: integrable_spike assms)
have get_n: "\<exists>n. \<forall>m\<ge>n. x \<in> cbox (u m) (v m) \<longleftrightarrow> x \<in> cbox c d" if x: "x \<notin> S" for x
proof -
define \<epsilon> where "\<epsilon> \<equiv> Min ((\<lambda>i. min \<bar>x \<bullet> i - c \<bullet> i\<bar> \<bar>x \<bullet> i - d \<bullet> i\<bar>) ` Basis)"
have "\<epsilon> > 0"
using \<open>x \<notin> S\<close> by (auto simp: S_def \<epsilon>_def)
then obtain n where "n \<noteq> 0" and n: "1 / (real n) < \<epsilon>"
by (metis inverse_eq_divide real_arch_inverse)
have emin: "\<epsilon> \<le> min \<bar>x \<bullet> i - c \<bullet> i\<bar> \<bar>x \<bullet> i - d \<bullet> i\<bar>" if "i \<in> Basis" for i
unfolding \<epsilon>_def
apply (rule Min.coboundedI)
using that by force+
have "1 / real (Suc n) < \<epsilon>"
using n \<open>n \<noteq> 0\<close> \<open>\<epsilon> > 0\<close> by (simp add: field_simps)
have "x \<in> cbox (u m) (v m) \<longleftrightarrow> x \<in> cbox c d" if "m \<ge> n" for m
proof -
have *: "\<lbrakk>\<bar>u - c\<bar> \<le> n; \<bar>v - d\<bar> \<le> n; N < \<bar>x - c\<bar>; N < \<bar>x - d\<bar>; n \<le> N\<rbrakk>
\<Longrightarrow> u \<le> x \<and> x \<le> v \<longleftrightarrow> c \<le> x \<and> x \<le> d" for N n u v c d and x::real
by linarith
have "(u m \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> v m \<bullet> i) = (c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i)"
if "i \<in> Basis" for i
proof (rule *)
show "\<bar>u m \<bullet> i - c \<bullet> i\<bar> \<le> 1 / Suc m"
using norm_u [of m]
by (metis (full_types) order_trans Basis_le_norm inner_commute inner_diff_right that)
show "\<bar>v m \<bullet> i - d \<bullet> i\<bar> \<le> 1 / real (Suc m)"
using norm_v [of m]
by (metis (full_types) order_trans Basis_le_norm inner_commute inner_diff_right that)
show "1/n < \<bar>x \<bullet> i - c \<bullet> i\<bar>" "1/n < \<bar>x \<bullet> i - d \<bullet> i\<bar>"
using n \<open>n \<noteq> 0\<close> emin [OF \<open>i \<in> Basis\<close>]
by (simp_all add: inverse_eq_divide)
show "1 / real (Suc m) \<le> 1 / real n"
using \<open>n \<noteq> 0\<close> \<open>m \<ge> n\<close> by (simp add: field_split_simps)
qed
then show ?thesis by (simp add: mem_box)
qed
then show ?thesis by blast
qed
have 1: "range (\<lambda>n x. if x \<in> cbox (u n) (v n) then if x \<in> S then 0 else f x else 0) equiintegrable_on cbox a b"
by (blast intro: equiintegrable_on_subset [OF equiintegrable_closed_interval_restrictions [OF int_f']])
have 2: "(\<lambda>n. if x \<in> cbox (u n) (v n) then if x \<in> S then 0 else f x else 0)
\<longlonglongrightarrow> (if x \<in> cbox c d then if x \<in> S then 0 else f x else 0)" for x
by (fastforce simp: dest: get_n intro: tendsto_eventually eventually_sequentiallyI)
have [simp]: "cbox c d \<inter> cbox a b = cbox c d"
using c d by (force simp: mem_box)
have [simp]: "cbox (u n) (v n) \<inter> cbox a b = cbox (u n) (v n)" for n
using u v by (fastforce simp: mem_box intro: order.trans)
have "\<And>y A. y \<in> A - S \<Longrightarrow> f y = (\<lambda>x. if x \<in> S then 0 else f x) y"
by simp
then have "\<And>A. integral A (\<lambda>x. if x \<in> S then 0 else f (x)) = integral A (\<lambda>x. f (x))"
by (blast intro: integral_spike [OF \<open>negligible S\<close>])
moreover
obtain N where "dist (integral (cbox (u N) (v N)) (\<lambda>x. if x \<in> S then 0 else f x))
(integral (cbox c d) (\<lambda>x. if x \<in> S then 0 else f x)) < \<epsilon>"
using equiintegrable_limit [OF 1 2] \<open>0 < \<epsilon>\<close> by (force simp: integral_restrict_Int lim_sequentially)
ultimately have "dist (integral (cbox (u N) (v N)) f) (integral (cbox c d) f) < \<epsilon>"
by simp
then show False
by (metis dist_norm not_le \<epsilon>)
qed
}
then show ?thesis
by (meson not_le that)
qed
corollary indefinite_integral_uniformly_continuous:
fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
assumes "f integrable_on cbox a b"
shows "uniformly_continuous_on (cbox (Pair a a) (Pair b b)) (\<lambda>y. integral (cbox (fst y) (snd y)) f)"
proof -
show ?thesis
proof (rule compact_uniformly_continuous, clarsimp simp add: continuous_on_iff)
fix c d and \<epsilon>::real
assume c: "c \<in> cbox a b" and d: "d \<in> cbox a b" and "0 < \<epsilon>"
obtain \<delta> where "0 < \<delta>" and \<delta>:
"\<And>c' d'. \<lbrakk>c' \<in> cbox a b; d' \<in> cbox a b; norm(c' - c) \<le> \<delta>; norm(d' - d) \<le> \<delta>\<rbrakk>
\<Longrightarrow> norm(integral(cbox c' d') f -
integral(cbox c d) f) < \<epsilon>"
using indefinite_integral_continuous \<open>0 < \<epsilon>\<close> assms c d by blast
show "\<exists>\<delta> > 0. \<forall>x' \<in> cbox (a, a) (b, b).
dist x' (c, d) < \<delta> \<longrightarrow>
dist (integral (cbox (fst x') (snd x')) f)
(integral (cbox c d) f)
< \<epsilon>"
using \<open>0 < \<delta>\<close>
by (force simp: dist_norm intro: \<delta> order_trans [OF norm_fst_le] order_trans [OF norm_snd_le] less_imp_le)
qed auto
qed
corollary bounded_integrals_over_subintervals:
fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
assumes "f integrable_on cbox a b"
shows "bounded {integral (cbox c d) f |c d. cbox c d \<subseteq> cbox a b}"
proof -
have "bounded ((\<lambda>y. integral (cbox (fst y) (snd y)) f) ` cbox (a, a) (b, b))"
(is "bounded ?I")
by (blast intro: bounded_cbox bounded_uniformly_continuous_image indefinite_integral_uniformly_continuous [OF assms])
then obtain B where "B > 0" and B: "\<And>x. x \<in> ?I \<Longrightarrow> norm x \<le> B"
by (auto simp: bounded_pos)
have "norm x \<le> B" if "x = integral (cbox c d) f" "cbox c d \<subseteq> cbox a b" for x c d
proof (cases "cbox c d = {}")
case True
with \<open>0 < B\<close> that show ?thesis by auto
next
case False
show ?thesis
apply (rule B)
using that \<open>B > 0\<close> False apply (clarsimp simp: image_def)
by (metis cbox_Pair_iff interval_subset_is_interval is_interval_cbox prod.sel)
qed
then show ?thesis
by (blast intro: boundedI)
qed
text\<open>An existence theorem for "improper" integrals.
Hake's theorem implies that if the integrals over subintervals have a limit, the integral exists.
We only need to assume that the integrals are bounded, and we get absolute integrability,
but we also need a (rather weak) bound assumption on the function.\<close>
theorem absolutely_integrable_improper:
fixes f :: "'M::euclidean_space \<Rightarrow> 'N::euclidean_space"
assumes int_f: "\<And>c d. cbox c d \<subseteq> box a b \<Longrightarrow> f integrable_on cbox c d"
and bo: "bounded {integral (cbox c d) f |c d. cbox c d \<subseteq> box a b}"
and absi: "\<And>i. i \<in> Basis
\<Longrightarrow> \<exists>g. g absolutely_integrable_on cbox a b \<and>
((\<forall>x \<in> cbox a b. f x \<bullet> i \<le> g x) \<or> (\<forall>x \<in> cbox a b. f x \<bullet> i \<ge> g x))"
shows "f absolutely_integrable_on cbox a b"
proof (cases "content(cbox a b) = 0")
case True
then show ?thesis
by auto
next
case False
then have pos: "content(cbox a b) > 0"
using zero_less_measure_iff by blast
show ?thesis
unfolding absolutely_integrable_componentwise_iff [where f = f]
proof
fix j::'N
assume "j \<in> Basis"
then obtain g where absint_g: "g absolutely_integrable_on cbox a b"
and g: "(\<forall>x \<in> cbox a b. f x \<bullet> j \<le> g x) \<or> (\<forall>x \<in> cbox a b. f x \<bullet> j \<ge> g x)"
using absi by blast
have int_gab: "g integrable_on cbox a b"
using absint_g set_lebesgue_integral_eq_integral(1) by blast
have 1: "cbox (a + (b - a) /\<^sub>R real (Suc n)) (b - (b - a) /\<^sub>R real (Suc n)) \<subseteq> box a b" for n
apply (rule subset_box_imp)
using pos apply (auto simp: content_pos_lt_eq algebra_simps)
done
have 2: "cbox (a + (b - a) /\<^sub>R real (Suc n)) (b - (b - a) /\<^sub>R real (Suc n)) \<subseteq>
cbox (a + (b - a) /\<^sub>R real (Suc n + 1)) (b - (b - a) /\<^sub>R real (Suc n + 1))" for n
apply (rule subset_box_imp)
using pos apply (simp add: content_pos_lt_eq algebra_simps)
apply (simp add: divide_simps)
apply (auto simp: field_simps)
done
have getN: "\<exists>N::nat. \<forall>k. k \<ge> N \<longrightarrow> x \<in> cbox (a + (b - a) /\<^sub>R real k) (b - (b - a) /\<^sub>R real k)"
if x: "x \<in> box a b" for x
proof -
let ?\<Delta> = "(\<Union>i \<in> Basis. {((x - a) \<bullet> i) / ((b - a) \<bullet> i), (b - x) \<bullet> i / ((b - a) \<bullet> i)})"
obtain N where N: "real N > 1 / Inf ?\<Delta>"
using reals_Archimedean2 by blast
moreover have \<Delta>: "Inf ?\<Delta> > 0"
using that by (auto simp: finite_less_Inf_iff mem_box algebra_simps divide_simps)
ultimately have "N > 0"
using of_nat_0_less_iff by fastforce
show ?thesis
proof (intro exI impI allI)
fix k assume "N \<le> k"
with \<open>0 < N\<close> have "k > 0"
by linarith
have xa_gt: "(x - a) \<bullet> i > ((b - a) \<bullet> i) / (real k)" if "i \<in> Basis" for i
proof -
have *: "Inf ?\<Delta> \<le> ((x - a) \<bullet> i) / ((b - a) \<bullet> i)"
using that by (force intro: cInf_le_finite)
have "1 / Inf ?\<Delta> \<ge> ((b - a) \<bullet> i) / ((x - a) \<bullet> i)"
using le_imp_inverse_le [OF * \<Delta>]
by (simp add: field_simps)
with N have "k > ((b - a) \<bullet> i) / ((x - a) \<bullet> i)"
using \<open>N \<le> k\<close> by linarith
with x that show ?thesis
by (auto simp: mem_box algebra_simps field_split_simps)
qed
have bx_gt: "(b - x) \<bullet> i > ((b - a) \<bullet> i) / k" if "i \<in> Basis" for i
proof -
have *: "Inf ?\<Delta> \<le> ((b - x) \<bullet> i) / ((b - a) \<bullet> i)"
using that by (force intro: cInf_le_finite)
have "1 / Inf ?\<Delta> \<ge> ((b - a) \<bullet> i) / ((b - x) \<bullet> i)"
using le_imp_inverse_le [OF * \<Delta>]
by (simp add: field_simps)
with N have "k > ((b - a) \<bullet> i) / ((b - x) \<bullet> i)"
using \<open>N \<le> k\<close> by linarith
with x that show ?thesis
by (auto simp: mem_box algebra_simps field_split_simps)
qed
show "x \<in> cbox (a + (b - a) /\<^sub>R k) (b - (b - a) /\<^sub>R k)"
using that \<Delta> \<open>k > 0\<close>
by (auto simp: mem_box algebra_simps divide_inverse dest: xa_gt bx_gt)
qed
qed
obtain Bf where "Bf > 0" and Bf: "\<And>c d. cbox c d \<subseteq> box a b \<Longrightarrow> norm (integral (cbox c d) f) \<le> Bf"
using bo unfolding bounded_pos by blast
obtain Bg where "Bg > 0" and Bg:"\<And>c d. cbox c d \<subseteq> cbox a b \<Longrightarrow> \<bar>integral (cbox c d) g\<bar> \<le> Bg"
using bounded_integrals_over_subintervals [OF int_gab] unfolding bounded_pos real_norm_def by blast
show "(\<lambda>x. f x \<bullet> j) absolutely_integrable_on cbox a b"
using g
proof \<comment> \<open>A lot of duplication in the two proofs\<close>
assume fg [rule_format]: "\<forall>x\<in>cbox a b. f x \<bullet> j \<le> g x"
have "(\<lambda>x. (f x \<bullet> j)) = (\<lambda>x. g x - (g x - (f x \<bullet> j)))"
by simp
moreover have "(\<lambda>x. g x - (g x - (f x \<bullet> j))) integrable_on cbox a b"
proof (rule Henstock_Kurzweil_Integration.integrable_diff [OF int_gab])
let ?\<phi> = "\<lambda>k x. if x \<in> cbox (a + (b - a) /\<^sub>R (Suc k)) (b - (b - a) /\<^sub>R (Suc k))
then g x - f x \<bullet> j else 0"
have "(\<lambda>x. g x - f x \<bullet> j) integrable_on box a b"
proof (rule monotone_convergence_increasing [of ?\<phi>, THEN conjunct1])
have *: "cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k)) \<inter> box a b
= cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k))" for k
using box_subset_cbox "1" by fastforce
show "?\<phi> k integrable_on box a b" for k
apply (simp add: integrable_restrict_Int integral_restrict_Int *)
apply (rule integrable_diff [OF integrable_on_subcbox [OF int_gab]])
using "*" box_subset_cbox apply blast
by (metis "1" int_f integrable_component of_nat_Suc)
have cb12: "cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k))
\<subseteq> cbox (a + (b - a) /\<^sub>R (2 + real k)) (b - (b - a) /\<^sub>R (2 + real k))" for k
using False content_eq_0
apply (simp add: subset_box algebra_simps)
apply (simp add: divide_simps)
apply (fastforce simp: field_simps)
done
show "?\<phi> k x \<le> ?\<phi> (Suc k) x" if "x \<in> box a b" for k x
using cb12 box_subset_cbox that by (force simp: intro!: fg)
show "(\<lambda>k. ?\<phi> k x) \<longlonglongrightarrow> g x - f x \<bullet> j" if x: "x \<in> box a b" for x
proof (rule tendsto_eventually)
obtain N::nat where N: "\<And>k. k \<ge> N \<Longrightarrow> x \<in> cbox (a + (b - a) /\<^sub>R real k) (b - (b - a) /\<^sub>R real k)"
using getN [OF x] by blast
show "\<forall>\<^sub>F k in sequentially. ?\<phi> k x = g x - f x \<bullet> j"
proof
fix k::nat assume "N \<le> k"
have "x \<in> cbox (a + (b - a) /\<^sub>R (Suc k)) (b - (b - a) /\<^sub>R (Suc k))"
by (metis \<open>N \<le> k\<close> le_Suc_eq N)
then show "?\<phi> k x = g x - f x \<bullet> j"
by simp
qed
qed
have "\<bar>integral (box a b)
(\<lambda>x. if x \<in> cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k))
then g x - f x \<bullet> j else 0)\<bar> \<le> Bg + Bf" for k
proof -
let ?I = "cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k))"
have I_int [simp]: "?I \<inter> box a b = ?I"
using 1 by (simp add: Int_absorb2)
have int_fI: "f integrable_on ?I"
by (metis I_int inf_le2 int_f)
then have "(\<lambda>x. f x \<bullet> j) integrable_on ?I"
by (simp add: integrable_component)
moreover have "g integrable_on ?I"
by (metis I_int inf.orderI int_gab integrable_on_open_interval integrable_on_subcbox)
moreover
have "\<bar>integral ?I (\<lambda>x. f x \<bullet> j)\<bar> \<le> norm (integral ?I f)"
by (simp add: Basis_le_norm int_fI \<open>j \<in> Basis\<close>)
with 1 I_int have "\<bar>integral ?I (\<lambda>x. f x \<bullet> j)\<bar> \<le> Bf"
by (blast intro: order_trans [OF _ Bf])
ultimately show ?thesis
apply (simp add: integral_restrict_Int integral_diff)
using "*" box_subset_cbox by (blast intro: Bg add_mono order_trans [OF abs_triangle_ineq4])
qed
then show "bounded (range (\<lambda>k. integral (box a b) (?\<phi> k)))"
apply (simp add: bounded_pos)
apply (rule_tac x="Bg+Bf" in exI)
using \<open>0 < Bf\<close> \<open>0 < Bg\<close> apply auto
done
qed
then show "(\<lambda>x. g x - f x \<bullet> j) integrable_on cbox a b"
by (simp add: integrable_on_open_interval)
qed
ultimately have "(\<lambda>x. f x \<bullet> j) integrable_on cbox a b"
by auto
then show ?thesis
apply (rule absolutely_integrable_component_ubound [OF _ absint_g])
by (simp add: fg)
next
assume gf [rule_format]: "\<forall>x\<in>cbox a b. g x \<le> f x \<bullet> j"
have "(\<lambda>x. (f x \<bullet> j)) = (\<lambda>x. ((f x \<bullet> j) - g x) + g x)"
by simp
moreover have "(\<lambda>x. (f x \<bullet> j - g x) + g x) integrable_on cbox a b"
proof (rule Henstock_Kurzweil_Integration.integrable_add [OF _ int_gab])
let ?\<phi> = "\<lambda>k x. if x \<in> cbox (a + (b - a) /\<^sub>R (Suc k)) (b - (b - a) /\<^sub>R (Suc k))
then f x \<bullet> j - g x else 0"
have "(\<lambda>x. f x \<bullet> j - g x) integrable_on box a b"
proof (rule monotone_convergence_increasing [of ?\<phi>, THEN conjunct1])
have *: "cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k)) \<inter> box a b
= cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k))" for k
using box_subset_cbox "1" by fastforce
show "?\<phi> k integrable_on box a b" for k
apply (simp add: integrable_restrict_Int integral_restrict_Int *)
apply (rule integrable_diff)
apply (metis "1" int_f integrable_component of_nat_Suc)
apply (rule integrable_on_subcbox [OF int_gab])
using "*" box_subset_cbox apply blast
done
have cb12: "cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k))
\<subseteq> cbox (a + (b - a) /\<^sub>R (2 + real k)) (b - (b - a) /\<^sub>R (2 + real k))" for k
using False content_eq_0
apply (simp add: subset_box algebra_simps)
apply (simp add: divide_simps)
apply (fastforce simp: field_simps)
done
show "?\<phi> k x \<le> ?\<phi> (Suc k) x" if "x \<in> box a b" for k x
using cb12 box_subset_cbox that by (force simp: intro!: gf)
show "(\<lambda>k. ?\<phi> k x) \<longlonglongrightarrow> f x \<bullet> j - g x" if x: "x \<in> box a b" for x
proof (rule tendsto_eventually)
obtain N::nat where N: "\<And>k. k \<ge> N \<Longrightarrow> x \<in> cbox (a + (b - a) /\<^sub>R real k) (b - (b - a) /\<^sub>R real k)"
using getN [OF x] by blast
show "\<forall>\<^sub>F k in sequentially. ?\<phi> k x = f x \<bullet> j - g x"
proof
fix k::nat assume "N \<le> k"
have "x \<in> cbox (a + (b - a) /\<^sub>R (Suc k)) (b - (b - a) /\<^sub>R (Suc k))"
by (metis \<open>N \<le> k\<close> le_Suc_eq N)
then show "?\<phi> k x = f x \<bullet> j - g x"
by simp
qed
qed
have "\<bar>integral (box a b)
(\<lambda>x. if x \<in> cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k))
then f x \<bullet> j - g x else 0)\<bar> \<le> Bf + Bg" for k
proof -
let ?I = "cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k))"
have I_int [simp]: "?I \<inter> box a b = ?I"
using 1 by (simp add: Int_absorb2)
have int_fI: "f integrable_on ?I"
by (simp add: inf.orderI int_f)
then have "(\<lambda>x. f x \<bullet> j) integrable_on ?I"
by (simp add: integrable_component)
moreover have "g integrable_on ?I"
by (metis I_int inf.orderI int_gab integrable_on_open_interval integrable_on_subcbox)
moreover
have "\<bar>integral ?I (\<lambda>x. f x \<bullet> j)\<bar> \<le> norm (integral ?I f)"
by (simp add: Basis_le_norm int_fI \<open>j \<in> Basis\<close>)
with 1 I_int have "\<bar>integral ?I (\<lambda>x. f x \<bullet> j)\<bar> \<le> Bf"
by (blast intro: order_trans [OF _ Bf])
ultimately show ?thesis
apply (simp add: integral_restrict_Int integral_diff)
using "*" box_subset_cbox by (blast intro: Bg add_mono order_trans [OF abs_triangle_ineq4])
qed
then show "bounded (range (\<lambda>k. integral (box a b) (?\<phi> k)))"
apply (simp add: bounded_pos)
apply (rule_tac x="Bf+Bg" in exI)
using \<open>0 < Bf\<close> \<open>0 < Bg\<close> by auto
qed
then show "(\<lambda>x. f x \<bullet> j - g x) integrable_on cbox a b"
by (simp add: integrable_on_open_interval)
qed
ultimately have "(\<lambda>x. f x \<bullet> j) integrable_on cbox a b"
by auto
then show ?thesis
apply (rule absolutely_integrable_component_lbound [OF absint_g])
by (simp add: gf)
qed
qed
qed
subsection\<open>Second Mean Value Theorem\<close>
subsection\<open>Second mean value theorem and corollaries\<close>
lemma level_approx:
fixes f :: "real \<Rightarrow> real" and n::nat
assumes f: "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f x \<and> f x \<le> 1" and "x \<in> S" "n \<noteq> 0"
shows "\<bar>f x - (\<Sum>k = Suc 0..n. if k / n \<le> f x then inverse n else 0)\<bar> < inverse n"
(is "?lhs < _")
proof -
have "n * f x \<ge> 0"
using assms by auto
then obtain m::nat where m: "floor(n * f x) = int m"
using nonneg_int_cases zero_le_floor by blast
then have kn: "real k / real n \<le> f x \<longleftrightarrow> k \<le> m" for k
- using \<open>n \<noteq> 0\<close> by (simp add: field_split_simps mult.commute) linarith
+ using \<open>n \<noteq> 0\<close> by (simp add: field_split_simps) linarith
then have "Suc n / real n \<le> f x \<longleftrightarrow> Suc n \<le> m"
by blast
have "real n * f x \<le> real n"
by (simp add: \<open>x \<in> S\<close> f mult_left_le)
then have "m \<le> n"
using m by linarith
have "?lhs = \<bar>f x - (\<Sum>k \<in> {Suc 0..n} \<inter> {..m}. inverse n)\<bar>"
by (subst sum.inter_restrict) (auto simp: kn)
also have "\<dots> < inverse n"
using \<open>m \<le> n\<close> \<open>n \<noteq> 0\<close> m
- by (simp add: min_absorb2 field_split_simps mult.commute) linarith
+ by (simp add: min_absorb2 field_split_simps) linarith
finally show ?thesis .
qed
lemma SMVT_lemma2:
fixes f :: "real \<Rightarrow> real"
assumes f: "f integrable_on {a..b}"
and g: "\<And>x y. x \<le> y \<Longrightarrow> g x \<le> g y"
shows "(\<Union>y::real. {\<lambda>x. if g x \<ge> y then f x else 0}) equiintegrable_on {a..b}"
proof -
have ffab: "{f} equiintegrable_on {a..b}"
by (metis equiintegrable_on_sing f interval_cbox)
then have ff: "{f} equiintegrable_on (cbox a b)"
by simp
have ge: "(\<Union>c. {\<lambda>x. if x \<ge> c then f x else 0}) equiintegrable_on {a..b}"
using equiintegrable_halfspace_restrictions_ge [OF ff] by auto
have gt: "(\<Union>c. {\<lambda>x. if x > c then f x else 0}) equiintegrable_on {a..b}"
using equiintegrable_halfspace_restrictions_gt [OF ff] by auto
have 0: "{(\<lambda>x. 0)} equiintegrable_on {a..b}"
by (metis box_real(2) equiintegrable_on_sing integrable_0)
have \<dagger>: "(\<lambda>x. if g x \<ge> y then f x else 0) \<in> {(\<lambda>x. 0), f} \<union> (\<Union>z. {\<lambda>x. if z < x then f x else 0}) \<union> (\<Union>z. {\<lambda>x. if z \<le> x then f x else 0})"
for y
proof (cases "(\<forall>x. g x \<ge> y) \<or> (\<forall>x. \<not> (g x \<ge> y))")
let ?\<mu> = "Inf {x. g x \<ge> y}"
case False
have lower: "?\<mu> \<le> x" if "g x \<ge> y" for x
apply (rule cInf_lower)
using False
apply (auto simp: that bdd_below_def)
by (meson dual_order.trans g linear)
have greatest: "?\<mu> \<ge> z" if "(\<And>x. g x \<ge> y \<Longrightarrow> z \<le> x)" for z
by (metis False cInf_greatest empty_iff mem_Collect_eq that)
show ?thesis
proof (cases "g ?\<mu> \<ge> y")
case True
then obtain \<zeta> where \<zeta>: "\<And>x. g x \<ge> y \<longleftrightarrow> x \<ge> \<zeta>"
by (metis g lower order.trans) \<comment> \<open>in fact y is @{term ?\<mu>}\<close>
then show ?thesis
by (force simp: \<zeta>)
next
case False
have "(y \<le> g x) = (?\<mu> < x)" for x
proof
show "?\<mu> < x" if "y \<le> g x"
using that False less_eq_real_def lower by blast
show "y \<le> g x" if "?\<mu> < x"
by (metis g greatest le_less_trans that less_le_trans linear not_less)
qed
then obtain \<zeta> where \<zeta>: "\<And>x. g x \<ge> y \<longleftrightarrow> x > \<zeta>" ..
then show ?thesis
by (force simp: \<zeta>)
qed
qed auto
show ?thesis
apply (rule equiintegrable_on_subset [OF equiintegrable_on_Un [OF gt equiintegrable_on_Un [OF ge equiintegrable_on_Un [OF ffab 0]]]])
using \<dagger> apply (simp add: UN_subset_iff)
done
qed
lemma SMVT_lemma4:
fixes f :: "real \<Rightarrow> real"
assumes f: "f integrable_on {a..b}"
and "a \<le> b"
and g: "\<And>x y. x \<le> y \<Longrightarrow> g x \<le> g y"
and 01: "\<And>x. \<lbrakk>a \<le> x; x \<le> b\<rbrakk> \<Longrightarrow> 0 \<le> g x \<and> g x \<le> 1"
obtains c where "a \<le> c" "c \<le> b" "((\<lambda>x. g x *\<^sub>R f x) has_integral integral {c..b} f) {a..b}"
proof -
have "connected ((\<lambda>x. integral {x..b} f) ` {a..b})"
by (simp add: f indefinite_integral_continuous_1' connected_continuous_image)
moreover have "compact ((\<lambda>x. integral {x..b} f) ` {a..b})"
by (simp add: compact_continuous_image f indefinite_integral_continuous_1')
ultimately obtain m M where int_fab: "(\<lambda>x. integral {x..b} f) ` {a..b} = {m..M}"
using connected_compact_interval_1 by meson
have "\<exists>c. c \<in> {a..b} \<and>
integral {c..b} f =
integral {a..b} (\<lambda>x. (\<Sum>k = 1..n. if g x \<ge> real k / real n then inverse n *\<^sub>R f x else 0))" for n
proof (cases "n=0")
case True
then show ?thesis
using \<open>a \<le> b\<close> by auto
next
case False
have "(\<Union>c::real. {\<lambda>x. if g x \<ge> c then f x else 0}) equiintegrable_on {a..b}"
using SMVT_lemma2 [OF f g] .
then have int: "(\<lambda>x. if g x \<ge> c then f x else 0) integrable_on {a..b}" for c
by (simp add: equiintegrable_on_def)
have int': "(\<lambda>x. if g x \<ge> c then u * f x else 0) integrable_on {a..b}" for c u
proof -
have "(\<lambda>x. if g x \<ge> c then u * f x else 0) = (\<lambda>x. u * (if g x \<ge> c then f x else 0))"
by (force simp: if_distrib)
then show ?thesis
using integrable_on_cmult_left [OF int] by simp
qed
have "\<exists>d. d \<in> {a..b} \<and> integral {a..b} (\<lambda>x. if g x \<ge> y then f x else 0) = integral {d..b} f" for y
proof -
let ?X = "{x. g x \<ge> y}"
have *: "\<exists>a. ?X = {x. a \<le> x} \<or> ?X = {x. a < x}"
if 1: "?X \<noteq> {}" and 2: "?X \<noteq> UNIV"
proof -
let ?\<mu> = "Inf{x. g x \<ge> y}"
have lower: "?\<mu> \<le> x" if "g x \<ge> y" for x
apply (rule cInf_lower)
using 1 2 apply (auto simp: that bdd_below_def)
by (meson dual_order.trans g linear)
have greatest: "?\<mu> \<ge> z" if "(\<And>x. g x \<ge> y \<Longrightarrow> z \<le> x)" for z
by (metis cInf_greatest mem_Collect_eq that 1)
show ?thesis
proof (cases "g ?\<mu> \<ge> y")
case True
then obtain \<zeta> where \<zeta>: "\<And>x. g x \<ge> y \<longleftrightarrow> x \<ge> \<zeta>"
by (metis g lower order.trans) \<comment> \<open>in fact y is @{term ?\<mu>}\<close>
then show ?thesis
by (force simp: \<zeta>)
next
case False
have "(y \<le> g x) = (?\<mu> < x)" for x
proof
show "?\<mu> < x" if "y \<le> g x"
using that False less_eq_real_def lower by blast
show "y \<le> g x" if "?\<mu> < x"
by (metis g greatest le_less_trans that less_le_trans linear not_less)
qed
then obtain \<zeta> where \<zeta>: "\<And>x. g x \<ge> y \<longleftrightarrow> x > \<zeta>" ..
then show ?thesis
by (force simp: \<zeta>)
qed
qed
then consider "?X = {}" | "?X = UNIV" | d where "?X = {x. d \<le> x} \<or>?X = {x. d < x}"
by metis
then have "\<exists>d. d \<in> {a..b} \<and> integral {a..b} (\<lambda>x. if x \<in> ?X then f x else 0) = integral {d..b} f"
proof cases
case 1
then show ?thesis
using \<open>a \<le> b\<close> by auto
next
case 2
then show ?thesis
using \<open>a \<le> b\<close> by auto
next
case (3 d)
show ?thesis
proof (cases "d < a")
case True
with 3 show ?thesis
apply (rule_tac x=a in exI)
apply (auto simp: \<open>a \<le> b\<close> iff del: mem_Collect_eq intro!: Henstock_Kurzweil_Integration.integral_cong, simp_all)
done
next
case False
show ?thesis
proof (cases "b < d")
case True
have "integral {a..b} (\<lambda>x. if x \<in> {x. y \<le> g x} then f x else 0) = integral {a..b} (\<lambda>x. 0)"
by (rule Henstock_Kurzweil_Integration.integral_cong) (use 3 True in fastforce)
then show ?thesis
using \<open>a \<le> b\<close> by auto
next
case False
with \<open>\<not> d < a\<close> have eq: "Collect ((\<le>) d) \<inter> {a..b} = {d..b}" "Collect ((<) d) \<inter> {a..b} = {d<..b}"
by force+
moreover have "integral {d<..b} f = integral {d..b} f"
by (rule integral_spike_set [OF empty_imp_negligible negligible_subset [OF negligible_sing [of d]]]) auto
ultimately
show ?thesis
using \<open>\<not> d < a\<close> False 3
apply (rule_tac x=d in exI)
apply (auto simp: \<open>a \<le> b\<close> iff del: mem_Collect_eq; subst Henstock_Kurzweil_Integration.integral_restrict_Int)
apply (simp_all add: \<open>a \<le> b\<close> not_less eq)
done
qed
qed
qed
then show ?thesis
by auto
qed
then have "\<forall>k. \<exists>d. d \<in> {a..b} \<and> integral {a..b} (\<lambda>x. if real k / real n \<le> g x then f x else 0) = integral {d..b} f"
by meson
then obtain d where dab: "\<And>k. d k \<in> {a..b}"
and deq: "\<And>k::nat. integral {a..b} (\<lambda>x. if k/n \<le> g x then f x else 0) = integral {d k..b} f"
by metis
have "(\<Sum>k = 1..n. integral {a..b} (\<lambda>x. if real k / real n \<le> g x then f x else 0)) /\<^sub>R n \<in> {m..M}"
unfolding scaleR_right.sum
proof (intro conjI allI impI convex [THEN iffD1, rule_format])
show "integral {a..b} (\<lambda>xa. if real k / real n \<le> g xa then f xa else 0) \<in> {m..M}" for k
by (metis (no_types, lifting) deq image_eqI int_fab dab)
qed (use \<open>n \<noteq> 0\<close> in auto)
then have "\<exists>c. c \<in> {a..b} \<and>
integral {c..b} f = inverse n *\<^sub>R (\<Sum>k = 1..n. integral {a..b} (\<lambda>x. if g x \<ge> real k / real n then f x else 0))"
by (metis (no_types, lifting) int_fab imageE)
then show ?thesis
by (simp add: sum_distrib_left if_distrib integral_sum int' flip: integral_mult_right cong: if_cong)
qed
then obtain c where cab: "\<And>n. c n \<in> {a..b}"
and c: "\<And>n. integral {c n..b} f = integral {a..b} (\<lambda>x. (\<Sum>k = 1..n. if g x \<ge> real k / real n then f x /\<^sub>R n else 0))"
by metis
obtain d and \<sigma> :: "nat\<Rightarrow>nat"
where "d \<in> {a..b}" and \<sigma>: "strict_mono \<sigma>" and d: "(c \<circ> \<sigma>) \<longlonglongrightarrow> d" and non0: "\<And>n. \<sigma> n \<ge> Suc 0"
proof -
have "compact{a..b}"
by auto
with cab obtain d and s0
where "d \<in> {a..b}" and s0: "strict_mono s0" and tends: "(c \<circ> s0) \<longlonglongrightarrow> d"
unfolding compact_def
using that by blast
show thesis
proof
show "d \<in> {a..b}"
by fact
show "strict_mono (s0 \<circ> Suc)"
using s0 by (auto simp: strict_mono_def)
show "(c \<circ> (s0 \<circ> Suc)) \<longlonglongrightarrow> d"
by (metis tends LIMSEQ_subseq_LIMSEQ Suc_less_eq comp_assoc strict_mono_def)
show "\<And>n. (s0 \<circ> Suc) n \<ge> Suc 0"
by (metis comp_apply le0 not_less_eq_eq old.nat.exhaust s0 seq_suble)
qed
qed
define \<phi> where "\<phi> \<equiv> \<lambda>n x. \<Sum>k = Suc 0..\<sigma> n. if k/(\<sigma> n) \<le> g x then f x /\<^sub>R (\<sigma> n) else 0"
define \<psi> where "\<psi> \<equiv> \<lambda>n x. \<Sum>k = Suc 0..\<sigma> n. if k/(\<sigma> n) \<le> g x then inverse (\<sigma> n) else 0"
have **: "(\<lambda>x. g x *\<^sub>R f x) integrable_on cbox a b \<and>
(\<lambda>n. integral (cbox a b) (\<phi> n)) \<longlonglongrightarrow> integral (cbox a b) (\<lambda>x. g x *\<^sub>R f x)"
proof (rule equiintegrable_limit)
have \<dagger>: "((\<lambda>n. \<lambda>x. (\<Sum>k = Suc 0..n. if k / n \<le> g x then inverse n *\<^sub>R f x else 0)) ` {Suc 0..}) equiintegrable_on {a..b}"
proof -
have *: "(\<Union>c::real. {\<lambda>x. if g x \<ge> c then f x else 0}) equiintegrable_on {a..b}"
using SMVT_lemma2 [OF f g] .
show ?thesis
apply (rule equiintegrable_on_subset [OF equiintegrable_sum_real [OF *]], clarify)
apply (rule_tac a="{Suc 0..n}" in UN_I, force)
apply (rule_tac a="\<lambda>k. inverse n" in UN_I, auto)
apply (rule_tac x="\<lambda>k x. if real k / real n \<le> g x then f x else 0" in bexI)
apply (force intro: sum.cong)+
done
qed
show "range \<phi> equiintegrable_on cbox a b"
unfolding \<phi>_def
by (auto simp: non0 intro: equiintegrable_on_subset [OF \<dagger>])
show "(\<lambda>n. \<phi> n x) \<longlonglongrightarrow> g x *\<^sub>R f x"
if x: "x \<in> cbox a b" for x
proof -
have eq: "\<phi> n x = \<psi> n x *\<^sub>R f x" for n
by (auto simp: \<phi>_def \<psi>_def sum_distrib_right if_distrib intro: sum.cong)
show ?thesis
unfolding eq
proof (rule tendsto_scaleR [OF _ tendsto_const])
show "(\<lambda>n. \<psi> n x) \<longlonglongrightarrow> g x"
unfolding lim_sequentially dist_real_def
proof (intro allI impI)
fix e :: real
assume "e > 0"
then obtain N where "N \<noteq> 0" "0 < inverse (real N)" and N: "inverse (real N) < e"
using real_arch_inverse by metis
moreover have "\<bar>\<psi> n x - g x\<bar> < inverse (real N)" if "n\<ge>N" for n
proof -
have "\<bar>g x - \<psi> n x\<bar> < inverse (real (\<sigma> n))"
unfolding \<psi>_def
proof (rule level_approx [of "{a..b}" g])
show "\<sigma> n \<noteq> 0"
by (metis Suc_n_not_le_n non0)
qed (use x 01 non0 in auto)
also have "\<dots> \<le> inverse N"
using seq_suble [OF \<sigma>] \<open>N \<noteq> 0\<close> non0 that by (auto intro: order_trans simp: field_split_simps)
finally show ?thesis
by linarith
qed
ultimately show "\<exists>N. \<forall>n\<ge>N. \<bar>\<psi> n x - g x\<bar> < e"
using less_trans by blast
qed
qed
qed
qed
show thesis
proof
show "a \<le> d" "d \<le> b"
using \<open>d \<in> {a..b}\<close> atLeastAtMost_iff by blast+
show "((\<lambda>x. g x *\<^sub>R f x) has_integral integral {d..b} f) {a..b}"
unfolding has_integral_iff
proof
show "(\<lambda>x. g x *\<^sub>R f x) integrable_on {a..b}"
using ** by simp
show "integral {a..b} (\<lambda>x. g x *\<^sub>R f x) = integral {d..b} f"
proof (rule tendsto_unique)
show "(\<lambda>n. integral {c(\<sigma> n)..b} f) \<longlonglongrightarrow> integral {a..b} (\<lambda>x. g x *\<^sub>R f x)"
using ** by (simp add: c \<phi>_def)
show "(\<lambda>n. integral {c(\<sigma> n)..b} f) \<longlonglongrightarrow> integral {d..b} f"
using indefinite_integral_continuous_1' [OF f]
using d unfolding o_def
apply (simp add: continuous_on_eq_continuous_within)
apply (drule_tac x=d in bspec)
using \<open>d \<in> {a..b}\<close> apply blast
apply (simp add: continuous_within_sequentially o_def)
using cab by auto
qed auto
qed
qed
qed
theorem second_mean_value_theorem_full:
fixes f :: "real \<Rightarrow> real"
assumes f: "f integrable_on {a..b}" and "a \<le> b"
and g: "\<And>x y. \<lbrakk>a \<le> x; x \<le> y; y \<le> b\<rbrakk> \<Longrightarrow> g x \<le> g y"
obtains c where "c \<in> {a..b}"
and "((\<lambda>x. g x * f x) has_integral (g a * integral {a..c} f + g b * integral {c..b} f)) {a..b}"
proof -
have gab: "g a \<le> g b"
using \<open>a \<le> b\<close> g by blast
then consider "g a < g b" | "g a = g b"
by linarith
then show thesis
proof cases
case 1
define h where "h \<equiv> \<lambda>x. if x < a then 0 else if b < x then 1
else (g x - g a) / (g b - g a)"
obtain c where "a \<le> c" "c \<le> b" and c: "((\<lambda>x. h x *\<^sub>R f x) has_integral integral {c..b} f) {a..b}"
proof (rule SMVT_lemma4 [OF f \<open>a \<le> b\<close>, of h])
show "h x \<le> h y" "0 \<le> h x \<and> h x \<le> 1" if "x \<le> y" for x y
using that gab by (auto simp: divide_simps g h_def)
qed
show ?thesis
proof
show "c \<in> {a..b}"
using \<open>a \<le> c\<close> \<open>c \<le> b\<close> by auto
have I: "((\<lambda>x. g x * f x - g a * f x) has_integral (g b - g a) * integral {c..b} f) {a..b}"
proof (subst has_integral_cong)
show "g x * f x - g a * f x = (g b - g a) * h x *\<^sub>R f x"
if "x \<in> {a..b}" for x
- using 1 that by (simp add: h_def field_split_simps algebra_simps)
+ using 1 that by (simp add: h_def field_split_simps)
show "((\<lambda>x. (g b - g a) * h x *\<^sub>R f x) has_integral (g b - g a) * integral {c..b} f) {a..b}"
using has_integral_mult_right [OF c, of "g b - g a"] .
qed
have II: "((\<lambda>x. g a * f x) has_integral g a * integral {a..b} f) {a..b}"
using has_integral_mult_right [where c = "g a", OF integrable_integral [OF f]] .
have "((\<lambda>x. g x * f x) has_integral (g b - g a) * integral {c..b} f + g a * integral {a..b} f) {a..b}"
using has_integral_add [OF I II] by simp
then show "((\<lambda>x. g x * f x) has_integral g a * integral {a..c} f + g b * integral {c..b} f) {a..b}"
by (simp add: algebra_simps flip: integral_combine [OF \<open>a \<le> c\<close> \<open>c \<le> b\<close> f])
qed
next
case 2
show ?thesis
proof
show "a \<in> {a..b}"
by (simp add: \<open>a \<le> b\<close>)
have "((\<lambda>x. g x * f x) has_integral g a * integral {a..b} f) {a..b}"
proof (rule has_integral_eq)
show "((\<lambda>x. g a * f x) has_integral g a * integral {a..b} f) {a..b}"
using f has_integral_mult_right by blast
show "g a * f x = g x * f x"
if "x \<in> {a..b}" for x
by (metis atLeastAtMost_iff g less_eq_real_def not_le that 2)
qed
then show "((\<lambda>x. g x * f x) has_integral g a * integral {a..a} f + g b * integral {a..b} f) {a..b}"
by (simp add: 2)
qed
qed
qed
corollary second_mean_value_theorem:
fixes f :: "real \<Rightarrow> real"
assumes f: "f integrable_on {a..b}" and "a \<le> b"
and g: "\<And>x y. \<lbrakk>a \<le> x; x \<le> y; y \<le> b\<rbrakk> \<Longrightarrow> g x \<le> g y"
obtains c where "c \<in> {a..b}"
"integral {a..b} (\<lambda>x. g x * f x)
= g a * integral {a..c} f + g b * integral {c..b} f"
using second_mean_value_theorem_full [where g=g, OF assms]
by (metis (full_types) integral_unique)
end
diff --git a/src/HOL/Analysis/Infinite_Set_Sum.thy b/src/HOL/Analysis/Infinite_Set_Sum.thy
--- a/src/HOL/Analysis/Infinite_Set_Sum.thy
+++ b/src/HOL/Analysis/Infinite_Set_Sum.thy
@@ -1,868 +1,868 @@
(*
Title: HOL/Analysis/Infinite_Set_Sum.thy
Author: Manuel Eberl, TU München
A theory of sums over possible infinite sets. (Only works for absolute summability)
*)
section \<open>Sums over Infinite Sets\<close>
theory Infinite_Set_Sum
imports Set_Integral
begin
(* TODO Move *)
lemma sets_eq_countable:
assumes "countable A" "space M = A" "\<And>x. x \<in> A \<Longrightarrow> {x} \<in> sets M"
shows "sets M = Pow A"
proof (intro equalityI subsetI)
fix X assume "X \<in> Pow A"
hence "(\<Union>x\<in>X. {x}) \<in> sets M"
by (intro sets.countable_UN' countable_subset[OF _ assms(1)]) (auto intro!: assms(3))
also have "(\<Union>x\<in>X. {x}) = X" by auto
finally show "X \<in> sets M" .
next
fix X assume "X \<in> sets M"
from sets.sets_into_space[OF this] and assms
show "X \<in> Pow A" by simp
qed
lemma measure_eqI_countable':
assumes spaces: "space M = A" "space N = A"
assumes sets: "\<And>x. x \<in> A \<Longrightarrow> {x} \<in> sets M" "\<And>x. x \<in> A \<Longrightarrow> {x} \<in> sets N"
assumes A: "countable A"
assumes eq: "\<And>a. a \<in> A \<Longrightarrow> emeasure M {a} = emeasure N {a}"
shows "M = N"
proof (rule measure_eqI_countable)
show "sets M = Pow A"
by (intro sets_eq_countable assms)
show "sets N = Pow A"
by (intro sets_eq_countable assms)
qed fact+
lemma count_space_PiM_finite:
fixes B :: "'a \<Rightarrow> 'b set"
assumes "finite A" "\<And>i. countable (B i)"
shows "PiM A (\<lambda>i. count_space (B i)) = count_space (PiE A B)"
proof (rule measure_eqI_countable')
show "space (PiM A (\<lambda>i. count_space (B i))) = PiE A B"
by (simp add: space_PiM)
show "space (count_space (PiE A B)) = PiE A B" by simp
next
fix f assume f: "f \<in> PiE A B"
hence "PiE A (\<lambda>x. {f x}) \<in> sets (Pi\<^sub>M A (\<lambda>i. count_space (B i)))"
by (intro sets_PiM_I_finite assms) auto
also from f have "PiE A (\<lambda>x. {f x}) = {f}"
by (intro PiE_singleton) (auto simp: PiE_def)
finally show "{f} \<in> sets (Pi\<^sub>M A (\<lambda>i. count_space (B i)))" .
next
interpret product_sigma_finite "(\<lambda>i. count_space (B i))"
by (intro product_sigma_finite.intro sigma_finite_measure_count_space_countable assms)
thm sigma_finite_measure_count_space
fix f assume f: "f \<in> PiE A B"
hence "{f} = PiE A (\<lambda>x. {f x})"
by (intro PiE_singleton [symmetric]) (auto simp: PiE_def)
also have "emeasure (Pi\<^sub>M A (\<lambda>i. count_space (B i))) \<dots> =
(\<Prod>i\<in>A. emeasure (count_space (B i)) {f i})"
using f assms by (subst emeasure_PiM) auto
also have "\<dots> = (\<Prod>i\<in>A. 1)"
by (intro prod.cong refl, subst emeasure_count_space_finite) (use f in auto)
also have "\<dots> = emeasure (count_space (PiE A B)) {f}"
using f by (subst emeasure_count_space_finite) auto
finally show "emeasure (Pi\<^sub>M A (\<lambda>i. count_space (B i))) {f} =
emeasure (count_space (Pi\<^sub>E A B)) {f}" .
qed (simp_all add: countable_PiE assms)
definition\<^marker>\<open>tag important\<close> abs_summable_on ::
"('a \<Rightarrow> 'b :: {banach, second_countable_topology}) \<Rightarrow> 'a set \<Rightarrow> bool"
(infix "abs'_summable'_on" 50)
where
"f abs_summable_on A \<longleftrightarrow> integrable (count_space A) f"
definition\<^marker>\<open>tag important\<close> infsetsum ::
"('a \<Rightarrow> 'b :: {banach, second_countable_topology}) \<Rightarrow> 'a set \<Rightarrow> 'b"
where
"infsetsum f A = lebesgue_integral (count_space A) f"
syntax (ASCII)
"_infsetsum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::{banach, second_countable_topology}"
("(3INFSETSUM _:_./ _)" [0, 51, 10] 10)
syntax
"_infsetsum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::{banach, second_countable_topology}"
("(2\<Sum>\<^sub>a_\<in>_./ _)" [0, 51, 10] 10)
translations \<comment> \<open>Beware of argument permutation!\<close>
"\<Sum>\<^sub>ai\<in>A. b" \<rightleftharpoons> "CONST infsetsum (\<lambda>i. b) A"
syntax (ASCII)
"_uinfsetsum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::{banach, second_countable_topology}"
("(3INFSETSUM _:_./ _)" [0, 51, 10] 10)
syntax
"_uinfsetsum" :: "pttrn \<Rightarrow> 'b \<Rightarrow> 'b::{banach, second_countable_topology}"
("(2\<Sum>\<^sub>a_./ _)" [0, 10] 10)
translations \<comment> \<open>Beware of argument permutation!\<close>
"\<Sum>\<^sub>ai. b" \<rightleftharpoons> "CONST infsetsum (\<lambda>i. b) (CONST UNIV)"
syntax (ASCII)
"_qinfsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a::{banach, second_countable_topology}"
("(3INFSETSUM _ |/ _./ _)" [0, 0, 10] 10)
syntax
"_qinfsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a::{banach, second_countable_topology}"
("(2\<Sum>\<^sub>a_ | (_)./ _)" [0, 0, 10] 10)
translations
"\<Sum>\<^sub>ax|P. t" => "CONST infsetsum (\<lambda>x. t) {x. P}"
print_translation \<open>
let
fun sum_tr' [Abs (x, Tx, t), Const (\<^const_syntax>\<open>Collect\<close>, _) $ Abs (y, Ty, P)] =
if x <> y then raise Match
else
let
val x' = Syntax_Trans.mark_bound_body (x, Tx);
val t' = subst_bound (x', t);
val P' = subst_bound (x', P);
in
Syntax.const \<^syntax_const>\<open>_qinfsetsum\<close> $ Syntax_Trans.mark_bound_abs (x, Tx) $ P' $ t'
end
| sum_tr' _ = raise Match;
in [(\<^const_syntax>\<open>infsetsum\<close>, K sum_tr')] end
\<close>
lemma restrict_count_space_subset:
"A \<subseteq> B \<Longrightarrow> restrict_space (count_space B) A = count_space A"
by (subst restrict_count_space) (simp_all add: Int_absorb2)
lemma abs_summable_on_restrict:
fixes f :: "'a \<Rightarrow> 'b :: {banach, second_countable_topology}"
assumes "A \<subseteq> B"
shows "f abs_summable_on A \<longleftrightarrow> (\<lambda>x. indicator A x *\<^sub>R f x) abs_summable_on B"
proof -
have "count_space A = restrict_space (count_space B) A"
by (rule restrict_count_space_subset [symmetric]) fact+
also have "integrable \<dots> f \<longleftrightarrow> set_integrable (count_space B) A f"
by (simp add: integrable_restrict_space set_integrable_def)
finally show ?thesis
unfolding abs_summable_on_def set_integrable_def .
qed
lemma abs_summable_on_altdef: "f abs_summable_on A \<longleftrightarrow> set_integrable (count_space UNIV) A f"
unfolding abs_summable_on_def set_integrable_def
by (metis (no_types) inf_top.right_neutral integrable_restrict_space restrict_count_space sets_UNIV)
lemma abs_summable_on_altdef':
"A \<subseteq> B \<Longrightarrow> f abs_summable_on A \<longleftrightarrow> set_integrable (count_space B) A f"
unfolding abs_summable_on_def set_integrable_def
- by (metis (no_types) Pow_iff abs_summable_on_def inf.orderE integrable_restrict_space restrict_count_space_subset set_integrable_def sets_count_space space_count_space)
+ by (metis (no_types) Pow_iff abs_summable_on_def inf.orderE integrable_restrict_space restrict_count_space_subset sets_count_space space_count_space)
lemma abs_summable_on_norm_iff [simp]:
"(\<lambda>x. norm (f x)) abs_summable_on A \<longleftrightarrow> f abs_summable_on A"
by (simp add: abs_summable_on_def integrable_norm_iff)
lemma abs_summable_on_normI: "f abs_summable_on A \<Longrightarrow> (\<lambda>x. norm (f x)) abs_summable_on A"
by simp
lemma abs_summable_complex_of_real [simp]: "(\<lambda>n. complex_of_real (f n)) abs_summable_on A \<longleftrightarrow> f abs_summable_on A"
by (simp add: abs_summable_on_def complex_of_real_integrable_eq)
lemma abs_summable_on_comparison_test:
assumes "g abs_summable_on A"
assumes "\<And>x. x \<in> A \<Longrightarrow> norm (f x) \<le> norm (g x)"
shows "f abs_summable_on A"
using assms Bochner_Integration.integrable_bound[of "count_space A" g f]
unfolding abs_summable_on_def by (auto simp: AE_count_space)
lemma abs_summable_on_comparison_test':
assumes "g abs_summable_on A"
assumes "\<And>x. x \<in> A \<Longrightarrow> norm (f x) \<le> g x"
shows "f abs_summable_on A"
proof (rule abs_summable_on_comparison_test[OF assms(1), of f])
fix x assume "x \<in> A"
with assms(2) have "norm (f x) \<le> g x" .
also have "\<dots> \<le> norm (g x)" by simp
finally show "norm (f x) \<le> norm (g x)" .
qed
lemma abs_summable_on_cong [cong]:
"(\<And>x. x \<in> A \<Longrightarrow> f x = g x) \<Longrightarrow> A = B \<Longrightarrow> (f abs_summable_on A) \<longleftrightarrow> (g abs_summable_on B)"
unfolding abs_summable_on_def by (intro integrable_cong) auto
lemma abs_summable_on_cong_neutral:
assumes "\<And>x. x \<in> A - B \<Longrightarrow> f x = 0"
assumes "\<And>x. x \<in> B - A \<Longrightarrow> g x = 0"
assumes "\<And>x. x \<in> A \<inter> B \<Longrightarrow> f x = g x"
shows "f abs_summable_on A \<longleftrightarrow> g abs_summable_on B"
unfolding abs_summable_on_altdef set_integrable_def using assms
by (intro Bochner_Integration.integrable_cong refl)
(auto simp: indicator_def split: if_splits)
lemma abs_summable_on_restrict':
fixes f :: "'a \<Rightarrow> 'b :: {banach, second_countable_topology}"
assumes "A \<subseteq> B"
shows "f abs_summable_on A \<longleftrightarrow> (\<lambda>x. if x \<in> A then f x else 0) abs_summable_on B"
by (subst abs_summable_on_restrict[OF assms]) (intro abs_summable_on_cong, auto)
lemma abs_summable_on_nat_iff:
"f abs_summable_on (A :: nat set) \<longleftrightarrow> summable (\<lambda>n. if n \<in> A then norm (f n) else 0)"
proof -
have "f abs_summable_on A \<longleftrightarrow> summable (\<lambda>x. norm (if x \<in> A then f x else 0))"
by (subst abs_summable_on_restrict'[of _ UNIV])
(simp_all add: abs_summable_on_def integrable_count_space_nat_iff)
also have "(\<lambda>x. norm (if x \<in> A then f x else 0)) = (\<lambda>x. if x \<in> A then norm (f x) else 0)"
by auto
finally show ?thesis .
qed
lemma abs_summable_on_nat_iff':
"f abs_summable_on (UNIV :: nat set) \<longleftrightarrow> summable (\<lambda>n. norm (f n))"
by (subst abs_summable_on_nat_iff) auto
lemma nat_abs_summable_on_comparison_test:
fixes f :: "nat \<Rightarrow> 'a :: {banach, second_countable_topology}"
assumes "g abs_summable_on I"
assumes "\<And>n. \<lbrakk>n\<ge>N; n \<in> I\<rbrakk> \<Longrightarrow> norm (f n) \<le> g n"
shows "f abs_summable_on I"
using assms by (fastforce simp add: abs_summable_on_nat_iff intro: summable_comparison_test')
lemma abs_summable_comparison_test_ev:
assumes "g abs_summable_on I"
assumes "eventually (\<lambda>x. x \<in> I \<longrightarrow> norm (f x) \<le> g x) sequentially"
shows "f abs_summable_on I"
by (metis (no_types, lifting) nat_abs_summable_on_comparison_test eventually_at_top_linorder assms)
lemma abs_summable_on_Cauchy:
"f abs_summable_on (UNIV :: nat set) \<longleftrightarrow> (\<forall>e>0. \<exists>N. \<forall>m\<ge>N. \<forall>n. (\<Sum>x = m..<n. norm (f x)) < e)"
by (simp add: abs_summable_on_nat_iff' summable_Cauchy sum_nonneg)
lemma abs_summable_on_finite [simp]: "finite A \<Longrightarrow> f abs_summable_on A"
unfolding abs_summable_on_def by (rule integrable_count_space)
lemma abs_summable_on_empty [simp, intro]: "f abs_summable_on {}"
by simp
lemma abs_summable_on_subset:
assumes "f abs_summable_on B" and "A \<subseteq> B"
shows "f abs_summable_on A"
unfolding abs_summable_on_altdef
by (rule set_integrable_subset) (insert assms, auto simp: abs_summable_on_altdef)
lemma abs_summable_on_union [intro]:
assumes "f abs_summable_on A" and "f abs_summable_on B"
shows "f abs_summable_on (A \<union> B)"
using assms unfolding abs_summable_on_altdef by (intro set_integrable_Un) auto
lemma abs_summable_on_insert_iff [simp]:
"f abs_summable_on insert x A \<longleftrightarrow> f abs_summable_on A"
proof safe
assume "f abs_summable_on insert x A"
thus "f abs_summable_on A"
by (rule abs_summable_on_subset) auto
next
assume "f abs_summable_on A"
from abs_summable_on_union[OF this, of "{x}"]
show "f abs_summable_on insert x A" by simp
qed
lemma abs_summable_sum:
assumes "\<And>x. x \<in> A \<Longrightarrow> f x abs_summable_on B"
shows "(\<lambda>y. \<Sum>x\<in>A. f x y) abs_summable_on B"
using assms unfolding abs_summable_on_def by (intro Bochner_Integration.integrable_sum)
lemma abs_summable_Re: "f abs_summable_on A \<Longrightarrow> (\<lambda>x. Re (f x)) abs_summable_on A"
by (simp add: abs_summable_on_def)
lemma abs_summable_Im: "f abs_summable_on A \<Longrightarrow> (\<lambda>x. Im (f x)) abs_summable_on A"
by (simp add: abs_summable_on_def)
lemma abs_summable_on_finite_diff:
assumes "f abs_summable_on A" "A \<subseteq> B" "finite (B - A)"
shows "f abs_summable_on B"
proof -
have "f abs_summable_on (A \<union> (B - A))"
by (intro abs_summable_on_union assms abs_summable_on_finite)
also from assms have "A \<union> (B - A) = B" by blast
finally show ?thesis .
qed
lemma abs_summable_on_reindex_bij_betw:
assumes "bij_betw g A B"
shows "(\<lambda>x. f (g x)) abs_summable_on A \<longleftrightarrow> f abs_summable_on B"
proof -
have *: "count_space B = distr (count_space A) (count_space B) g"
by (rule distr_bij_count_space [symmetric]) fact
show ?thesis unfolding abs_summable_on_def
by (subst *, subst integrable_distr_eq[of _ _ "count_space B"])
(insert assms, auto simp: bij_betw_def)
qed
lemma abs_summable_on_reindex:
assumes "(\<lambda>x. f (g x)) abs_summable_on A"
shows "f abs_summable_on (g ` A)"
proof -
define g' where "g' = inv_into A g"
from assms have "(\<lambda>x. f (g x)) abs_summable_on (g' ` g ` A)"
by (rule abs_summable_on_subset) (auto simp: g'_def inv_into_into)
also have "?this \<longleftrightarrow> (\<lambda>x. f (g (g' x))) abs_summable_on (g ` A)" unfolding g'_def
by (intro abs_summable_on_reindex_bij_betw [symmetric] inj_on_imp_bij_betw inj_on_inv_into) auto
also have "\<dots> \<longleftrightarrow> f abs_summable_on (g ` A)"
by (intro abs_summable_on_cong refl) (auto simp: g'_def f_inv_into_f)
finally show ?thesis .
qed
lemma abs_summable_on_reindex_iff:
"inj_on g A \<Longrightarrow> (\<lambda>x. f (g x)) abs_summable_on A \<longleftrightarrow> f abs_summable_on (g ` A)"
by (intro abs_summable_on_reindex_bij_betw inj_on_imp_bij_betw)
lemma abs_summable_on_Sigma_project2:
fixes A :: "'a set" and B :: "'a \<Rightarrow> 'b set"
assumes "f abs_summable_on (Sigma A B)" "x \<in> A"
shows "(\<lambda>y. f (x, y)) abs_summable_on (B x)"
proof -
from assms(2) have "f abs_summable_on (Sigma {x} B)"
by (intro abs_summable_on_subset [OF assms(1)]) auto
also have "?this \<longleftrightarrow> (\<lambda>z. f (x, snd z)) abs_summable_on (Sigma {x} B)"
by (rule abs_summable_on_cong) auto
finally have "(\<lambda>y. f (x, y)) abs_summable_on (snd ` Sigma {x} B)"
by (rule abs_summable_on_reindex)
also have "snd ` Sigma {x} B = B x"
using assms by (auto simp: image_iff)
finally show ?thesis .
qed
lemma abs_summable_on_Times_swap:
"f abs_summable_on A \<times> B \<longleftrightarrow> (\<lambda>(x,y). f (y,x)) abs_summable_on B \<times> A"
proof -
have bij: "bij_betw (\<lambda>(x,y). (y,x)) (B \<times> A) (A \<times> B)"
by (auto simp: bij_betw_def inj_on_def)
show ?thesis
by (subst abs_summable_on_reindex_bij_betw[OF bij, of f, symmetric])
(simp_all add: case_prod_unfold)
qed
lemma abs_summable_on_0 [simp, intro]: "(\<lambda>_. 0) abs_summable_on A"
by (simp add: abs_summable_on_def)
lemma abs_summable_on_uminus [intro]:
"f abs_summable_on A \<Longrightarrow> (\<lambda>x. -f x) abs_summable_on A"
unfolding abs_summable_on_def by (rule Bochner_Integration.integrable_minus)
lemma abs_summable_on_add [intro]:
assumes "f abs_summable_on A" and "g abs_summable_on A"
shows "(\<lambda>x. f x + g x) abs_summable_on A"
using assms unfolding abs_summable_on_def by (rule Bochner_Integration.integrable_add)
lemma abs_summable_on_diff [intro]:
assumes "f abs_summable_on A" and "g abs_summable_on A"
shows "(\<lambda>x. f x - g x) abs_summable_on A"
using assms unfolding abs_summable_on_def by (rule Bochner_Integration.integrable_diff)
lemma abs_summable_on_scaleR_left [intro]:
assumes "c \<noteq> 0 \<Longrightarrow> f abs_summable_on A"
shows "(\<lambda>x. f x *\<^sub>R c) abs_summable_on A"
using assms unfolding abs_summable_on_def by (intro Bochner_Integration.integrable_scaleR_left)
lemma abs_summable_on_scaleR_right [intro]:
assumes "c \<noteq> 0 \<Longrightarrow> f abs_summable_on A"
shows "(\<lambda>x. c *\<^sub>R f x) abs_summable_on A"
using assms unfolding abs_summable_on_def by (intro Bochner_Integration.integrable_scaleR_right)
lemma abs_summable_on_cmult_right [intro]:
fixes f :: "'a \<Rightarrow> 'b :: {banach, real_normed_algebra, second_countable_topology}"
assumes "c \<noteq> 0 \<Longrightarrow> f abs_summable_on A"
shows "(\<lambda>x. c * f x) abs_summable_on A"
using assms unfolding abs_summable_on_def by (intro Bochner_Integration.integrable_mult_right)
lemma abs_summable_on_cmult_left [intro]:
fixes f :: "'a \<Rightarrow> 'b :: {banach, real_normed_algebra, second_countable_topology}"
assumes "c \<noteq> 0 \<Longrightarrow> f abs_summable_on A"
shows "(\<lambda>x. f x * c) abs_summable_on A"
using assms unfolding abs_summable_on_def by (intro Bochner_Integration.integrable_mult_left)
lemma abs_summable_on_prod_PiE:
fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c :: {real_normed_field,banach,second_countable_topology}"
assumes finite: "finite A" and countable: "\<And>x. x \<in> A \<Longrightarrow> countable (B x)"
assumes summable: "\<And>x. x \<in> A \<Longrightarrow> f x abs_summable_on B x"
shows "(\<lambda>g. \<Prod>x\<in>A. f x (g x)) abs_summable_on PiE A B"
proof -
define B' where "B' = (\<lambda>x. if x \<in> A then B x else {})"
from assms have [simp]: "countable (B' x)" for x
by (auto simp: B'_def)
then interpret product_sigma_finite "count_space \<circ> B'"
unfolding o_def by (intro product_sigma_finite.intro sigma_finite_measure_count_space_countable)
from assms have "integrable (PiM A (count_space \<circ> B')) (\<lambda>g. \<Prod>x\<in>A. f x (g x))"
by (intro product_integrable_prod) (auto simp: abs_summable_on_def B'_def)
also have "PiM A (count_space \<circ> B') = count_space (PiE A B')"
unfolding o_def using finite by (intro count_space_PiM_finite) simp_all
also have "PiE A B' = PiE A B" by (intro PiE_cong) (simp_all add: B'_def)
finally show ?thesis by (simp add: abs_summable_on_def)
qed
lemma not_summable_infsetsum_eq:
"\<not>f abs_summable_on A \<Longrightarrow> infsetsum f A = 0"
by (simp add: abs_summable_on_def infsetsum_def not_integrable_integral_eq)
lemma infsetsum_altdef:
"infsetsum f A = set_lebesgue_integral (count_space UNIV) A f"
unfolding set_lebesgue_integral_def
by (subst integral_restrict_space [symmetric])
(auto simp: restrict_count_space_subset infsetsum_def)
lemma infsetsum_altdef':
"A \<subseteq> B \<Longrightarrow> infsetsum f A = set_lebesgue_integral (count_space B) A f"
unfolding set_lebesgue_integral_def
by (subst integral_restrict_space [symmetric])
(auto simp: restrict_count_space_subset infsetsum_def)
lemma nn_integral_conv_infsetsum:
assumes "f abs_summable_on A" "\<And>x. x \<in> A \<Longrightarrow> f x \<ge> 0"
shows "nn_integral (count_space A) f = ennreal (infsetsum f A)"
using assms unfolding infsetsum_def abs_summable_on_def
by (subst nn_integral_eq_integral) auto
lemma infsetsum_conv_nn_integral:
assumes "nn_integral (count_space A) f \<noteq> \<infinity>" "\<And>x. x \<in> A \<Longrightarrow> f x \<ge> 0"
shows "infsetsum f A = enn2real (nn_integral (count_space A) f)"
unfolding infsetsum_def using assms
by (subst integral_eq_nn_integral) auto
lemma infsetsum_cong [cong]:
"(\<And>x. x \<in> A \<Longrightarrow> f x = g x) \<Longrightarrow> A = B \<Longrightarrow> infsetsum f A = infsetsum g B"
unfolding infsetsum_def by (intro Bochner_Integration.integral_cong) auto
lemma infsetsum_0 [simp]: "infsetsum (\<lambda>_. 0) A = 0"
by (simp add: infsetsum_def)
lemma infsetsum_all_0: "(\<And>x. x \<in> A \<Longrightarrow> f x = 0) \<Longrightarrow> infsetsum f A = 0"
by simp
lemma infsetsum_nonneg: "(\<And>x. x \<in> A \<Longrightarrow> f x \<ge> (0::real)) \<Longrightarrow> infsetsum f A \<ge> 0"
unfolding infsetsum_def by (rule Bochner_Integration.integral_nonneg) auto
lemma sum_infsetsum:
assumes "\<And>x. x \<in> A \<Longrightarrow> f x abs_summable_on B"
shows "(\<Sum>x\<in>A. \<Sum>\<^sub>ay\<in>B. f x y) = (\<Sum>\<^sub>ay\<in>B. \<Sum>x\<in>A. f x y)"
using assms by (simp add: infsetsum_def abs_summable_on_def Bochner_Integration.integral_sum)
lemma Re_infsetsum: "f abs_summable_on A \<Longrightarrow> Re (infsetsum f A) = (\<Sum>\<^sub>ax\<in>A. Re (f x))"
by (simp add: infsetsum_def abs_summable_on_def)
lemma Im_infsetsum: "f abs_summable_on A \<Longrightarrow> Im (infsetsum f A) = (\<Sum>\<^sub>ax\<in>A. Im (f x))"
by (simp add: infsetsum_def abs_summable_on_def)
lemma infsetsum_of_real:
shows "infsetsum (\<lambda>x. of_real (f x)
:: 'a :: {real_normed_algebra_1,banach,second_countable_topology,real_inner}) A =
of_real (infsetsum f A)"
unfolding infsetsum_def
by (rule integral_bounded_linear'[OF bounded_linear_of_real bounded_linear_inner_left[of 1]]) auto
lemma infsetsum_finite [simp]: "finite A \<Longrightarrow> infsetsum f A = (\<Sum>x\<in>A. f x)"
by (simp add: infsetsum_def lebesgue_integral_count_space_finite)
lemma infsetsum_nat:
assumes "f abs_summable_on A"
shows "infsetsum f A = (\<Sum>n. if n \<in> A then f n else 0)"
proof -
from assms have "infsetsum f A = (\<Sum>n. indicator A n *\<^sub>R f n)"
unfolding infsetsum_altdef abs_summable_on_altdef set_lebesgue_integral_def set_integrable_def
by (subst integral_count_space_nat) auto
also have "(\<lambda>n. indicator A n *\<^sub>R f n) = (\<lambda>n. if n \<in> A then f n else 0)"
by auto
finally show ?thesis .
qed
lemma infsetsum_nat':
assumes "f abs_summable_on UNIV"
shows "infsetsum f UNIV = (\<Sum>n. f n)"
using assms by (subst infsetsum_nat) auto
lemma sums_infsetsum_nat:
assumes "f abs_summable_on A"
shows "(\<lambda>n. if n \<in> A then f n else 0) sums infsetsum f A"
proof -
from assms have "summable (\<lambda>n. if n \<in> A then norm (f n) else 0)"
by (simp add: abs_summable_on_nat_iff)
also have "(\<lambda>n. if n \<in> A then norm (f n) else 0) = (\<lambda>n. norm (if n \<in> A then f n else 0))"
by auto
finally have "summable (\<lambda>n. if n \<in> A then f n else 0)"
by (rule summable_norm_cancel)
with assms show ?thesis
by (auto simp: sums_iff infsetsum_nat)
qed
lemma sums_infsetsum_nat':
assumes "f abs_summable_on UNIV"
shows "f sums infsetsum f UNIV"
using sums_infsetsum_nat [OF assms] by simp
lemma infsetsum_Un_disjoint:
assumes "f abs_summable_on A" "f abs_summable_on B" "A \<inter> B = {}"
shows "infsetsum f (A \<union> B) = infsetsum f A + infsetsum f B"
using assms unfolding infsetsum_altdef abs_summable_on_altdef
by (subst set_integral_Un) auto
lemma infsetsum_Diff:
assumes "f abs_summable_on B" "A \<subseteq> B"
shows "infsetsum f (B - A) = infsetsum f B - infsetsum f A"
proof -
have "infsetsum f ((B - A) \<union> A) = infsetsum f (B - A) + infsetsum f A"
using assms(2) by (intro infsetsum_Un_disjoint abs_summable_on_subset[OF assms(1)]) auto
also from assms(2) have "(B - A) \<union> A = B"
by auto
ultimately show ?thesis
by (simp add: algebra_simps)
qed
lemma infsetsum_Un_Int:
assumes "f abs_summable_on (A \<union> B)"
shows "infsetsum f (A \<union> B) = infsetsum f A + infsetsum f B - infsetsum f (A \<inter> B)"
proof -
have "A \<union> B = A \<union> (B - A \<inter> B)"
by auto
also have "infsetsum f \<dots> = infsetsum f A + infsetsum f (B - A \<inter> B)"
by (intro infsetsum_Un_disjoint abs_summable_on_subset[OF assms]) auto
also have "infsetsum f (B - A \<inter> B) = infsetsum f B - infsetsum f (A \<inter> B)"
by (intro infsetsum_Diff abs_summable_on_subset[OF assms]) auto
finally show ?thesis
by (simp add: algebra_simps)
qed
lemma infsetsum_reindex_bij_betw:
assumes "bij_betw g A B"
shows "infsetsum (\<lambda>x. f (g x)) A = infsetsum f B"
proof -
have *: "count_space B = distr (count_space A) (count_space B) g"
by (rule distr_bij_count_space [symmetric]) fact
show ?thesis unfolding infsetsum_def
by (subst *, subst integral_distr[of _ _ "count_space B"])
(insert assms, auto simp: bij_betw_def)
qed
theorem infsetsum_reindex:
assumes "inj_on g A"
shows "infsetsum f (g ` A) = infsetsum (\<lambda>x. f (g x)) A"
by (intro infsetsum_reindex_bij_betw [symmetric] inj_on_imp_bij_betw assms)
lemma infsetsum_cong_neutral:
assumes "\<And>x. x \<in> A - B \<Longrightarrow> f x = 0"
assumes "\<And>x. x \<in> B - A \<Longrightarrow> g x = 0"
assumes "\<And>x. x \<in> A \<inter> B \<Longrightarrow> f x = g x"
shows "infsetsum f A = infsetsum g B"
unfolding infsetsum_altdef set_lebesgue_integral_def using assms
by (intro Bochner_Integration.integral_cong refl)
(auto simp: indicator_def split: if_splits)
lemma infsetsum_mono_neutral:
fixes f g :: "'a \<Rightarrow> real"
assumes "f abs_summable_on A" and "g abs_summable_on B"
assumes "\<And>x. x \<in> A \<Longrightarrow> f x \<le> g x"
assumes "\<And>x. x \<in> A - B \<Longrightarrow> f x \<le> 0"
assumes "\<And>x. x \<in> B - A \<Longrightarrow> g x \<ge> 0"
shows "infsetsum f A \<le> infsetsum g B"
using assms unfolding infsetsum_altdef set_lebesgue_integral_def abs_summable_on_altdef set_integrable_def
by (intro Bochner_Integration.integral_mono) (auto simp: indicator_def)
lemma infsetsum_mono_neutral_left:
fixes f g :: "'a \<Rightarrow> real"
assumes "f abs_summable_on A" and "g abs_summable_on B"
assumes "\<And>x. x \<in> A \<Longrightarrow> f x \<le> g x"
assumes "A \<subseteq> B"
assumes "\<And>x. x \<in> B - A \<Longrightarrow> g x \<ge> 0"
shows "infsetsum f A \<le> infsetsum g B"
using \<open>A \<subseteq> B\<close> by (intro infsetsum_mono_neutral assms) auto
lemma infsetsum_mono_neutral_right:
fixes f g :: "'a \<Rightarrow> real"
assumes "f abs_summable_on A" and "g abs_summable_on B"
assumes "\<And>x. x \<in> A \<Longrightarrow> f x \<le> g x"
assumes "B \<subseteq> A"
assumes "\<And>x. x \<in> A - B \<Longrightarrow> f x \<le> 0"
shows "infsetsum f A \<le> infsetsum g B"
using \<open>B \<subseteq> A\<close> by (intro infsetsum_mono_neutral assms) auto
lemma infsetsum_mono:
fixes f g :: "'a \<Rightarrow> real"
assumes "f abs_summable_on A" and "g abs_summable_on A"
assumes "\<And>x. x \<in> A \<Longrightarrow> f x \<le> g x"
shows "infsetsum f A \<le> infsetsum g A"
by (intro infsetsum_mono_neutral assms) auto
lemma norm_infsetsum_bound:
"norm (infsetsum f A) \<le> infsetsum (\<lambda>x. norm (f x)) A"
unfolding abs_summable_on_def infsetsum_def
by (rule Bochner_Integration.integral_norm_bound)
theorem infsetsum_Sigma:
fixes A :: "'a set" and B :: "'a \<Rightarrow> 'b set"
assumes [simp]: "countable A" and "\<And>i. countable (B i)"
assumes summable: "f abs_summable_on (Sigma A B)"
shows "infsetsum f (Sigma A B) = infsetsum (\<lambda>x. infsetsum (\<lambda>y. f (x, y)) (B x)) A"
proof -
define B' where "B' = (\<Union>i\<in>A. B i)"
have [simp]: "countable B'"
unfolding B'_def by (intro countable_UN assms)
interpret pair_sigma_finite "count_space A" "count_space B'"
by (intro pair_sigma_finite.intro sigma_finite_measure_count_space_countable) fact+
have "integrable (count_space (A \<times> B')) (\<lambda>z. indicator (Sigma A B) z *\<^sub>R f z)"
using summable
by (metis (mono_tags, lifting) abs_summable_on_altdef abs_summable_on_def integrable_cong integrable_mult_indicator set_integrable_def sets_UNIV)
also have "?this \<longleftrightarrow> integrable (count_space A \<Otimes>\<^sub>M count_space B') (\<lambda>(x, y). indicator (B x) y *\<^sub>R f (x, y))"
by (intro Bochner_Integration.integrable_cong)
(auto simp: pair_measure_countable indicator_def split: if_splits)
finally have integrable: \<dots> .
have "infsetsum (\<lambda>x. infsetsum (\<lambda>y. f (x, y)) (B x)) A =
(\<integral>x. infsetsum (\<lambda>y. f (x, y)) (B x) \<partial>count_space A)"
unfolding infsetsum_def by simp
also have "\<dots> = (\<integral>x. \<integral>y. indicator (B x) y *\<^sub>R f (x, y) \<partial>count_space B' \<partial>count_space A)"
proof (rule Bochner_Integration.integral_cong [OF refl])
show "\<And>x. x \<in> space (count_space A) \<Longrightarrow>
(\<Sum>\<^sub>ay\<in>B x. f (x, y)) = LINT y|count_space B'. indicat_real (B x) y *\<^sub>R f (x, y)"
using infsetsum_altdef'[of _ B']
unfolding set_lebesgue_integral_def B'_def
by auto
qed
also have "\<dots> = (\<integral>(x,y). indicator (B x) y *\<^sub>R f (x, y) \<partial>(count_space A \<Otimes>\<^sub>M count_space B'))"
by (subst integral_fst [OF integrable]) auto
also have "\<dots> = (\<integral>z. indicator (Sigma A B) z *\<^sub>R f z \<partial>count_space (A \<times> B'))"
by (intro Bochner_Integration.integral_cong)
(auto simp: pair_measure_countable indicator_def split: if_splits)
also have "\<dots> = infsetsum f (Sigma A B)"
unfolding set_lebesgue_integral_def [symmetric]
by (rule infsetsum_altdef' [symmetric]) (auto simp: B'_def)
finally show ?thesis ..
qed
lemma infsetsum_Sigma':
fixes A :: "'a set" and B :: "'a \<Rightarrow> 'b set"
assumes [simp]: "countable A" and "\<And>i. countable (B i)"
assumes summable: "(\<lambda>(x,y). f x y) abs_summable_on (Sigma A B)"
shows "infsetsum (\<lambda>x. infsetsum (\<lambda>y. f x y) (B x)) A = infsetsum (\<lambda>(x,y). f x y) (Sigma A B)"
using assms by (subst infsetsum_Sigma) auto
lemma infsetsum_Times:
fixes A :: "'a set" and B :: "'b set"
assumes [simp]: "countable A" and "countable B"
assumes summable: "f abs_summable_on (A \<times> B)"
shows "infsetsum f (A \<times> B) = infsetsum (\<lambda>x. infsetsum (\<lambda>y. f (x, y)) B) A"
using assms by (subst infsetsum_Sigma) auto
lemma infsetsum_Times':
fixes A :: "'a set" and B :: "'b set"
fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c :: {banach, second_countable_topology}"
assumes [simp]: "countable A" and [simp]: "countable B"
assumes summable: "(\<lambda>(x,y). f x y) abs_summable_on (A \<times> B)"
shows "infsetsum (\<lambda>x. infsetsum (\<lambda>y. f x y) B) A = infsetsum (\<lambda>(x,y). f x y) (A \<times> B)"
using assms by (subst infsetsum_Times) auto
lemma infsetsum_swap:
fixes A :: "'a set" and B :: "'b set"
fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c :: {banach, second_countable_topology}"
assumes [simp]: "countable A" and [simp]: "countable B"
assumes summable: "(\<lambda>(x,y). f x y) abs_summable_on A \<times> B"
shows "infsetsum (\<lambda>x. infsetsum (\<lambda>y. f x y) B) A = infsetsum (\<lambda>y. infsetsum (\<lambda>x. f x y) A) B"
proof -
from summable have summable': "(\<lambda>(x,y). f y x) abs_summable_on B \<times> A"
by (subst abs_summable_on_Times_swap) auto
have bij: "bij_betw (\<lambda>(x, y). (y, x)) (B \<times> A) (A \<times> B)"
by (auto simp: bij_betw_def inj_on_def)
have "infsetsum (\<lambda>x. infsetsum (\<lambda>y. f x y) B) A = infsetsum (\<lambda>(x,y). f x y) (A \<times> B)"
using summable by (subst infsetsum_Times) auto
also have "\<dots> = infsetsum (\<lambda>(x,y). f y x) (B \<times> A)"
by (subst infsetsum_reindex_bij_betw[OF bij, of "\<lambda>(x,y). f x y", symmetric])
(simp_all add: case_prod_unfold)
also have "\<dots> = infsetsum (\<lambda>y. infsetsum (\<lambda>x. f x y) A) B"
using summable' by (subst infsetsum_Times) auto
finally show ?thesis .
qed
theorem abs_summable_on_Sigma_iff:
assumes [simp]: "countable A" and "\<And>x. x \<in> A \<Longrightarrow> countable (B x)"
shows "f abs_summable_on Sigma A B \<longleftrightarrow>
(\<forall>x\<in>A. (\<lambda>y. f (x, y)) abs_summable_on B x) \<and>
((\<lambda>x. infsetsum (\<lambda>y. norm (f (x, y))) (B x)) abs_summable_on A)"
proof safe
define B' where "B' = (\<Union>x\<in>A. B x)"
have [simp]: "countable B'"
unfolding B'_def using assms by auto
interpret pair_sigma_finite "count_space A" "count_space B'"
by (intro pair_sigma_finite.intro sigma_finite_measure_count_space_countable) fact+
{
assume *: "f abs_summable_on Sigma A B"
thus "(\<lambda>y. f (x, y)) abs_summable_on B x" if "x \<in> A" for x
using that by (rule abs_summable_on_Sigma_project2)
have "set_integrable (count_space (A \<times> B')) (Sigma A B) (\<lambda>z. norm (f z))"
using abs_summable_on_normI[OF *]
by (subst abs_summable_on_altdef' [symmetric]) (auto simp: B'_def)
also have "count_space (A \<times> B') = count_space A \<Otimes>\<^sub>M count_space B'"
by (simp add: pair_measure_countable)
finally have "integrable (count_space A)
(\<lambda>x. lebesgue_integral (count_space B')
(\<lambda>y. indicator (Sigma A B) (x, y) *\<^sub>R norm (f (x, y))))"
unfolding set_integrable_def by (rule integrable_fst')
also have "?this \<longleftrightarrow> integrable (count_space A)
(\<lambda>x. lebesgue_integral (count_space B')
(\<lambda>y. indicator (B x) y *\<^sub>R norm (f (x, y))))"
by (intro integrable_cong refl) (simp_all add: indicator_def)
also have "\<dots> \<longleftrightarrow> integrable (count_space A) (\<lambda>x. infsetsum (\<lambda>y. norm (f (x, y))) (B x))"
unfolding set_lebesgue_integral_def [symmetric]
by (intro integrable_cong refl infsetsum_altdef' [symmetric]) (auto simp: B'_def)
also have "\<dots> \<longleftrightarrow> (\<lambda>x. infsetsum (\<lambda>y. norm (f (x, y))) (B x)) abs_summable_on A"
by (simp add: abs_summable_on_def)
finally show \<dots> .
}
{
assume *: "\<forall>x\<in>A. (\<lambda>y. f (x, y)) abs_summable_on B x"
assume "(\<lambda>x. \<Sum>\<^sub>ay\<in>B x. norm (f (x, y))) abs_summable_on A"
also have "?this \<longleftrightarrow> (\<lambda>x. \<integral>y\<in>B x. norm (f (x, y)) \<partial>count_space B') abs_summable_on A"
by (intro abs_summable_on_cong refl infsetsum_altdef') (auto simp: B'_def)
also have "\<dots> \<longleftrightarrow> (\<lambda>x. \<integral>y. indicator (Sigma A B) (x, y) *\<^sub>R norm (f (x, y)) \<partial>count_space B')
abs_summable_on A" (is "_ \<longleftrightarrow> ?h abs_summable_on _")
unfolding set_lebesgue_integral_def
by (intro abs_summable_on_cong) (auto simp: indicator_def)
also have "\<dots> \<longleftrightarrow> integrable (count_space A) ?h"
by (simp add: abs_summable_on_def)
finally have **: \<dots> .
have "integrable (count_space A \<Otimes>\<^sub>M count_space B') (\<lambda>z. indicator (Sigma A B) z *\<^sub>R f z)"
proof (rule Fubini_integrable, goal_cases)
case 3
{
fix x assume x: "x \<in> A"
with * have "(\<lambda>y. f (x, y)) abs_summable_on B x"
by blast
also have "?this \<longleftrightarrow> integrable (count_space B')
(\<lambda>y. indicator (B x) y *\<^sub>R f (x, y))"
unfolding set_integrable_def [symmetric]
using x by (intro abs_summable_on_altdef') (auto simp: B'_def)
also have "(\<lambda>y. indicator (B x) y *\<^sub>R f (x, y)) =
(\<lambda>y. indicator (Sigma A B) (x, y) *\<^sub>R f (x, y))"
using x by (auto simp: indicator_def)
finally have "integrable (count_space B')
(\<lambda>y. indicator (Sigma A B) (x, y) *\<^sub>R f (x, y))" .
}
thus ?case by (auto simp: AE_count_space)
qed (insert **, auto simp: pair_measure_countable)
moreover have "count_space A \<Otimes>\<^sub>M count_space B' = count_space (A \<times> B')"
by (simp add: pair_measure_countable)
moreover have "set_integrable (count_space (A \<times> B')) (Sigma A B) f \<longleftrightarrow>
f abs_summable_on Sigma A B"
by (rule abs_summable_on_altdef' [symmetric]) (auto simp: B'_def)
ultimately show "f abs_summable_on Sigma A B"
by (simp add: set_integrable_def)
}
qed
lemma abs_summable_on_Sigma_project1:
assumes "(\<lambda>(x,y). f x y) abs_summable_on Sigma A B"
assumes [simp]: "countable A" and "\<And>x. x \<in> A \<Longrightarrow> countable (B x)"
shows "(\<lambda>x. infsetsum (\<lambda>y. norm (f x y)) (B x)) abs_summable_on A"
using assms by (subst (asm) abs_summable_on_Sigma_iff) auto
lemma abs_summable_on_Sigma_project1':
assumes "(\<lambda>(x,y). f x y) abs_summable_on Sigma A B"
assumes [simp]: "countable A" and "\<And>x. x \<in> A \<Longrightarrow> countable (B x)"
shows "(\<lambda>x. infsetsum (\<lambda>y. f x y) (B x)) abs_summable_on A"
by (intro abs_summable_on_comparison_test' [OF abs_summable_on_Sigma_project1[OF assms]]
norm_infsetsum_bound)
theorem infsetsum_prod_PiE:
fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c :: {real_normed_field,banach,second_countable_topology}"
assumes finite: "finite A" and countable: "\<And>x. x \<in> A \<Longrightarrow> countable (B x)"
assumes summable: "\<And>x. x \<in> A \<Longrightarrow> f x abs_summable_on B x"
shows "infsetsum (\<lambda>g. \<Prod>x\<in>A. f x (g x)) (PiE A B) = (\<Prod>x\<in>A. infsetsum (f x) (B x))"
proof -
define B' where "B' = (\<lambda>x. if x \<in> A then B x else {})"
from assms have [simp]: "countable (B' x)" for x
by (auto simp: B'_def)
then interpret product_sigma_finite "count_space \<circ> B'"
unfolding o_def by (intro product_sigma_finite.intro sigma_finite_measure_count_space_countable)
have "infsetsum (\<lambda>g. \<Prod>x\<in>A. f x (g x)) (PiE A B) =
(\<integral>g. (\<Prod>x\<in>A. f x (g x)) \<partial>count_space (PiE A B))"
by (simp add: infsetsum_def)
also have "PiE A B = PiE A B'"
by (intro PiE_cong) (simp_all add: B'_def)
hence "count_space (PiE A B) = count_space (PiE A B')"
by simp
also have "\<dots> = PiM A (count_space \<circ> B')"
unfolding o_def using finite by (intro count_space_PiM_finite [symmetric]) simp_all
also have "(\<integral>g. (\<Prod>x\<in>A. f x (g x)) \<partial>\<dots>) = (\<Prod>x\<in>A. infsetsum (f x) (B' x))"
by (subst product_integral_prod)
(insert summable finite, simp_all add: infsetsum_def B'_def abs_summable_on_def)
also have "\<dots> = (\<Prod>x\<in>A. infsetsum (f x) (B x))"
by (intro prod.cong refl) (simp_all add: B'_def)
finally show ?thesis .
qed
lemma infsetsum_uminus: "infsetsum (\<lambda>x. -f x) A = -infsetsum f A"
unfolding infsetsum_def abs_summable_on_def
by (rule Bochner_Integration.integral_minus)
lemma infsetsum_add:
assumes "f abs_summable_on A" and "g abs_summable_on A"
shows "infsetsum (\<lambda>x. f x + g x) A = infsetsum f A + infsetsum g A"
using assms unfolding infsetsum_def abs_summable_on_def
by (rule Bochner_Integration.integral_add)
lemma infsetsum_diff:
assumes "f abs_summable_on A" and "g abs_summable_on A"
shows "infsetsum (\<lambda>x. f x - g x) A = infsetsum f A - infsetsum g A"
using assms unfolding infsetsum_def abs_summable_on_def
by (rule Bochner_Integration.integral_diff)
lemma infsetsum_scaleR_left:
assumes "c \<noteq> 0 \<Longrightarrow> f abs_summable_on A"
shows "infsetsum (\<lambda>x. f x *\<^sub>R c) A = infsetsum f A *\<^sub>R c"
using assms unfolding infsetsum_def abs_summable_on_def
by (rule Bochner_Integration.integral_scaleR_left)
lemma infsetsum_scaleR_right:
"infsetsum (\<lambda>x. c *\<^sub>R f x) A = c *\<^sub>R infsetsum f A"
unfolding infsetsum_def abs_summable_on_def
by (subst Bochner_Integration.integral_scaleR_right) auto
lemma infsetsum_cmult_left:
fixes f :: "'a \<Rightarrow> 'b :: {banach, real_normed_algebra, second_countable_topology}"
assumes "c \<noteq> 0 \<Longrightarrow> f abs_summable_on A"
shows "infsetsum (\<lambda>x. f x * c) A = infsetsum f A * c"
using assms unfolding infsetsum_def abs_summable_on_def
by (rule Bochner_Integration.integral_mult_left)
lemma infsetsum_cmult_right:
fixes f :: "'a \<Rightarrow> 'b :: {banach, real_normed_algebra, second_countable_topology}"
assumes "c \<noteq> 0 \<Longrightarrow> f abs_summable_on A"
shows "infsetsum (\<lambda>x. c * f x) A = c * infsetsum f A"
using assms unfolding infsetsum_def abs_summable_on_def
by (rule Bochner_Integration.integral_mult_right)
lemma infsetsum_cdiv:
fixes f :: "'a \<Rightarrow> 'b :: {banach, real_normed_field, second_countable_topology}"
assumes "c \<noteq> 0 \<Longrightarrow> f abs_summable_on A"
shows "infsetsum (\<lambda>x. f x / c) A = infsetsum f A / c"
using assms unfolding infsetsum_def abs_summable_on_def by auto
(* TODO Generalise with bounded_linear *)
lemma
fixes f :: "'a \<Rightarrow> 'c :: {banach, real_normed_field, second_countable_topology}"
assumes [simp]: "countable A" and [simp]: "countable B"
assumes "f abs_summable_on A" and "g abs_summable_on B"
shows abs_summable_on_product: "(\<lambda>(x,y). f x * g y) abs_summable_on A \<times> B"
and infsetsum_product: "infsetsum (\<lambda>(x,y). f x * g y) (A \<times> B) =
infsetsum f A * infsetsum g B"
proof -
from assms show "(\<lambda>(x,y). f x * g y) abs_summable_on A \<times> B"
by (subst abs_summable_on_Sigma_iff)
(auto intro!: abs_summable_on_cmult_right simp: norm_mult infsetsum_cmult_right)
with assms show "infsetsum (\<lambda>(x,y). f x * g y) (A \<times> B) = infsetsum f A * infsetsum g B"
by (subst infsetsum_Sigma)
(auto simp: infsetsum_cmult_left infsetsum_cmult_right)
qed
end
diff --git a/src/HOL/Analysis/Jordan_Curve.thy b/src/HOL/Analysis/Jordan_Curve.thy
--- a/src/HOL/Analysis/Jordan_Curve.thy
+++ b/src/HOL/Analysis/Jordan_Curve.thy
@@ -1,675 +1,675 @@
(* Title: HOL/Analysis/Jordan_Curve.thy
Authors: LC Paulson, based on material from HOL Light
*)
section \<open>The Jordan Curve Theorem and Applications\<close>
theory Jordan_Curve
imports Arcwise_Connected Further_Topology
begin
subsection\<open>Janiszewski's theorem\<close>
lemma Janiszewski_weak:
fixes a b::complex
assumes "compact S" "compact T" and conST: "connected(S \<inter> T)"
and ccS: "connected_component (- S) a b" and ccT: "connected_component (- T) a b"
shows "connected_component (- (S \<union> T)) a b"
proof -
have [simp]: "a \<notin> S" "a \<notin> T" "b \<notin> S" "b \<notin> T"
by (meson ComplD ccS ccT connected_component_in)+
have clo: "closedin (top_of_set (S \<union> T)) S" "closedin (top_of_set (S \<union> T)) T"
by (simp_all add: assms closed_subset compact_imp_closed)
obtain g where contg: "continuous_on S g"
and g: "\<And>x. x \<in> S \<Longrightarrow> exp (\<i>* of_real (g x)) = (x - a) /\<^sub>R cmod (x - a) / ((x - b) /\<^sub>R cmod (x - b))"
using ccS \<open>compact S\<close>
apply (simp add: Borsuk_maps_homotopic_in_connected_component_eq [symmetric])
apply (subst (asm) homotopic_circlemaps_divide)
apply (auto simp: inessential_eq_continuous_logarithm_circle)
done
obtain h where conth: "continuous_on T h"
and h: "\<And>x. x \<in> T \<Longrightarrow> exp (\<i>* of_real (h x)) = (x - a) /\<^sub>R cmod (x - a) / ((x - b) /\<^sub>R cmod (x - b))"
using ccT \<open>compact T\<close>
apply (simp add: Borsuk_maps_homotopic_in_connected_component_eq [symmetric])
apply (subst (asm) homotopic_circlemaps_divide)
apply (auto simp: inessential_eq_continuous_logarithm_circle)
done
have "continuous_on (S \<union> T) (\<lambda>x. (x - a) /\<^sub>R cmod (x - a))" "continuous_on (S \<union> T) (\<lambda>x. (x - b) /\<^sub>R cmod (x - b))"
by (intro continuous_intros; force)+
moreover have "(\<lambda>x. (x - a) /\<^sub>R cmod (x - a)) ` (S \<union> T) \<subseteq> sphere 0 1" "(\<lambda>x. (x - b) /\<^sub>R cmod (x - b)) ` (S \<union> T) \<subseteq> sphere 0 1"
by (auto simp: divide_simps)
moreover have "\<exists>g. continuous_on (S \<union> T) g \<and>
(\<forall>x\<in>S \<union> T. (x - a) /\<^sub>R cmod (x - a) / ((x - b) /\<^sub>R cmod (x - b)) = exp (\<i>*complex_of_real (g x)))"
proof (cases "S \<inter> T = {}")
case True
have "continuous_on (S \<union> T) (\<lambda>x. if x \<in> S then g x else h x)"
apply (rule continuous_on_cases_local [OF clo contg conth])
using True by auto
then show ?thesis
by (rule_tac x="(\<lambda>x. if x \<in> S then g x else h x)" in exI) (auto simp: g h)
next
case False
have diffpi: "\<exists>n. g x = h x + 2* of_int n*pi" if "x \<in> S \<inter> T" for x
proof -
have "exp (\<i>* of_real (g x)) = exp (\<i>* of_real (h x))"
using that by (simp add: g h)
then obtain n where "complex_of_real (g x) = complex_of_real (h x) + 2* of_int n*complex_of_real pi"
apply (auto simp: exp_eq)
by (metis complex_i_not_zero distrib_left mult.commute mult_cancel_left)
then show ?thesis
apply (rule_tac x=n in exI)
using of_real_eq_iff by fastforce
qed
have contgh: "continuous_on (S \<inter> T) (\<lambda>x. g x - h x)"
by (intro continuous_intros continuous_on_subset [OF contg] continuous_on_subset [OF conth]) auto
moreover have disc:
"\<exists>e>0. \<forall>y. y \<in> S \<inter> T \<and> g y - h y \<noteq> g x - h x \<longrightarrow> e \<le> norm ((g y - h y) - (g x - h x))"
if "x \<in> S \<inter> T" for x
proof -
obtain nx where nx: "g x = h x + 2* of_int nx*pi"
using \<open>x \<in> S \<inter> T\<close> diffpi by blast
have "2*pi \<le> norm (g y - h y - (g x - h x))" if y: "y \<in> S \<inter> T" and neq: "g y - h y \<noteq> g x - h x" for y
proof -
obtain ny where ny: "g y = h y + 2* of_int ny*pi"
using \<open>y \<in> S \<inter> T\<close> diffpi by blast
{ assume "nx \<noteq> ny"
then have "1 \<le> \<bar>real_of_int ny - real_of_int nx\<bar>"
by linarith
then have "(2*pi)*1 \<le> (2*pi)*\<bar>real_of_int ny - real_of_int nx\<bar>"
by simp
also have "... = \<bar>2*real_of_int ny*pi - 2*real_of_int nx*pi\<bar>"
by (simp add: algebra_simps abs_if)
finally have "2*pi \<le> \<bar>2*real_of_int ny*pi - 2*real_of_int nx*pi\<bar>" by simp
}
with neq show ?thesis
by (simp add: nx ny)
qed
then show ?thesis
by (rule_tac x="2*pi" in exI) auto
qed
ultimately have "(\<lambda>x. g x - h x) constant_on S \<inter> T"
using continuous_discrete_range_constant [OF conST contgh] by blast
then obtain z where z: "\<And>x. x \<in> S \<inter> T \<Longrightarrow> g x - h x = z"
by (auto simp: constant_on_def)
obtain w where "exp(\<i> * of_real(h w)) = exp (\<i> * of_real(z + h w))"
using disc z False
by auto (metis diff_add_cancel g h of_real_add)
then have [simp]: "exp (\<i>* of_real z) = 1"
by (metis cis_conv_exp cis_mult exp_not_eq_zero mult_cancel_right1)
show ?thesis
proof (intro exI conjI)
show "continuous_on (S \<union> T) (\<lambda>x. if x \<in> S then g x else z + h x)"
apply (intro continuous_intros continuous_on_cases_local [OF clo contg] conth)
using z by fastforce
qed (auto simp: g h algebra_simps exp_add)
qed
ultimately have *: "homotopic_with_canon (\<lambda>x. True) (S \<union> T) (sphere 0 1)
(\<lambda>x. (x - a) /\<^sub>R cmod (x - a)) (\<lambda>x. (x - b) /\<^sub>R cmod (x - b))"
by (subst homotopic_circlemaps_divide) (auto simp: inessential_eq_continuous_logarithm_circle)
show ?thesis
apply (rule Borsuk_maps_homotopic_in_connected_component_eq [THEN iffD1])
using assms by (auto simp: *)
qed
theorem Janiszewski:
fixes a b :: complex
assumes "compact S" "closed T" and conST: "connected (S \<inter> T)"
and ccS: "connected_component (- S) a b" and ccT: "connected_component (- T) a b"
shows "connected_component (- (S \<union> T)) a b"
proof -
have "path_component(- T) a b"
by (simp add: \<open>closed T\<close> ccT open_Compl open_path_connected_component)
then obtain g where g: "path g" "path_image g \<subseteq> - T" "pathstart g = a" "pathfinish g = b"
by (auto simp: path_component_def)
obtain C where C: "compact C" "connected C" "a \<in> C" "b \<in> C" "C \<inter> T = {}"
proof
show "compact (path_image g)"
by (simp add: \<open>path g\<close> compact_path_image)
show "connected (path_image g)"
by (simp add: \<open>path g\<close> connected_path_image)
qed (use g in auto)
obtain r where "0 < r" and r: "C \<union> S \<subseteq> ball 0 r"
by (metis \<open>compact C\<close> \<open>compact S\<close> bounded_Un compact_imp_bounded bounded_subset_ballD)
have "connected_component (- (S \<union> (T \<inter> cball 0 r \<union> sphere 0 r))) a b"
proof (rule Janiszewski_weak [OF \<open>compact S\<close>])
show comT': "compact ((T \<inter> cball 0 r) \<union> sphere 0 r)"
by (simp add: \<open>closed T\<close> closed_Int_compact compact_Un)
have "S \<inter> (T \<inter> cball 0 r \<union> sphere 0 r) = S \<inter> T"
using r by auto
with conST show "connected (S \<inter> (T \<inter> cball 0 r \<union> sphere 0 r))"
by simp
show "connected_component (- (T \<inter> cball 0 r \<union> sphere 0 r)) a b"
using conST C r
apply (simp add: connected_component_def)
apply (rule_tac x=C in exI)
by auto
qed (simp add: ccS)
then obtain U where U: "connected U" "U \<subseteq> - S" "U \<subseteq> - T \<union> - cball 0 r" "U \<subseteq> - sphere 0 r" "a \<in> U" "b \<in> U"
by (auto simp: connected_component_def)
show ?thesis
unfolding connected_component_def
proof (intro exI conjI)
show "U \<subseteq> - (S \<union> T)"
using U r \<open>0 < r\<close> \<open>a \<in> C\<close> connected_Int_frontier [of U "cball 0 r"]
apply simp
by (metis ball_subset_cball compl_inf disjoint_eq_subset_Compl disjoint_iff_not_equal inf.orderE inf_sup_aci(3) subsetCE)
qed (auto simp: U)
qed
lemma Janiszewski_connected:
fixes S :: "complex set"
assumes ST: "compact S" "closed T" "connected(S \<inter> T)"
and notST: "connected (- S)" "connected (- T)"
shows "connected(- (S \<union> T))"
using Janiszewski [OF ST]
by (metis IntD1 IntD2 notST compl_sup connected_iff_connected_component)
subsection\<open>The Jordan Curve theorem\<close>
lemma exists_double_arc:
fixes g :: "real \<Rightarrow> 'a::real_normed_vector"
assumes "simple_path g" "pathfinish g = pathstart g" "a \<in> path_image g" "b \<in> path_image g" "a \<noteq> b"
obtains u d where "arc u" "arc d" "pathstart u = a" "pathfinish u = b"
"pathstart d = b" "pathfinish d = a"
"(path_image u) \<inter> (path_image d) = {a,b}"
"(path_image u) \<union> (path_image d) = path_image g"
proof -
obtain u where u: "0 \<le> u" "u \<le> 1" "g u = a"
using assms by (auto simp: path_image_def)
define h where "h \<equiv> shiftpath u g"
have "simple_path h"
using \<open>simple_path g\<close> simple_path_shiftpath \<open>0 \<le> u\<close> \<open>u \<le> 1\<close> assms(2) h_def by blast
have "pathstart h = g u"
by (simp add: \<open>u \<le> 1\<close> h_def pathstart_shiftpath)
have "pathfinish h = g u"
by (simp add: \<open>0 \<le> u\<close> assms h_def pathfinish_shiftpath)
have pihg: "path_image h = path_image g"
by (simp add: \<open>0 \<le> u\<close> \<open>u \<le> 1\<close> assms h_def path_image_shiftpath)
then obtain v where v: "0 \<le> v" "v \<le> 1" "h v = b"
using assms by (metis (mono_tags, lifting) atLeastAtMost_iff imageE path_image_def)
show ?thesis
proof
show "arc (subpath 0 v h)"
by (metis (no_types) \<open>pathstart h = g u\<close> \<open>simple_path h\<close> arc_simple_path_subpath \<open>a \<noteq> b\<close> atLeastAtMost_iff zero_le_one order_refl pathstart_def u(3) v)
show "arc (subpath v 1 h)"
by (metis (no_types) \<open>pathfinish h = g u\<close> \<open>simple_path h\<close> arc_simple_path_subpath \<open>a \<noteq> b\<close> atLeastAtMost_iff zero_le_one order_refl pathfinish_def u(3) v)
show "pathstart (subpath 0 v h) = a"
by (metis \<open>pathstart h = g u\<close> pathstart_def pathstart_subpath u(3))
show "pathfinish (subpath 0 v h) = b" "pathstart (subpath v 1 h) = b"
by (simp_all add: v(3))
show "pathfinish (subpath v 1 h) = a"
by (metis \<open>pathfinish h = g u\<close> pathfinish_def pathfinish_subpath u(3))
show "path_image (subpath 0 v h) \<inter> path_image (subpath v 1 h) = {a, b}"
proof
show "path_image (subpath 0 v h) \<inter> path_image (subpath v 1 h) \<subseteq> {a, b}"
using v \<open>pathfinish (subpath v 1 h) = a\<close> \<open>simple_path h\<close>
apply (auto simp: simple_path_def path_image_subpath image_iff Ball_def)
by (metis (full_types) less_eq_real_def less_irrefl less_le_trans)
show "{a, b} \<subseteq> path_image (subpath 0 v h) \<inter> path_image (subpath v 1 h)"
using v \<open>pathstart (subpath 0 v h) = a\<close> \<open>pathfinish (subpath v 1 h) = a\<close>
apply (auto simp: path_image_subpath image_iff)
by (metis atLeastAtMost_iff order_refl)
qed
show "path_image (subpath 0 v h) \<union> path_image (subpath v 1 h) = path_image g"
using v apply (simp add: path_image_subpath pihg [symmetric])
using path_image_def by fastforce
qed
qed
theorem\<^marker>\<open>tag unimportant\<close> Jordan_curve:
fixes c :: "real \<Rightarrow> complex"
assumes "simple_path c" and loop: "pathfinish c = pathstart c"
obtains inner outer where
"inner \<noteq> {}" "open inner" "connected inner"
"outer \<noteq> {}" "open outer" "connected outer"
"bounded inner" "\<not> bounded outer" "inner \<inter> outer = {}"
"inner \<union> outer = - path_image c"
"frontier inner = path_image c"
"frontier outer = path_image c"
proof -
have "path c"
by (simp add: assms simple_path_imp_path)
have hom: "(path_image c) homeomorphic (sphere(0::complex) 1)"
by (simp add: assms homeomorphic_simple_path_image_circle)
with Jordan_Brouwer_separation have "\<not> connected (- (path_image c))"
by fastforce
then obtain inner where inner: "inner \<in> components (- path_image c)" and "bounded inner"
using cobounded_has_bounded_component [of "- (path_image c)"]
using \<open>\<not> connected (- path_image c)\<close> \<open>simple_path c\<close> bounded_simple_path_image by force
obtain outer where outer: "outer \<in> components (- path_image c)" and "\<not> bounded outer"
using cobounded_unbounded_components [of "- (path_image c)"]
using \<open>path c\<close> bounded_path_image by auto
show ?thesis
proof
show "inner \<noteq> {}"
using inner in_components_nonempty by auto
show "open inner"
by (meson \<open>simple_path c\<close> compact_imp_closed compact_simple_path_image inner open_Compl open_components)
show "connected inner"
using in_components_connected inner by blast
show "outer \<noteq> {}"
using outer in_components_nonempty by auto
show "open outer"
by (meson \<open>simple_path c\<close> compact_imp_closed compact_simple_path_image outer open_Compl open_components)
show "connected outer"
using in_components_connected outer by blast
show "inner \<inter> outer = {}"
by (meson \<open>\<not> bounded outer\<close> \<open>bounded inner\<close> \<open>connected outer\<close> bounded_subset components_maximal in_components_subset inner outer)
show fro_inner: "frontier inner = path_image c"
by (simp add: Jordan_Brouwer_frontier [OF hom inner])
show fro_outer: "frontier outer = path_image c"
by (simp add: Jordan_Brouwer_frontier [OF hom outer])
have False if m: "middle \<in> components (- path_image c)" and "middle \<noteq> inner" "middle \<noteq> outer" for middle
proof -
have "frontier middle = path_image c"
by (simp add: Jordan_Brouwer_frontier [OF hom] that)
have middle: "open middle" "connected middle" "middle \<noteq> {}"
apply (meson \<open>simple_path c\<close> compact_imp_closed compact_simple_path_image m open_Compl open_components)
using in_components_connected in_components_nonempty m by blast+
obtain a0 b0 where "a0 \<in> path_image c" "b0 \<in> path_image c" "a0 \<noteq> b0"
using simple_path_image_uncountable [OF \<open>simple_path c\<close>]
by (metis Diff_cancel countable_Diff_eq countable_empty insert_iff subsetI subset_singleton_iff)
obtain a b g where ab: "a \<in> path_image c" "b \<in> path_image c" "a \<noteq> b"
and "arc g" "pathstart g = a" "pathfinish g = b"
and pag_sub: "path_image g - {a,b} \<subseteq> middle"
proof (rule dense_accessible_frontier_point_pairs [OF \<open>open middle\<close> \<open>connected middle\<close>, of "path_image c \<inter> ball a0 (dist a0 b0)" "path_image c \<inter> ball b0 (dist a0 b0)"])
show "openin (top_of_set (frontier middle)) (path_image c \<inter> ball a0 (dist a0 b0))"
"openin (top_of_set (frontier middle)) (path_image c \<inter> ball b0 (dist a0 b0))"
by (simp_all add: \<open>frontier middle = path_image c\<close> openin_open_Int)
show "path_image c \<inter> ball a0 (dist a0 b0) \<noteq> path_image c \<inter> ball b0 (dist a0 b0)"
using \<open>a0 \<noteq> b0\<close> \<open>b0 \<in> path_image c\<close> by auto
show "path_image c \<inter> ball a0 (dist a0 b0) \<noteq> {}"
using \<open>a0 \<in> path_image c\<close> \<open>a0 \<noteq> b0\<close> by auto
show "path_image c \<inter> ball b0 (dist a0 b0) \<noteq> {}"
using \<open>b0 \<in> path_image c\<close> \<open>a0 \<noteq> b0\<close> by auto
qed (use arc_distinct_ends arc_imp_simple_path simple_path_endless that in fastforce)
obtain u d where "arc u" "arc d"
and "pathstart u = a" "pathfinish u = b" "pathstart d = b" "pathfinish d = a"
and ud_ab: "(path_image u) \<inter> (path_image d) = {a,b}"
and ud_Un: "(path_image u) \<union> (path_image d) = path_image c"
using exists_double_arc [OF assms ab] by blast
obtain x y where "x \<in> inner" "y \<in> outer"
using \<open>inner \<noteq> {}\<close> \<open>outer \<noteq> {}\<close> by auto
have "inner \<inter> middle = {}" "middle \<inter> outer = {}"
using components_nonoverlap inner outer m that by blast+
have "connected_component (- (path_image u \<union> path_image g \<union> (path_image d \<union> path_image g))) x y"
proof (rule Janiszewski)
show "compact (path_image u \<union> path_image g)"
by (simp add: \<open>arc g\<close> \<open>arc u\<close> compact_Un compact_arc_image)
show "closed (path_image d \<union> path_image g)"
by (simp add: \<open>arc d\<close> \<open>arc g\<close> closed_Un closed_arc_image)
show "connected ((path_image u \<union> path_image g) \<inter> (path_image d \<union> path_image g))"
by (metis Un_Diff_cancel \<open>arc g\<close> \<open>path_image u \<inter> path_image d = {a, b}\<close> \<open>pathfinish g = b\<close> \<open>pathstart g = a\<close> connected_arc_image insert_Diff1 pathfinish_in_path_image pathstart_in_path_image sup_bot.right_neutral sup_commute sup_inf_distrib1)
show "connected_component (- (path_image u \<union> path_image g)) x y"
unfolding connected_component_def
proof (intro exI conjI)
have "connected ((inner \<union> (path_image c - path_image u)) \<union> (outer \<union> (path_image c - path_image u)))"
proof (rule connected_Un)
show "connected (inner \<union> (path_image c - path_image u))"
apply (rule connected_intermediate_closure [OF \<open>connected inner\<close>])
using fro_inner [symmetric] apply (auto simp: closure_subset frontier_def)
done
show "connected (outer \<union> (path_image c - path_image u))"
apply (rule connected_intermediate_closure [OF \<open>connected outer\<close>])
using fro_outer [symmetric] apply (auto simp: closure_subset frontier_def)
done
have "(inner \<inter> outer) \<union> (path_image c - path_image u) \<noteq> {}"
by (metis \<open>arc d\<close> ud_ab Diff_Int Diff_cancel Un_Diff \<open>inner \<inter> outer = {}\<close> \<open>pathfinish d = a\<close> \<open>pathstart d = b\<close> arc_simple_path insert_commute nonempty_simple_path_endless sup_bot_left ud_Un)
then show "(inner \<union> (path_image c - path_image u)) \<inter> (outer \<union> (path_image c - path_image u)) \<noteq> {}"
by auto
qed
then show "connected (inner \<union> outer \<union> (path_image c - path_image u))"
by (metis sup.right_idem sup_assoc sup_commute)
have "inner \<subseteq> - path_image u" "outer \<subseteq> - path_image u"
using in_components_subset inner outer ud_Un by auto
moreover have "inner \<subseteq> - path_image g" "outer \<subseteq> - path_image g"
using \<open>inner \<inter> middle = {}\<close> \<open>inner \<subseteq> - path_image u\<close>
using \<open>middle \<inter> outer = {}\<close> \<open>outer \<subseteq> - path_image u\<close> pag_sub ud_ab by fastforce+
moreover have "path_image c - path_image u \<subseteq> - path_image g"
using in_components_subset m pag_sub ud_ab by fastforce
ultimately show "inner \<union> outer \<union> (path_image c - path_image u) \<subseteq> - (path_image u \<union> path_image g)"
by force
show "x \<in> inner \<union> outer \<union> (path_image c - path_image u)"
by (auto simp: \<open>x \<in> inner\<close>)
show "y \<in> inner \<union> outer \<union> (path_image c - path_image u)"
by (auto simp: \<open>y \<in> outer\<close>)
qed
show "connected_component (- (path_image d \<union> path_image g)) x y"
unfolding connected_component_def
proof (intro exI conjI)
have "connected ((inner \<union> (path_image c - path_image d)) \<union> (outer \<union> (path_image c - path_image d)))"
proof (rule connected_Un)
show "connected (inner \<union> (path_image c - path_image d))"
apply (rule connected_intermediate_closure [OF \<open>connected inner\<close>])
using fro_inner [symmetric] apply (auto simp: closure_subset frontier_def)
done
show "connected (outer \<union> (path_image c - path_image d))"
apply (rule connected_intermediate_closure [OF \<open>connected outer\<close>])
using fro_outer [symmetric] apply (auto simp: closure_subset frontier_def)
done
have "(inner \<inter> outer) \<union> (path_image c - path_image d) \<noteq> {}"
using \<open>arc u\<close> \<open>pathfinish u = b\<close> \<open>pathstart u = a\<close> arc_imp_simple_path nonempty_simple_path_endless ud_Un ud_ab by fastforce
then show "(inner \<union> (path_image c - path_image d)) \<inter> (outer \<union> (path_image c - path_image d)) \<noteq> {}"
by auto
qed
then show "connected (inner \<union> outer \<union> (path_image c - path_image d))"
by (metis sup.right_idem sup_assoc sup_commute)
have "inner \<subseteq> - path_image d" "outer \<subseteq> - path_image d"
using in_components_subset inner outer ud_Un by auto
moreover have "inner \<subseteq> - path_image g" "outer \<subseteq> - path_image g"
using \<open>inner \<inter> middle = {}\<close> \<open>inner \<subseteq> - path_image d\<close>
using \<open>middle \<inter> outer = {}\<close> \<open>outer \<subseteq> - path_image d\<close> pag_sub ud_ab by fastforce+
moreover have "path_image c - path_image d \<subseteq> - path_image g"
using in_components_subset m pag_sub ud_ab by fastforce
ultimately show "inner \<union> outer \<union> (path_image c - path_image d) \<subseteq> - (path_image d \<union> path_image g)"
by force
show "x \<in> inner \<union> outer \<union> (path_image c - path_image d)"
by (auto simp: \<open>x \<in> inner\<close>)
show "y \<in> inner \<union> outer \<union> (path_image c - path_image d)"
by (auto simp: \<open>y \<in> outer\<close>)
qed
qed
then have "connected_component (- (path_image u \<union> path_image d \<union> path_image g)) x y"
by (simp add: Un_ac)
moreover have "\<not>(connected_component (- (path_image c)) x y)"
by (metis (no_types, lifting) \<open>\<not> bounded outer\<close> \<open>bounded inner\<close> \<open>x \<in> inner\<close> \<open>y \<in> outer\<close> componentsE connected_component_eq inner mem_Collect_eq outer)
ultimately show False
by (auto simp: ud_Un [symmetric] connected_component_def)
qed
then have "components (- path_image c) = {inner,outer}"
using inner outer by blast
then have "Union (components (- path_image c)) = inner \<union> outer"
by simp
then show "inner \<union> outer = - path_image c"
by auto
qed (auto simp: \<open>bounded inner\<close> \<open>\<not> bounded outer\<close>)
qed
corollary\<^marker>\<open>tag unimportant\<close> Jordan_disconnected:
fixes c :: "real \<Rightarrow> complex"
assumes "simple_path c" "pathfinish c = pathstart c"
shows "\<not> connected(- path_image c)"
using Jordan_curve [OF assms]
by (metis Jordan_Brouwer_separation assms homeomorphic_simple_path_image_circle zero_less_one)
corollary Jordan_inside_outside:
fixes c :: "real \<Rightarrow> complex"
assumes "simple_path c" "pathfinish c = pathstart c"
shows "inside(path_image c) \<noteq> {} \<and>
open(inside(path_image c)) \<and>
connected(inside(path_image c)) \<and>
outside(path_image c) \<noteq> {} \<and>
open(outside(path_image c)) \<and>
connected(outside(path_image c)) \<and>
bounded(inside(path_image c)) \<and>
\<not> bounded(outside(path_image c)) \<and>
inside(path_image c) \<inter> outside(path_image c) = {} \<and>
inside(path_image c) \<union> outside(path_image c) =
- path_image c \<and>
frontier(inside(path_image c)) = path_image c \<and>
frontier(outside(path_image c)) = path_image c"
proof -
obtain inner outer
where *: "inner \<noteq> {}" "open inner" "connected inner"
"outer \<noteq> {}" "open outer" "connected outer"
"bounded inner" "\<not> bounded outer" "inner \<inter> outer = {}"
"inner \<union> outer = - path_image c"
"frontier inner = path_image c"
"frontier outer = path_image c"
using Jordan_curve [OF assms] by blast
then have inner: "inside(path_image c) = inner"
by (metis dual_order.antisym inside_subset interior_eq interior_inside_frontier)
have outer: "outside(path_image c) = outer"
using \<open>inner \<union> outer = - path_image c\<close> \<open>inside (path_image c) = inner\<close>
outside_inside \<open>inner \<inter> outer = {}\<close> by auto
show ?thesis
using * by (auto simp: inner outer)
qed
subsubsection\<open>Triple-curve or "theta-curve" theorem\<close>
text\<open>Proof that there is no fourth component taken from
Kuratowski's Topology vol 2, para 61, II.\<close>
theorem split_inside_simple_closed_curve:
fixes c :: "real \<Rightarrow> complex"
assumes "simple_path c1" and c1: "pathstart c1 = a" "pathfinish c1 = b"
and "simple_path c2" and c2: "pathstart c2 = a" "pathfinish c2 = b"
and "simple_path c" and c: "pathstart c = a" "pathfinish c = b"
and "a \<noteq> b"
and c1c2: "path_image c1 \<inter> path_image c2 = {a,b}"
and c1c: "path_image c1 \<inter> path_image c = {a,b}"
and c2c: "path_image c2 \<inter> path_image c = {a,b}"
and ne_12: "path_image c \<inter> inside(path_image c1 \<union> path_image c2) \<noteq> {}"
obtains "inside(path_image c1 \<union> path_image c) \<inter> inside(path_image c2 \<union> path_image c) = {}"
"inside(path_image c1 \<union> path_image c) \<union> inside(path_image c2 \<union> path_image c) \<union>
(path_image c - {a,b}) = inside(path_image c1 \<union> path_image c2)"
proof -
let ?\<Theta> = "path_image c" let ?\<Theta>1 = "path_image c1" let ?\<Theta>2 = "path_image c2"
have sp: "simple_path (c1 +++ reversepath c2)" "simple_path (c1 +++ reversepath c)" "simple_path (c2 +++ reversepath c)"
using assms by (auto simp: simple_path_join_loop_eq arc_simple_path simple_path_reversepath)
then have op_in12: "open (inside (?\<Theta>1 \<union> ?\<Theta>2))"
and op_out12: "open (outside (?\<Theta>1 \<union> ?\<Theta>2))"
and op_in1c: "open (inside (?\<Theta>1 \<union> ?\<Theta>))"
and op_in2c: "open (inside (?\<Theta>2 \<union> ?\<Theta>))"
and op_out1c: "open (outside (?\<Theta>1 \<union> ?\<Theta>))"
and op_out2c: "open (outside (?\<Theta>2 \<union> ?\<Theta>))"
and co_in1c: "connected (inside (?\<Theta>1 \<union> ?\<Theta>))"
and co_in2c: "connected (inside (?\<Theta>2 \<union> ?\<Theta>))"
and co_out12c: "connected (outside (?\<Theta>1 \<union> ?\<Theta>2))"
and co_out1c: "connected (outside (?\<Theta>1 \<union> ?\<Theta>))"
and co_out2c: "connected (outside (?\<Theta>2 \<union> ?\<Theta>))"
and pa_c: "?\<Theta> - {pathstart c, pathfinish c} \<subseteq> - ?\<Theta>1"
"?\<Theta> - {pathstart c, pathfinish c} \<subseteq> - ?\<Theta>2"
and pa_c1: "?\<Theta>1 - {pathstart c1, pathfinish c1} \<subseteq> - ?\<Theta>2"
"?\<Theta>1 - {pathstart c1, pathfinish c1} \<subseteq> - ?\<Theta>"
and pa_c2: "?\<Theta>2 - {pathstart c2, pathfinish c2} \<subseteq> - ?\<Theta>1"
"?\<Theta>2 - {pathstart c2, pathfinish c2} \<subseteq> - ?\<Theta>"
and co_c: "connected(?\<Theta> - {pathstart c,pathfinish c})"
and co_c1: "connected(?\<Theta>1 - {pathstart c1,pathfinish c1})"
and co_c2: "connected(?\<Theta>2 - {pathstart c2,pathfinish c2})"
and fr_in: "frontier(inside(?\<Theta>1 \<union> ?\<Theta>2)) = ?\<Theta>1 \<union> ?\<Theta>2"
"frontier(inside(?\<Theta>2 \<union> ?\<Theta>)) = ?\<Theta>2 \<union> ?\<Theta>"
"frontier(inside(?\<Theta>1 \<union> ?\<Theta>)) = ?\<Theta>1 \<union> ?\<Theta>"
and fr_out: "frontier(outside(?\<Theta>1 \<union> ?\<Theta>2)) = ?\<Theta>1 \<union> ?\<Theta>2"
"frontier(outside(?\<Theta>2 \<union> ?\<Theta>)) = ?\<Theta>2 \<union> ?\<Theta>"
"frontier(outside(?\<Theta>1 \<union> ?\<Theta>)) = ?\<Theta>1 \<union> ?\<Theta>"
using Jordan_inside_outside [of "c1 +++ reversepath c2"]
using Jordan_inside_outside [of "c1 +++ reversepath c"]
using Jordan_inside_outside [of "c2 +++ reversepath c"] assms
apply (simp_all add: path_image_join closed_Un closed_simple_path_image open_inside open_outside)
apply (blast elim: | metis connected_simple_path_endless)+
done
have inout_12: "inside (?\<Theta>1 \<union> ?\<Theta>2) \<inter> (?\<Theta> - {pathstart c, pathfinish c}) \<noteq> {}"
by (metis (no_types, lifting) c c1c ne_12 Diff_Int_distrib Diff_empty Int_empty_right Int_left_commute inf_sup_absorb inf_sup_aci(1) inside_no_overlap)
have pi_disjoint: "?\<Theta> \<inter> outside(?\<Theta>1 \<union> ?\<Theta>2) = {}"
proof (rule ccontr)
assume "?\<Theta> \<inter> outside (?\<Theta>1 \<union> ?\<Theta>2) \<noteq> {}"
then show False
using connectedD [OF co_c, of "inside(?\<Theta>1 \<union> ?\<Theta>2)" "outside(?\<Theta>1 \<union> ?\<Theta>2)"]
using c c1c2 pa_c op_in12 op_out12 inout_12
apply auto
apply (metis Un_Diff_cancel2 Un_iff compl_sup disjoint_insert(1) inf_commute inf_compl_bot_left2 inside_Un_outside mk_disjoint_insert sup_inf_absorb)
done
qed
have out_sub12: "outside(?\<Theta>1 \<union> ?\<Theta>2) \<subseteq> outside(?\<Theta>1 \<union> ?\<Theta>)" "outside(?\<Theta>1 \<union> ?\<Theta>2) \<subseteq> outside(?\<Theta>2 \<union> ?\<Theta>)"
by (metis Un_commute pi_disjoint outside_Un_outside_Un)+
have pa1_disj_in2: "?\<Theta>1 \<inter> inside (?\<Theta>2 \<union> ?\<Theta>) = {}"
proof (rule ccontr)
assume ne: "?\<Theta>1 \<inter> inside (?\<Theta>2 \<union> ?\<Theta>) \<noteq> {}"
have 1: "inside (?\<Theta> \<union> ?\<Theta>2) \<inter> ?\<Theta> = {}"
by (metis (no_types) Diff_Int_distrib Diff_cancel inf_sup_absorb inf_sup_aci(3) inside_no_overlap)
have 2: "outside (?\<Theta> \<union> ?\<Theta>2) \<inter> ?\<Theta> = {}"
by (metis (no_types) Int_empty_right Int_left_commute inf_sup_absorb outside_no_overlap)
have "outside (?\<Theta>2 \<union> ?\<Theta>) \<subseteq> outside (?\<Theta>1 \<union> ?\<Theta>2)"
apply (subst Un_commute, rule outside_Un_outside_Un)
using connectedD [OF co_c1, of "inside(?\<Theta>2 \<union> ?\<Theta>)" "outside(?\<Theta>2 \<union> ?\<Theta>)"]
pa_c1 op_in2c op_out2c ne c1 c2c 1 2 by (auto simp: inf_sup_aci)
with out_sub12
have "outside(?\<Theta>1 \<union> ?\<Theta>2) = outside(?\<Theta>2 \<union> ?\<Theta>)" by blast
then have "frontier(outside(?\<Theta>1 \<union> ?\<Theta>2)) = frontier(outside(?\<Theta>2 \<union> ?\<Theta>))"
by simp
then show False
using inout_12 pi_disjoint c c1c c2c fr_out by auto
qed
have pa2_disj_in1: "?\<Theta>2 \<inter> inside(?\<Theta>1 \<union> ?\<Theta>) = {}"
proof (rule ccontr)
assume ne: "?\<Theta>2 \<inter> inside (?\<Theta>1 \<union> ?\<Theta>) \<noteq> {}"
have 1: "inside (?\<Theta> \<union> ?\<Theta>1) \<inter> ?\<Theta> = {}"
by (metis (no_types) Diff_Int_distrib Diff_cancel inf_sup_absorb inf_sup_aci(3) inside_no_overlap)
have 2: "outside (?\<Theta> \<union> ?\<Theta>1) \<inter> ?\<Theta> = {}"
by (metis (no_types) Int_empty_right Int_left_commute inf_sup_absorb outside_no_overlap)
have "outside (?\<Theta>1 \<union> ?\<Theta>) \<subseteq> outside (?\<Theta>1 \<union> ?\<Theta>2)"
apply (rule outside_Un_outside_Un)
using connectedD [OF co_c2, of "inside(?\<Theta>1 \<union> ?\<Theta>)" "outside(?\<Theta>1 \<union> ?\<Theta>)"]
pa_c2 op_in1c op_out1c ne c2 c1c 1 2 by (auto simp: inf_sup_aci)
with out_sub12
have "outside(?\<Theta>1 \<union> ?\<Theta>2) = outside(?\<Theta>1 \<union> ?\<Theta>)"
by blast
then have "frontier(outside(?\<Theta>1 \<union> ?\<Theta>2)) = frontier(outside(?\<Theta>1 \<union> ?\<Theta>))"
by simp
then show False
using inout_12 pi_disjoint c c1c c2c fr_out by auto
qed
have in_sub_in1: "inside(?\<Theta>1 \<union> ?\<Theta>) \<subseteq> inside(?\<Theta>1 \<union> ?\<Theta>2)"
using pa2_disj_in1 out_sub12 by (auto simp: inside_outside)
have in_sub_in2: "inside(?\<Theta>2 \<union> ?\<Theta>) \<subseteq> inside(?\<Theta>1 \<union> ?\<Theta>2)"
using pa1_disj_in2 out_sub12 by (auto simp: inside_outside)
have in_sub_out12: "inside(?\<Theta>1 \<union> ?\<Theta>) \<subseteq> outside(?\<Theta>2 \<union> ?\<Theta>)"
proof
fix x
assume x: "x \<in> inside (?\<Theta>1 \<union> ?\<Theta>)"
then have xnot: "x \<notin> ?\<Theta>"
by (simp add: inside_def)
obtain z where zim: "z \<in> ?\<Theta>1" and zout: "z \<in> outside(?\<Theta>2 \<union> ?\<Theta>)"
apply (auto simp: outside_inside)
using nonempty_simple_path_endless [OF \<open>simple_path c1\<close>]
by (metis Diff_Diff_Int Diff_iff ex_in_conv c1 c1c c1c2 pa1_disj_in2)
obtain e where "e > 0" and e: "ball z e \<subseteq> outside(?\<Theta>2 \<union> ?\<Theta>)"
using zout op_out2c open_contains_ball_eq by blast
have "z \<in> frontier (inside (?\<Theta>1 \<union> ?\<Theta>))"
using zim by (auto simp: fr_in)
then obtain w where w1: "w \<in> inside (?\<Theta>1 \<union> ?\<Theta>)" and dwz: "dist w z < e"
using zim \<open>e > 0\<close> by (auto simp: frontier_def closure_approachable)
then have w2: "w \<in> outside (?\<Theta>2 \<union> ?\<Theta>)"
by (metis e dist_commute mem_ball subsetCE)
then have "connected_component (- ?\<Theta>2 \<inter> - ?\<Theta>) z w"
apply (simp add: connected_component_def)
apply (rule_tac x = "outside(?\<Theta>2 \<union> ?\<Theta>)" in exI)
using zout apply (auto simp: co_out2c)
apply (simp_all add: outside_inside)
done
moreover have "connected_component (- ?\<Theta>2 \<inter> - ?\<Theta>) w x"
unfolding connected_component_def
using pa2_disj_in1 co_in1c x w1 union_with_outside by fastforce
ultimately have eq: "connected_component_set (- ?\<Theta>2 \<inter> - ?\<Theta>) x =
connected_component_set (- ?\<Theta>2 \<inter> - ?\<Theta>) z"
by (metis (mono_tags, lifting) connected_component_eq mem_Collect_eq)
show "x \<in> outside (?\<Theta>2 \<union> ?\<Theta>)"
using zout x pa2_disj_in1 by (auto simp: outside_def eq xnot)
qed
have in_sub_out21: "inside(?\<Theta>2 \<union> ?\<Theta>) \<subseteq> outside(?\<Theta>1 \<union> ?\<Theta>)"
proof
fix x
assume x: "x \<in> inside (?\<Theta>2 \<union> ?\<Theta>)"
then have xnot: "x \<notin> ?\<Theta>"
by (simp add: inside_def)
obtain z where zim: "z \<in> ?\<Theta>2" and zout: "z \<in> outside(?\<Theta>1 \<union> ?\<Theta>)"
apply (auto simp: outside_inside)
using nonempty_simple_path_endless [OF \<open>simple_path c2\<close>]
by (metis (no_types, hide_lams) Diff_Diff_Int Diff_iff c1c2 c2 c2c ex_in_conv pa2_disj_in1)
obtain e where "e > 0" and e: "ball z e \<subseteq> outside(?\<Theta>1 \<union> ?\<Theta>)"
using zout op_out1c open_contains_ball_eq by blast
have "z \<in> frontier (inside (?\<Theta>2 \<union> ?\<Theta>))"
using zim by (auto simp: fr_in)
then obtain w where w2: "w \<in> inside (?\<Theta>2 \<union> ?\<Theta>)" and dwz: "dist w z < e"
using zim \<open>e > 0\<close> by (auto simp: frontier_def closure_approachable)
then have w1: "w \<in> outside (?\<Theta>1 \<union> ?\<Theta>)"
by (metis e dist_commute mem_ball subsetCE)
then have "connected_component (- ?\<Theta>1 \<inter> - ?\<Theta>) z w"
apply (simp add: connected_component_def)
apply (rule_tac x = "outside(?\<Theta>1 \<union> ?\<Theta>)" in exI)
using zout apply (auto simp: co_out1c)
apply (simp_all add: outside_inside)
done
moreover have "connected_component (- ?\<Theta>1 \<inter> - ?\<Theta>) w x"
unfolding connected_component_def
using pa1_disj_in2 co_in2c x w2 union_with_outside by fastforce
ultimately have eq: "connected_component_set (- ?\<Theta>1 \<inter> - ?\<Theta>) x =
connected_component_set (- ?\<Theta>1 \<inter> - ?\<Theta>) z"
by (metis (no_types, lifting) connected_component_eq mem_Collect_eq)
show "x \<in> outside (?\<Theta>1 \<union> ?\<Theta>)"
using zout x pa1_disj_in2 by (auto simp: outside_def eq xnot)
qed
show ?thesis
proof
show "inside (?\<Theta>1 \<union> ?\<Theta>) \<inter> inside (?\<Theta>2 \<union> ?\<Theta>) = {}"
by (metis Int_Un_distrib in_sub_out12 bot_eq_sup_iff disjoint_eq_subset_Compl outside_inside)
have *: "outside (?\<Theta>1 \<union> ?\<Theta>) \<inter> outside (?\<Theta>2 \<union> ?\<Theta>) \<subseteq> outside (?\<Theta>1 \<union> ?\<Theta>2)"
proof (rule components_maximal)
show out_in: "outside (?\<Theta>1 \<union> ?\<Theta>2) \<in> components (- (?\<Theta>1 \<union> ?\<Theta>2))"
apply (simp only: outside_in_components co_out12c)
by (metis bounded_empty fr_out(1) frontier_empty unbounded_outside)
have conn_U: "connected (- (closure (inside (?\<Theta>1 \<union> ?\<Theta>)) \<union> closure (inside (?\<Theta>2 \<union> ?\<Theta>))))"
proof (rule Janiszewski_connected, simp_all)
show "bounded (inside (?\<Theta>1 \<union> ?\<Theta>))"
by (simp add: \<open>simple_path c1\<close> \<open>simple_path c\<close> bounded_inside bounded_simple_path_image)
have if1: "- (inside (?\<Theta>1 \<union> ?\<Theta>) \<union> frontier (inside (?\<Theta>1 \<union> ?\<Theta>))) = - ?\<Theta>1 \<inter> - ?\<Theta> \<inter> - inside (?\<Theta>1 \<union> ?\<Theta>)"
by (metis (no_types, lifting) Int_commute Jordan_inside_outside c c1 compl_sup path_image_join path_image_reversepath pathfinish_join pathfinish_reversepath pathstart_join pathstart_reversepath sp(2) closure_Un_frontier fr_out(3))
then show "connected (- closure (inside (?\<Theta>1 \<union> ?\<Theta>)))"
by (metis Compl_Un outside_inside co_out1c closure_Un_frontier)
have if2: "- (inside (?\<Theta>2 \<union> ?\<Theta>) \<union> frontier (inside (?\<Theta>2 \<union> ?\<Theta>))) = - ?\<Theta>2 \<inter> - ?\<Theta> \<inter> - inside (?\<Theta>2 \<union> ?\<Theta>)"
- by (metis (no_types, lifting) Int_commute Jordan_inside_outside c c2 compl_sup path_image_join path_image_reversepath pathfinish_join pathfinish_reversepath pathstart_join pathstart_reversepath sp closure_Un_frontier fr_out(2))
+ by (metis (no_types, lifting) Int_commute Jordan_inside_outside c c2 compl_sup path_image_join path_image_reversepath pathfinish_join pathfinish_reversepath pathstart_join pathstart_reversepath sp(3) closure_Un_frontier fr_out(2))
then show "connected (- closure (inside (?\<Theta>2 \<union> ?\<Theta>)))"
by (metis Compl_Un outside_inside co_out2c closure_Un_frontier)
have "connected(?\<Theta>)"
by (metis \<open>simple_path c\<close> connected_simple_path_image)
moreover
have "closure (inside (?\<Theta>1 \<union> ?\<Theta>)) \<inter> closure (inside (?\<Theta>2 \<union> ?\<Theta>)) = ?\<Theta>"
(is "?lhs = ?rhs")
proof
show "?lhs \<subseteq> ?rhs"
proof clarify
fix x
assume x: "x \<in> closure (inside (?\<Theta>1 \<union> ?\<Theta>))" "x \<in> closure (inside (?\<Theta>2 \<union> ?\<Theta>))"
then have "x \<notin> inside (?\<Theta>1 \<union> ?\<Theta>)"
by (meson closure_iff_nhds_not_empty in_sub_out12 inside_Int_outside op_in1c)
with fr_in x show "x \<in> ?\<Theta>"
by (metis c1c c1c2 closure_Un_frontier pa1_disj_in2 Int_iff Un_iff insert_disjoint(2) insert_subset subsetI subset_antisym)
qed
show "?rhs \<subseteq> ?lhs"
using if1 if2 closure_Un_frontier by fastforce
qed
ultimately
show "connected (closure (inside (?\<Theta>1 \<union> ?\<Theta>)) \<inter> closure (inside (?\<Theta>2 \<union> ?\<Theta>)))"
by auto
qed
show "connected (outside (?\<Theta>1 \<union> ?\<Theta>) \<inter> outside (?\<Theta>2 \<union> ?\<Theta>))"
using fr_in conn_U by (simp add: closure_Un_frontier outside_inside Un_commute)
show "outside (?\<Theta>1 \<union> ?\<Theta>) \<inter> outside (?\<Theta>2 \<union> ?\<Theta>) \<subseteq> - (?\<Theta>1 \<union> ?\<Theta>2)"
by clarify (metis Diff_Compl Diff_iff Un_iff inf_sup_absorb outside_inside)
show "outside (?\<Theta>1 \<union> ?\<Theta>2) \<inter>
(outside (?\<Theta>1 \<union> ?\<Theta>) \<inter> outside (?\<Theta>2 \<union> ?\<Theta>)) \<noteq> {}"
by (metis Int_assoc out_in inf.orderE out_sub12(1) out_sub12(2) outside_in_components)
qed
show "inside (?\<Theta>1 \<union> ?\<Theta>) \<union> inside (?\<Theta>2 \<union> ?\<Theta>) \<union> (?\<Theta> - {a, b}) = inside (?\<Theta>1 \<union> ?\<Theta>2)"
(is "?lhs = ?rhs")
proof
show "?lhs \<subseteq> ?rhs"
apply (simp add: in_sub_in1 in_sub_in2)
using c1c c2c inside_outside pi_disjoint by fastforce
have "inside (?\<Theta>1 \<union> ?\<Theta>2) \<subseteq> inside (?\<Theta>1 \<union> ?\<Theta>) \<union> inside (?\<Theta>2 \<union> ?\<Theta>) \<union> (?\<Theta>)"
using Compl_anti_mono [OF *] by (force simp: inside_outside)
moreover have "inside (?\<Theta>1 \<union> ?\<Theta>2) \<subseteq> -{a,b}"
using c1 union_with_outside by fastforce
ultimately show "?rhs \<subseteq> ?lhs" by auto
qed
qed
qed
end
diff --git a/src/HOL/Analysis/Lipschitz.thy b/src/HOL/Analysis/Lipschitz.thy
--- a/src/HOL/Analysis/Lipschitz.thy
+++ b/src/HOL/Analysis/Lipschitz.thy
@@ -1,836 +1,836 @@
(* Title: HOL/Analysis/Lipschitz.thy
Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr
Author: Fabian Immler, TU München
*)
section \<open>Lipschitz Continuity\<close>
theory Lipschitz
imports
Derivative
begin
definition\<^marker>\<open>tag important\<close> lipschitz_on
where "lipschitz_on C U f \<longleftrightarrow> (0 \<le> C \<and> (\<forall>x \<in> U. \<forall>y\<in>U. dist (f x) (f y) \<le> C * dist x y))"
bundle lipschitz_syntax begin
notation\<^marker>\<open>tag important\<close> lipschitz_on ("_-lipschitz'_on" [1000])
end
bundle no_lipschitz_syntax begin
no_notation lipschitz_on ("_-lipschitz'_on" [1000])
end
unbundle lipschitz_syntax
lemma lipschitz_onI: "L-lipschitz_on X f"
if "\<And>x y. x \<in> X \<Longrightarrow> y \<in> X \<Longrightarrow> dist (f x) (f y) \<le> L * dist x y" "0 \<le> L"
using that by (auto simp: lipschitz_on_def)
lemma lipschitz_onD:
"dist (f x) (f y) \<le> L * dist x y"
if "L-lipschitz_on X f" "x \<in> X" "y \<in> X"
using that by (auto simp: lipschitz_on_def)
lemma lipschitz_on_nonneg:
"0 \<le> L" if "L-lipschitz_on X f"
using that by (auto simp: lipschitz_on_def)
lemma lipschitz_on_normD:
"norm (f x - f y) \<le> L * norm (x - y)"
if "lipschitz_on L X f" "x \<in> X" "y \<in> X"
using lipschitz_onD[OF that]
by (simp add: dist_norm)
lemma lipschitz_on_mono: "L-lipschitz_on D f" if "M-lipschitz_on E f" "D \<subseteq> E" "M \<le> L"
using that
by (force simp: lipschitz_on_def intro: order_trans[OF _ mult_right_mono])
lemmas lipschitz_on_subset = lipschitz_on_mono[OF _ _ order_refl]
and lipschitz_on_le = lipschitz_on_mono[OF _ order_refl]
lemma lipschitz_on_leI:
"L-lipschitz_on X f"
if "\<And>x y. x \<in> X \<Longrightarrow> y \<in> X \<Longrightarrow> x \<le> y \<Longrightarrow> dist (f x) (f y) \<le> L * dist x y"
"0 \<le> L"
for f::"'a::{linorder_topology, ordered_real_vector, metric_space} \<Rightarrow> 'b::metric_space"
proof (rule lipschitz_onI)
fix x y assume xy: "x \<in> X" "y \<in> X"
consider "y \<le> x" | "x \<le> y"
by (rule le_cases)
then show "dist (f x) (f y) \<le> L * dist x y"
proof cases
case 1
then have "dist (f y) (f x) \<le> L * dist y x"
by (auto intro!: that xy)
then show ?thesis
by (simp add: dist_commute)
qed (auto intro!: that xy)
qed fact
lemma lipschitz_on_concat:
fixes a b c::real
assumes f: "L-lipschitz_on {a .. b} f"
assumes g: "L-lipschitz_on {b .. c} g"
assumes fg: "f b = g b"
shows "lipschitz_on L {a .. c} (\<lambda>x. if x \<le> b then f x else g x)"
(is "lipschitz_on _ _ ?f")
proof (rule lipschitz_on_leI)
fix x y
assume x: "x \<in> {a..c}" and y: "y \<in> {a..c}" and xy: "x \<le> y"
consider "x \<le> b \<and> b < y" | "x \<ge> b \<or> y \<le> b" by arith
then show "dist (?f x) (?f y) \<le> L * dist x y"
proof cases
case 1
have "dist (f x) (g y) \<le> dist (f x) (f b) + dist (g b) (g y)"
unfolding fg by (rule dist_triangle)
also have "dist (f x) (f b) \<le> L * dist x b"
using 1 x
by (auto intro!: lipschitz_onD[OF f])
also have "dist (g b) (g y) \<le> L * dist b y"
using 1 x y
by (auto intro!: lipschitz_onD[OF g] lipschitz_onD[OF f])
finally have "dist (f x) (g y) \<le> L * dist x b + L * dist b y"
by simp
also have "\<dots> = L * (dist x b + dist b y)"
by (simp add: algebra_simps)
also have "dist x b + dist b y = dist x y"
using 1 x y
by (auto simp: dist_real_def abs_real_def)
finally show ?thesis
using 1 by simp
next
case 2
with lipschitz_onD[OF f, of x y] lipschitz_onD[OF g, of x y] x y xy
show ?thesis
by (auto simp: fg)
qed
qed (rule lipschitz_on_nonneg[OF f])
lemma lipschitz_on_concat_max:
fixes a b c::real
assumes f: "L-lipschitz_on {a .. b} f"
assumes g: "M-lipschitz_on {b .. c} g"
assumes fg: "f b = g b"
shows "(max L M)-lipschitz_on {a .. c} (\<lambda>x. if x \<le> b then f x else g x)"
proof -
have "lipschitz_on (max L M) {a .. b} f" "lipschitz_on (max L M) {b .. c} g"
by (auto intro!: lipschitz_on_mono[OF f order_refl] lipschitz_on_mono[OF g order_refl])
from lipschitz_on_concat[OF this fg] show ?thesis .
qed
subsubsection \<open>Continuity\<close>
proposition lipschitz_on_uniformly_continuous:
assumes "L-lipschitz_on X f"
shows "uniformly_continuous_on X f"
unfolding uniformly_continuous_on_def
proof safe
fix e::real
assume "0 < e"
from assms have l: "(L+1)-lipschitz_on X f"
by (rule lipschitz_on_mono) auto
show "\<exists>d>0. \<forall>x\<in>X. \<forall>x'\<in>X. dist x' x < d \<longrightarrow> dist (f x') (f x) < e"
using lipschitz_onD[OF l] lipschitz_on_nonneg[OF assms] \<open>0 < e\<close>
by (force intro!: exI[where x="e/(L + 1)"] simp: field_simps)
qed
proposition lipschitz_on_continuous_on:
"continuous_on X f" if "L-lipschitz_on X f"
by (rule uniformly_continuous_imp_continuous[OF lipschitz_on_uniformly_continuous[OF that]])
lemma lipschitz_on_continuous_within:
"continuous (at x within X) f" if "L-lipschitz_on X f" "x \<in> X"
using lipschitz_on_continuous_on[OF that(1)] that(2)
by (auto simp: continuous_on_eq_continuous_within)
subsubsection \<open>Differentiable functions\<close>
proposition bounded_derivative_imp_lipschitz:
assumes "\<And>x. x \<in> X \<Longrightarrow> (f has_derivative f' x) (at x within X)"
assumes convex: "convex X"
assumes "\<And>x. x \<in> X \<Longrightarrow> onorm (f' x) \<le> C" "0 \<le> C"
shows "C-lipschitz_on X f"
proof (rule lipschitz_onI)
show "\<And>x y. x \<in> X \<Longrightarrow> y \<in> X \<Longrightarrow> dist (f x) (f y) \<le> C * dist x y"
by (auto intro!: assms differentiable_bound[unfolded dist_norm[symmetric], OF convex])
qed fact
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Structural introduction rules\<close>
named_theorems lipschitz_intros "structural introduction rules for Lipschitz controls"
lemma lipschitz_on_compose [lipschitz_intros]:
"(D * C)-lipschitz_on U (g o f)"
if f: "C-lipschitz_on U f" and g: "D-lipschitz_on (f`U) g"
proof (rule lipschitz_onI)
show "D* C \<ge> 0" using lipschitz_on_nonneg[OF f] lipschitz_on_nonneg[OF g] by auto
fix x y assume H: "x \<in> U" "y \<in> U"
have "dist (g (f x)) (g (f y)) \<le> D * dist (f x) (f y)"
apply (rule lipschitz_onD[OF g]) using H by auto
also have "... \<le> D * C * dist x y"
using mult_left_mono[OF lipschitz_onD(1)[OF f H] lipschitz_on_nonneg[OF g]] by auto
finally show "dist ((g \<circ> f) x) ((g \<circ> f) y) \<le> D * C* dist x y"
unfolding comp_def by (auto simp add: mult.commute)
qed
lemma lipschitz_on_compose2:
"(D * C)-lipschitz_on U (\<lambda>x. g (f x))"
if "C-lipschitz_on U f" "D-lipschitz_on (f`U) g"
using lipschitz_on_compose[OF that] by (simp add: o_def)
lemma lipschitz_on_cong[cong]:
"C-lipschitz_on U g \<longleftrightarrow> D-lipschitz_on V f"
if "C = D" "U = V" "\<And>x. x \<in> V \<Longrightarrow> g x = f x"
using that by (auto simp: lipschitz_on_def)
lemma lipschitz_on_transform:
"C-lipschitz_on U g"
if "C-lipschitz_on U f"
"\<And>x. x \<in> U \<Longrightarrow> g x = f x"
using that
by simp
lemma lipschitz_on_empty_iff[simp]: "C-lipschitz_on {} f \<longleftrightarrow> C \<ge> 0"
by (auto simp: lipschitz_on_def)
lemma lipschitz_on_insert_iff[simp]:
"C-lipschitz_on (insert y X) f \<longleftrightarrow>
C-lipschitz_on X f \<and> (\<forall>x \<in> X. dist (f x) (f y) \<le> C * dist x y)"
by (auto simp: lipschitz_on_def dist_commute)
lemma lipschitz_on_singleton [lipschitz_intros]: "C \<ge> 0 \<Longrightarrow> C-lipschitz_on {x} f"
and lipschitz_on_empty [lipschitz_intros]: "C \<ge> 0 \<Longrightarrow> C-lipschitz_on {} f"
by simp_all
lemma lipschitz_on_id [lipschitz_intros]: "1-lipschitz_on U (\<lambda>x. x)"
by (auto simp: lipschitz_on_def)
lemma lipschitz_on_constant [lipschitz_intros]: "0-lipschitz_on U (\<lambda>x. c)"
by (auto simp: lipschitz_on_def)
lemma lipschitz_on_add [lipschitz_intros]:
fixes f::"'a::metric_space \<Rightarrow>'b::real_normed_vector"
assumes "C-lipschitz_on U f"
"D-lipschitz_on U g"
shows "(C+D)-lipschitz_on U (\<lambda>x. f x + g x)"
proof (rule lipschitz_onI)
show "C + D \<ge> 0"
using lipschitz_on_nonneg[OF assms(1)] lipschitz_on_nonneg[OF assms(2)] by auto
fix x y assume H: "x \<in> U" "y \<in> U"
have "dist (f x + g x) (f y + g y) \<le> dist (f x) (f y) + dist (g x) (g y)"
by (simp add: dist_triangle_add)
also have "... \<le> C * dist x y + D * dist x y"
using lipschitz_onD(1)[OF assms(1) H] lipschitz_onD(1)[OF assms(2) H] by auto
finally show "dist (f x + g x) (f y + g y) \<le> (C+D) * dist x y" by (auto simp add: algebra_simps)
qed
lemma lipschitz_on_cmult [lipschitz_intros]:
fixes f::"'a::metric_space \<Rightarrow> 'b::real_normed_vector"
assumes "C-lipschitz_on U f"
shows "(abs(a) * C)-lipschitz_on U (\<lambda>x. a *\<^sub>R f x)"
proof (rule lipschitz_onI)
show "abs(a) * C \<ge> 0" using lipschitz_on_nonneg[OF assms(1)] by auto
fix x y assume H: "x \<in> U" "y \<in> U"
have "dist (a *\<^sub>R f x) (a *\<^sub>R f y) = abs(a) * dist (f x) (f y)"
by (metis dist_norm norm_scaleR real_vector.scale_right_diff_distrib)
also have "... \<le> abs(a) * C * dist x y"
using lipschitz_onD(1)[OF assms(1) H] by (simp add: Groups.mult_ac(1) mult_left_mono)
finally show "dist (a *\<^sub>R f x) (a *\<^sub>R f y) \<le> \<bar>a\<bar> * C * dist x y" by auto
qed
lemma lipschitz_on_cmult_real [lipschitz_intros]:
fixes f::"'a::metric_space \<Rightarrow> real"
assumes "C-lipschitz_on U f"
shows "(abs(a) * C)-lipschitz_on U (\<lambda>x. a * f x)"
using lipschitz_on_cmult[OF assms] by auto
lemma lipschitz_on_cmult_nonneg [lipschitz_intros]:
fixes f::"'a::metric_space \<Rightarrow> 'b::real_normed_vector"
assumes "C-lipschitz_on U f"
"a \<ge> 0"
shows "(a * C)-lipschitz_on U (\<lambda>x. a *\<^sub>R f x)"
using lipschitz_on_cmult[OF assms(1), of a] assms(2) by auto
lemma lipschitz_on_cmult_real_nonneg [lipschitz_intros]:
fixes f::"'a::metric_space \<Rightarrow> real"
assumes "C-lipschitz_on U f"
"a \<ge> 0"
shows "(a * C)-lipschitz_on U (\<lambda>x. a * f x)"
using lipschitz_on_cmult_nonneg[OF assms] by auto
lemma lipschitz_on_cmult_upper [lipschitz_intros]:
fixes f::"'a::metric_space \<Rightarrow> 'b::real_normed_vector"
assumes "C-lipschitz_on U f"
"abs(a) \<le> D"
shows "(D * C)-lipschitz_on U (\<lambda>x. a *\<^sub>R f x)"
apply (rule lipschitz_on_mono[OF lipschitz_on_cmult[OF assms(1), of a], of _ "D * C"])
using assms(2) lipschitz_on_nonneg[OF assms(1)] mult_right_mono by auto
lemma lipschitz_on_cmult_real_upper [lipschitz_intros]:
fixes f::"'a::metric_space \<Rightarrow> real"
assumes "C-lipschitz_on U f"
"abs(a) \<le> D"
shows "(D * C)-lipschitz_on U (\<lambda>x. a * f x)"
using lipschitz_on_cmult_upper[OF assms] by auto
lemma lipschitz_on_minus[lipschitz_intros]:
fixes f::"'a::metric_space \<Rightarrow>'b::real_normed_vector"
assumes "C-lipschitz_on U f"
shows "C-lipschitz_on U (\<lambda>x. - f x)"
by (metis (mono_tags, lifting) assms dist_minus lipschitz_on_def)
lemma lipschitz_on_minus_iff[simp]:
"L-lipschitz_on X (\<lambda>x. - f x) \<longleftrightarrow> L-lipschitz_on X f"
"L-lipschitz_on X (- f) \<longleftrightarrow> L-lipschitz_on X f"
for f::"'a::metric_space \<Rightarrow>'b::real_normed_vector"
using lipschitz_on_minus[of L X f] lipschitz_on_minus[of L X "-f"]
by auto
lemma lipschitz_on_diff[lipschitz_intros]:
fixes f::"'a::metric_space \<Rightarrow>'b::real_normed_vector"
assumes "C-lipschitz_on U f" "D-lipschitz_on U g"
shows "(C + D)-lipschitz_on U (\<lambda>x. f x - g x)"
using lipschitz_on_add[OF assms(1) lipschitz_on_minus[OF assms(2)]] by auto
lemma lipschitz_on_closure [lipschitz_intros]:
assumes "C-lipschitz_on U f" "continuous_on (closure U) f"
shows "C-lipschitz_on (closure U) f"
proof (rule lipschitz_onI)
show "C \<ge> 0" using lipschitz_on_nonneg[OF assms(1)] by simp
fix x y assume "x \<in> closure U" "y \<in> closure U"
obtain u v::"nat \<Rightarrow> 'a" where *: "\<And>n. u n \<in> U" "u \<longlonglongrightarrow> x"
"\<And>n. v n \<in> U" "v \<longlonglongrightarrow> y"
using \<open>x \<in> closure U\<close> \<open>y \<in> closure U\<close> unfolding closure_sequential by blast
have a: "(\<lambda>n. f (u n)) \<longlonglongrightarrow> f x"
using *(1) *(2) \<open>x \<in> closure U\<close> \<open>continuous_on (closure U) f\<close>
unfolding comp_def continuous_on_closure_sequentially[of U f] by auto
have b: "(\<lambda>n. f (v n)) \<longlonglongrightarrow> f y"
using *(3) *(4) \<open>y \<in> closure U\<close> \<open>continuous_on (closure U) f\<close>
unfolding comp_def continuous_on_closure_sequentially[of U f] by auto
have l: "(\<lambda>n. C * dist (u n) (v n) - dist (f (u n)) (f (v n))) \<longlonglongrightarrow> C * dist x y - dist (f x) (f y)"
by (intro tendsto_intros * a b)
have "C * dist (u n) (v n) - dist (f (u n)) (f (v n)) \<ge> 0" for n
using lipschitz_onD(1)[OF assms(1) \<open>u n \<in> U\<close> \<open>v n \<in> U\<close>] by simp
then have "C * dist x y - dist (f x) (f y) \<ge> 0" using LIMSEQ_le_const[OF l, of 0] by auto
then show "dist (f x) (f y) \<le> C * dist x y" by auto
qed
lemma lipschitz_on_Pair[lipschitz_intros]:
assumes f: "L-lipschitz_on A f"
assumes g: "M-lipschitz_on A g"
shows "(sqrt (L\<^sup>2 + M\<^sup>2))-lipschitz_on A (\<lambda>a. (f a, g a))"
proof (rule lipschitz_onI, goal_cases)
case (1 x y)
have "dist (f x, g x) (f y, g y) = sqrt ((dist (f x) (f y))\<^sup>2 + (dist (g x) (g y))\<^sup>2)"
by (auto simp add: dist_Pair_Pair real_le_lsqrt)
also have "\<dots> \<le> sqrt ((L * dist x y)\<^sup>2 + (M * dist x y)\<^sup>2)"
by (auto intro!: real_sqrt_le_mono add_mono power_mono 1 lipschitz_onD f g)
also have "\<dots> \<le> sqrt (L\<^sup>2 + M\<^sup>2) * dist x y"
by (auto simp: power_mult_distrib ring_distribs[symmetric] real_sqrt_mult)
finally show ?case .
qed simp
lemma lipschitz_extend_closure:
fixes f::"('a::metric_space) \<Rightarrow> ('b::complete_space)"
assumes "C-lipschitz_on U f"
shows "\<exists>g. C-lipschitz_on (closure U) g \<and> (\<forall>x\<in>U. g x = f x)"
proof -
obtain g where g: "\<And>x. x \<in> U \<Longrightarrow> g x = f x" "uniformly_continuous_on (closure U) g"
using uniformly_continuous_on_extension_on_closure[OF lipschitz_on_uniformly_continuous[OF assms]] by metis
have "C-lipschitz_on (closure U) g"
apply (rule lipschitz_on_closure, rule lipschitz_on_transform[OF assms])
using g uniformly_continuous_imp_continuous[OF g(2)] by auto
then show ?thesis using g(1) by auto
qed
lemma (in bounded_linear) lipschitz_boundE:
obtains B where "B-lipschitz_on A f"
proof -
from nonneg_bounded
obtain B where B: "B \<ge> 0" "\<And>x. norm (f x) \<le> B * norm x"
by (auto simp: ac_simps)
have "B-lipschitz_on A f"
by (auto intro!: lipschitz_onI B simp: dist_norm diff[symmetric])
thus ?thesis ..
qed
subsection \<open>Local Lipschitz continuity\<close>
text \<open>Given a function defined on a real interval, it is Lipschitz-continuous if and only if
it is locally so, as proved in the following lemmas. It is useful especially for
piecewise-defined functions: if each piece is Lipschitz, then so is the whole function.
The same goes for functions defined on geodesic spaces, or more generally on geodesic subsets
in a metric space (for instance convex subsets in a real vector space), and this follows readily
from the real case, but we will not prove it explicitly.
We give several variations around this statement. This is essentially a connectedness argument.\<close>
lemma locally_lipschitz_imp_lipschitz_aux:
fixes f::"real \<Rightarrow> ('a::metric_space)"
assumes "a \<le> b"
"continuous_on {a..b} f"
"\<And>x. x \<in> {a..<b} \<Longrightarrow> \<exists>y \<in> {x<..b}. dist (f y) (f x) \<le> M * (y-x)"
shows "dist (f b) (f a) \<le> M * (b-a)"
proof -
define A where "A = {x \<in> {a..b}. dist (f x) (f a) \<le> M * (x-a)}"
have *: "A = (\<lambda>x. M * (x-a) - dist (f x) (f a))-`{0..} \<inter> {a..b}"
unfolding A_def by auto
have "a \<in> A" unfolding A_def using \<open>a \<le> b\<close> by auto
then have "A \<noteq> {}" by auto
moreover have "bdd_above A" unfolding A_def by auto
moreover have "closed A" unfolding * by (rule closed_vimage_Int, auto intro!: continuous_intros assms)
ultimately have "Sup A \<in> A" by (rule closed_contains_Sup)
have "Sup A = b"
proof (rule ccontr)
assume "Sup A \<noteq> b"
define x where "x = Sup A"
have I: "dist (f x) (f a) \<le> M * (x-a)" using \<open>Sup A \<in> A\<close> x_def A_def by auto
have "x \<in> {a..<b}" unfolding x_def using \<open>Sup A \<in> A\<close> \<open>Sup A \<noteq> b\<close> A_def by auto
then obtain y where J: "y \<in> {x<..b}" "dist (f y) (f x) \<le> M * (y-x)" using assms(3) by blast
have "dist (f y) (f a) \<le> dist (f y) (f x) + dist (f x) (f a)" by (rule dist_triangle)
also have "... \<le> M * (y-x) + M * (x-a)" using I J(2) by auto
finally have "dist (f y) (f a) \<le> M * (y-a)" by (auto simp add: algebra_simps)
then have "y \<in> A" unfolding A_def using \<open>y \<in> {x<..b}\<close> \<open>x \<in> {a..<b}\<close> by auto
then have "y \<le> Sup A" by (rule cSup_upper, auto simp: A_def)
then show False using \<open>y \<in> {x<..b}\<close> x_def by auto
qed
then show ?thesis using \<open>Sup A \<in> A\<close> A_def by auto
qed
lemma locally_lipschitz_imp_lipschitz:
fixes f::"real \<Rightarrow> ('a::metric_space)"
assumes "continuous_on {a..b} f"
"\<And>x y. x \<in> {a..<b} \<Longrightarrow> y > x \<Longrightarrow> \<exists>z \<in> {x<..y}. dist (f z) (f x) \<le> M * (z-x)"
"M \<ge> 0"
shows "lipschitz_on M {a..b} f"
proof (rule lipschitz_onI[OF _ \<open>M \<ge> 0\<close>])
have *: "dist (f t) (f s) \<le> M * (t-s)" if "s \<le> t" "s \<in> {a..b}" "t \<in> {a..b}" for s t
proof (rule locally_lipschitz_imp_lipschitz_aux, simp add: \<open>s \<le> t\<close>)
show "continuous_on {s..t} f" using continuous_on_subset[OF assms(1)] that by auto
fix x assume "x \<in> {s..<t}"
then have "x \<in> {a..<b}" using that by auto
show "\<exists>z\<in>{x<..t}. dist (f z) (f x) \<le> M * (z - x)"
using assms(2)[OF \<open>x \<in> {a..<b}\<close>, of t] \<open>x \<in> {s..<t}\<close> by auto
qed
fix x y assume "x \<in> {a..b}" "y \<in> {a..b}"
consider "x \<le> y" | "y \<le> x" by linarith
then show "dist (f x) (f y) \<le> M * dist x y"
apply (cases)
using *[OF _ \<open>x \<in> {a..b}\<close> \<open>y \<in> {a..b}\<close>] *[OF _ \<open>y \<in> {a..b}\<close> \<open>x \<in> {a..b}\<close>]
by (auto simp add: dist_commute dist_real_def)
qed
text \<open>We deduce that if a function is Lipschitz on finitely many closed sets on the real line, then
it is Lipschitz on any interval contained in their union. The difficulty in the proof is to show
that any point \<open>z\<close> in this interval (except the maximum) has a point arbitrarily close to it on its
right which is contained in a common initial closed set. Otherwise, we show that there is a small
interval \<open>(z, T)\<close> which does not intersect any of the initial closed sets, a contradiction.\<close>
proposition lipschitz_on_closed_Union:
assumes "\<And>i. i \<in> I \<Longrightarrow> lipschitz_on M (U i) f"
"\<And>i. i \<in> I \<Longrightarrow> closed (U i)"
"finite I"
"M \<ge> 0"
"{u..(v::real)} \<subseteq> (\<Union>i\<in>I. U i)"
shows "lipschitz_on M {u..v} f"
proof (rule locally_lipschitz_imp_lipschitz[OF _ _ \<open>M \<ge> 0\<close>])
have *: "continuous_on (U i) f" if "i \<in> I" for i
by (rule lipschitz_on_continuous_on[OF assms(1)[OF \<open>i\<in> I\<close>]])
have "continuous_on (\<Union>i\<in>I. U i) f"
apply (rule continuous_on_closed_Union) using \<open>finite I\<close> * assms(2) by auto
then show "continuous_on {u..v} f"
using \<open>{u..(v::real)} \<subseteq> (\<Union>i\<in>I. U i)\<close> continuous_on_subset by auto
fix z Z assume z: "z \<in> {u..<v}" "z < Z"
then have "u \<le> v" by auto
define T where "T = min Z v"
then have T: "T > z" "T \<le> v" "T \<ge> u" "T \<le> Z" using z by auto
define A where "A = (\<Union>i\<in> I \<inter> {i. U i \<inter> {z<..T} \<noteq> {}}. U i \<inter> {z..T})"
have a: "closed A"
unfolding A_def apply (rule closed_UN) using \<open>finite I\<close> \<open>\<And>i. i \<in> I \<Longrightarrow> closed (U i)\<close> by auto
have b: "bdd_below A" unfolding A_def using \<open>finite I\<close> by auto
have "\<exists>i \<in> I. T \<in> U i" using \<open>{u..v} \<subseteq> (\<Union>i\<in>I. U i)\<close> T by auto
then have c: "T \<in> A" unfolding A_def using T by (auto, fastforce)
have "Inf A \<ge> z"
apply (rule cInf_greatest, auto) using c unfolding A_def by auto
moreover have "Inf A \<le> z"
proof (rule ccontr)
assume "\<not>(Inf A \<le> z)"
then obtain w where w: "w > z" "w < Inf A" by (meson dense not_le_imp_less)
have "Inf A \<le> T" using a b c by (simp add: cInf_lower)
then have "w \<le> T" using w by auto
then have "w \<in> {u..v}" using w \<open>z \<in> {u..<v}\<close> T by auto
then obtain j where j: "j \<in> I" "w \<in> U j" using \<open>{u..v} \<subseteq> (\<Union>i\<in>I. U i)\<close> by fastforce
then have "w \<in> U j \<inter> {z..T}" "U j \<inter> {z<..T} \<noteq> {}" using j T w \<open>w \<le> T\<close> by auto
then have "w \<in> A" unfolding A_def using \<open>j \<in> I\<close> by auto
then have "Inf A \<le> w" using a b by (simp add: cInf_lower)
then show False using w by auto
qed
ultimately have "Inf A = z" by simp
moreover have "Inf A \<in> A"
apply (rule closed_contains_Inf) using a b c by auto
ultimately have "z \<in> A" by simp
then obtain i where i: "i \<in> I" "U i \<inter> {z<..T} \<noteq> {}" "z \<in> U i" unfolding A_def by auto
then obtain t where "t \<in> U i \<inter> {z<..T}" by blast
then have "dist (f t) (f z) \<le> M * (t - z)"
using lipschitz_onD(1)[OF assms(1)[of i], of t z] i dist_real_def by auto
then show "\<exists>t\<in>{z<..Z}. dist (f t) (f z) \<le> M * (t - z)" using \<open>T \<le> Z\<close> \<open>t \<in> U i \<inter> {z<..T}\<close> by auto
qed
subsection \<open>Local Lipschitz continuity (uniform for a family of functions)\<close>
definition\<^marker>\<open>tag important\<close> local_lipschitz::
"'a::metric_space set \<Rightarrow> 'b::metric_space set \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c::metric_space) \<Rightarrow> bool"
where
"local_lipschitz T X f \<equiv> \<forall>x \<in> X. \<forall>t \<in> T.
\<exists>u>0. \<exists>L. \<forall>t \<in> cball t u \<inter> T. L-lipschitz_on (cball x u \<inter> X) (f t)"
lemma local_lipschitzI:
assumes "\<And>t x. t \<in> T \<Longrightarrow> x \<in> X \<Longrightarrow> \<exists>u>0. \<exists>L. \<forall>t \<in> cball t u \<inter> T. L-lipschitz_on (cball x u \<inter> X) (f t)"
shows "local_lipschitz T X f"
using assms
unfolding local_lipschitz_def
by auto
lemma local_lipschitzE:
assumes local_lipschitz: "local_lipschitz T X f"
assumes "t \<in> T" "x \<in> X"
obtains u L where "u > 0" "\<And>s. s \<in> cball t u \<inter> T \<Longrightarrow> L-lipschitz_on (cball x u \<inter> X) (f s)"
using assms local_lipschitz_def
by metis
lemma local_lipschitz_continuous_on:
assumes local_lipschitz: "local_lipschitz T X f"
assumes "t \<in> T"
shows "continuous_on X (f t)"
unfolding continuous_on_def
proof safe
fix x assume "x \<in> X"
from local_lipschitzE[OF local_lipschitz \<open>t \<in> T\<close> \<open>x \<in> X\<close>] obtain u L
where "0 < u"
and L: "\<And>s. s \<in> cball t u \<inter> T \<Longrightarrow> L-lipschitz_on (cball x u \<inter> X) (f s)"
by metis
have "x \<in> ball x u" using \<open>0 < u\<close> by simp
from lipschitz_on_continuous_on[OF L]
have tendsto: "(f t \<longlongrightarrow> f t x) (at x within cball x u \<inter> X)"
using \<open>0 < u\<close> \<open>x \<in> X\<close> \<open>t \<in> T\<close>
by (auto simp: continuous_on_def)
moreover have "\<forall>\<^sub>F xa in at x. (xa \<in> cball x u \<inter> X) = (xa \<in> X)"
using eventually_at_ball[OF \<open>0 < u\<close>, of x UNIV]
by eventually_elim auto
ultimately show "(f t \<longlongrightarrow> f t x) (at x within X)"
by (rule Lim_transform_within_set)
qed
lemma
local_lipschitz_compose1:
assumes ll: "local_lipschitz (g ` T) X (\<lambda>t. f t)"
assumes g: "continuous_on T g"
shows "local_lipschitz T X (\<lambda>t. f (g t))"
proof (rule local_lipschitzI)
fix t x
assume "t \<in> T" "x \<in> X"
then have "g t \<in> g ` T" by simp
from local_lipschitzE[OF assms(1) this \<open>x \<in> X\<close>]
obtain u L where "0 < u" and l: "(\<And>s. s \<in> cball (g t) u \<inter> g ` T \<Longrightarrow> L-lipschitz_on (cball x u \<inter> X) (f s))"
by auto
from g[unfolded continuous_on_eq_continuous_within, rule_format, OF \<open>t \<in> T\<close>,
unfolded continuous_within_eps_delta, rule_format, OF \<open>0 < u\<close>]
obtain d where d: "d>0" "\<And>x'. x'\<in>T \<Longrightarrow> dist x' t < d \<Longrightarrow> dist (g x') (g t) < u"
by (auto)
show "\<exists>u>0. \<exists>L. \<forall>t\<in>cball t u \<inter> T. L-lipschitz_on (cball x u \<inter> X) (f (g t))"
using d \<open>0 < u\<close>
by (fastforce intro: exI[where x="(min d u)/2"] exI[where x=L]
intro!: less_imp_le[OF d(2)] lipschitz_on_subset[OF l] simp: dist_commute)
qed
context
fixes T::"'a::metric_space set" and X f
assumes local_lipschitz: "local_lipschitz T X f"
begin
lemma continuous_on_TimesI:
assumes y: "\<And>x. x \<in> X \<Longrightarrow> continuous_on T (\<lambda>t. f t x)"
shows "continuous_on (T \<times> X) (\<lambda>(t, x). f t x)"
unfolding continuous_on_iff
proof (safe, simp)
fix a b and e::real
assume H: "a \<in> T" "b \<in> X" "0 < e"
hence "0 < e/2" by simp
from y[unfolded continuous_on_iff, OF \<open>b \<in> X\<close>, rule_format, OF \<open>a \<in> T\<close> \<open>0 < e/2\<close>]
obtain d where d: "d > 0" "\<And>t. t \<in> T \<Longrightarrow> dist t a < d \<Longrightarrow> dist (f t b) (f a b) < e/2"
by auto
from \<open>a : T\<close> \<open>b \<in> X\<close>
obtain u L where u: "0 < u"
and L: "\<And>t. t \<in> cball a u \<inter> T \<Longrightarrow> L-lipschitz_on (cball b u \<inter> X) (f t)"
by (erule local_lipschitzE[OF local_lipschitz])
have "a \<in> cball a u \<inter> T" by (auto simp: \<open>0 < u\<close> \<open>a \<in> T\<close> less_imp_le)
from lipschitz_on_nonneg[OF L[OF \<open>a \<in> cball _ _ \<inter> _\<close>]] have "0 \<le> L" .
let ?d = "Min {d, u, (e/2/(L + 1))}"
show "\<exists>d>0. \<forall>x\<in>T. \<forall>y\<in>X. dist (x, y) (a, b) < d \<longrightarrow> dist (f x y) (f a b) < e"
proof (rule exI[where x = ?d], safe)
show "0 < ?d"
using \<open>0 \<le> L\<close> \<open>0 < u\<close> \<open>0 < e\<close> \<open>0 < d\<close>
by (auto intro!: divide_pos_pos )
fix x y
assume "x \<in> T" "y \<in> X"
assume dist_less: "dist (x, y) (a, b) < ?d"
have "dist y b \<le> dist (x, y) (a, b)"
using dist_snd_le[of "(x, y)" "(a, b)"]
by auto
also
note dist_less
also
{
note calculation
also have "?d \<le> u" by simp
finally have "dist y b < u" .
}
have "?d \<le> e/2/(L + 1)" by simp
also have "(L + 1) * \<dots> \<le> e / 2"
using \<open>0 < e\<close> \<open>L \<ge> 0\<close>
by (auto simp: field_split_simps)
finally have le1: "(L + 1) * dist y b < e / 2" using \<open>L \<ge> 0\<close> by simp
have "dist x a \<le> dist (x, y) (a, b)"
using dist_fst_le[of "(x, y)" "(a, b)"]
by auto
also note dist_less
finally have "dist x a < ?d" .
also have "?d \<le> d" by simp
finally have "dist x a < d" .
note \<open>dist x a < ?d\<close>
also have "?d \<le> u" by simp
finally have "dist x a < u" .
then have "x \<in> cball a u \<inter> T"
using \<open>x \<in> T\<close>
by (auto simp: dist_commute)
have "dist (f x y) (f a b) \<le> dist (f x y) (f x b) + dist (f x b) (f a b)"
by (rule dist_triangle)
also have "(L + 1)-lipschitz_on (cball b u \<inter> X) (f x)"
using L[OF \<open>x \<in> cball a u \<inter> T\<close>]
by (rule lipschitz_on_le) simp
then have "dist (f x y) (f x b) \<le> (L + 1) * dist y b"
apply (rule lipschitz_onD)
subgoal
using \<open>y \<in> X\<close> \<open>dist y b < u\<close>
by (simp add: dist_commute)
subgoal
using \<open>0 < u\<close> \<open>b \<in> X\<close>
by (simp add: )
done
also have "(L + 1) * dist y b \<le> e / 2"
using le1 \<open>0 \<le> L\<close> by simp
also have "dist (f x b) (f a b) < e / 2"
by (rule d; fact)
also have "e / 2 + e / 2 = e" by simp
finally show "dist (f x y) (f a b) < e" by simp
qed
qed
lemma local_lipschitz_compact_implies_lipschitz:
assumes "compact X" "compact T"
assumes cont: "\<And>x. x \<in> X \<Longrightarrow> continuous_on T (\<lambda>t. f t x)"
obtains L where "\<And>t. t \<in> T \<Longrightarrow> L-lipschitz_on X (f t)"
proof -
{
assume *: "\<And>n::nat. \<not>(\<forall>t\<in>T. n-lipschitz_on X (f t))"
{
fix n::nat
from *[of n] have "\<exists>x y t. t \<in> T \<and> x \<in> X \<and> y \<in> X \<and> dist (f t y) (f t x) > n * dist y x"
by (force simp: lipschitz_on_def)
} then obtain t and x y::"nat \<Rightarrow> 'b" where xy: "\<And>n. x n \<in> X" "\<And>n. y n \<in> X"
and t: "\<And>n. t n \<in> T"
and d: "\<And>n. dist (f (t n) (y n)) (f (t n) (x n)) > n * dist (y n) (x n)"
by metis
from xy assms obtain lx rx where lx': "lx \<in> X" "strict_mono (rx :: nat \<Rightarrow> nat)" "(x o rx) \<longlonglongrightarrow> lx"
by (metis compact_def)
with xy have "\<And>n. (y o rx) n \<in> X" by auto
with assms obtain ly ry where ly': "ly \<in> X" "strict_mono (ry :: nat \<Rightarrow> nat)" "((y o rx) o ry) \<longlonglongrightarrow> ly"
by (metis compact_def)
with t have "\<And>n. ((t o rx) o ry) n \<in> T" by simp
with assms obtain lt rt where lt': "lt \<in> T" "strict_mono (rt :: nat \<Rightarrow> nat)" "(((t o rx) o ry) o rt) \<longlonglongrightarrow> lt"
by (metis compact_def)
from lx' ly'
have lx: "(x o (rx o ry o rt)) \<longlonglongrightarrow> lx" (is "?x \<longlonglongrightarrow> _")
and ly: "(y o (rx o ry o rt)) \<longlonglongrightarrow> ly" (is "?y \<longlonglongrightarrow> _")
and lt: "(t o (rx o ry o rt)) \<longlonglongrightarrow> lt" (is "?t \<longlonglongrightarrow> _")
subgoal by (simp add: LIMSEQ_subseq_LIMSEQ o_assoc lt'(2))
subgoal by (simp add: LIMSEQ_subseq_LIMSEQ ly'(3) o_assoc lt'(2))
subgoal by (simp add: o_assoc lt'(3))
done
hence "(\<lambda>n. dist (?y n) (?x n)) \<longlonglongrightarrow> dist ly lx"
by (metis tendsto_dist)
moreover
let ?S = "(\<lambda>(t, x). f t x) ` (T \<times> X)"
have "eventually (\<lambda>n::nat. n > 0) sequentially"
by (metis eventually_at_top_dense)
hence "eventually (\<lambda>n. norm (dist (?y n) (?x n)) \<le> norm (\<bar>diameter ?S\<bar> / n) * 1) sequentially"
proof eventually_elim
case (elim n)
have "0 < rx (ry (rt n))" using \<open>0 < n\<close>
by (metis dual_order.strict_trans1 lt'(2) lx'(2) ly'(2) seq_suble)
have compact: "compact ?S"
by (auto intro!: compact_continuous_image continuous_on_subset[OF continuous_on_TimesI]
compact_Times \<open>compact X\<close> \<open>compact T\<close> cont)
have "norm (dist (?y n) (?x n)) = dist (?y n) (?x n)" by simp
also
from this elim d[of "rx (ry (rt n))"]
have "\<dots> < dist (f (?t n) (?y n)) (f (?t n) (?x n)) / rx (ry (rt (n)))"
using lx'(2) ly'(2) lt'(2) \<open>0 < rx _\<close>
- by (auto simp add: field_split_simps algebra_simps strict_mono_def)
+ by (auto simp add: field_split_simps strict_mono_def)
also have "\<dots> \<le> diameter ?S / n"
by (force intro!: \<open>0 < n\<close> strict_mono_def xy diameter_bounded_bound frac_le
compact_imp_bounded compact t
intro: le_trans[OF seq_suble[OF lt'(2)]]
le_trans[OF seq_suble[OF ly'(2)]]
le_trans[OF seq_suble[OF lx'(2)]])
also have "\<dots> \<le> abs (diameter ?S) / n"
by (auto intro!: divide_right_mono)
finally show ?case by simp
qed
with _ have "(\<lambda>n. dist (?y n) (?x n)) \<longlonglongrightarrow> 0"
by (rule tendsto_0_le)
(metis tendsto_divide_0[OF tendsto_const] filterlim_at_top_imp_at_infinity
filterlim_real_sequentially)
ultimately have "lx = ly"
using LIMSEQ_unique by fastforce
with assms lx' have "lx \<in> X" by auto
from \<open>lt \<in> T\<close> this obtain u L where L: "u > 0" "\<And>t. t \<in> cball lt u \<inter> T \<Longrightarrow> L-lipschitz_on (cball lx u \<inter> X) (f t)"
by (erule local_lipschitzE[OF local_lipschitz])
hence "L \<ge> 0" by (force intro!: lipschitz_on_nonneg \<open>lt \<in> T\<close>)
from L lt ly lx \<open>lx = ly\<close>
have
"eventually (\<lambda>n. ?t n \<in> ball lt u) sequentially"
"eventually (\<lambda>n. ?y n \<in> ball lx u) sequentially"
"eventually (\<lambda>n. ?x n \<in> ball lx u) sequentially"
by (auto simp: dist_commute Lim)
moreover have "eventually (\<lambda>n. n > L) sequentially"
by (metis filterlim_at_top_dense filterlim_real_sequentially)
ultimately
have "eventually (\<lambda>_. False) sequentially"
proof eventually_elim
case (elim n)
hence "dist (f (?t n) (?y n)) (f (?t n) (?x n)) \<le> L * dist (?y n) (?x n)"
using assms xy t
unfolding dist_norm[symmetric]
by (intro lipschitz_onD[OF L(2)]) (auto)
also have "\<dots> \<le> n * dist (?y n) (?x n)"
using elim by (intro mult_right_mono) auto
also have "\<dots> \<le> rx (ry (rt n)) * dist (?y n) (?x n)"
by (intro mult_right_mono[OF _ zero_le_dist])
(meson lt'(2) lx'(2) ly'(2) of_nat_le_iff order_trans seq_suble)
also have "\<dots> < dist (f (?t n) (?y n)) (f (?t n) (?x n))"
by (auto intro!: d)
finally show ?case by simp
qed
hence False
by simp
} then obtain L where "\<And>t. t \<in> T \<Longrightarrow> L-lipschitz_on X (f t)"
by metis
thus ?thesis ..
qed
lemma local_lipschitz_subset:
assumes "S \<subseteq> T" "Y \<subseteq> X"
shows "local_lipschitz S Y f"
proof (rule local_lipschitzI)
fix t x assume "t \<in> S" "x \<in> Y"
then have "t \<in> T" "x \<in> X" using assms by auto
from local_lipschitzE[OF local_lipschitz, OF this]
obtain u L where u: "0 < u" and L: "\<And>s. s \<in> cball t u \<inter> T \<Longrightarrow> L-lipschitz_on (cball x u \<inter> X) (f s)"
by blast
show "\<exists>u>0. \<exists>L. \<forall>t\<in>cball t u \<inter> S. L-lipschitz_on (cball x u \<inter> Y) (f t)"
using assms
by (auto intro: exI[where x=u] exI[where x=L]
intro!: u lipschitz_on_subset[OF _ Int_mono[OF order_refl \<open>Y \<subseteq> X\<close>]] L)
qed
end
lemma local_lipschitz_minus:
fixes f::"'a::metric_space \<Rightarrow> 'b::metric_space \<Rightarrow> 'c::real_normed_vector"
shows "local_lipschitz T X (\<lambda>t x. - f t x) = local_lipschitz T X f"
by (auto simp: local_lipschitz_def lipschitz_on_minus)
lemma local_lipschitz_PairI:
assumes f: "local_lipschitz A B (\<lambda>a b. f a b)"
assumes g: "local_lipschitz A B (\<lambda>a b. g a b)"
shows "local_lipschitz A B (\<lambda>a b. (f a b, g a b))"
proof (rule local_lipschitzI)
fix t x assume "t \<in> A" "x \<in> B"
from local_lipschitzE[OF f this] local_lipschitzE[OF g this]
obtain u L v M where "0 < u" "(\<And>s. s \<in> cball t u \<inter> A \<Longrightarrow> L-lipschitz_on (cball x u \<inter> B) (f s))"
"0 < v" "(\<And>s. s \<in> cball t v \<inter> A \<Longrightarrow> M-lipschitz_on (cball x v \<inter> B) (g s))"
by metis
then show "\<exists>u>0. \<exists>L. \<forall>t\<in>cball t u \<inter> A. L-lipschitz_on (cball x u \<inter> B) (\<lambda>b. (f t b, g t b))"
by (intro exI[where x="min u v"])
(force intro: lipschitz_on_subset intro!: lipschitz_on_Pair)
qed
lemma local_lipschitz_constI: "local_lipschitz S T (\<lambda>t x. f t)"
by (auto simp: intro!: local_lipschitzI lipschitz_on_constant intro: exI[where x=1])
lemma (in bounded_linear) local_lipschitzI:
shows "local_lipschitz A B (\<lambda>_. f)"
proof (rule local_lipschitzI, goal_cases)
case (1 t x)
from lipschitz_boundE[of "(cball x 1 \<inter> B)"] obtain C where "C-lipschitz_on (cball x 1 \<inter> B) f" by auto
then show ?case
by (auto intro: exI[where x=1])
qed
proposition c1_implies_local_lipschitz:
fixes T::"real set" and X::"'a::{banach,heine_borel} set"
and f::"real \<Rightarrow> 'a \<Rightarrow> 'a"
assumes f': "\<And>t x. t \<in> T \<Longrightarrow> x \<in> X \<Longrightarrow> (f t has_derivative blinfun_apply (f' (t, x))) (at x)"
assumes cont_f': "continuous_on (T \<times> X) f'"
assumes "open T"
assumes "open X"
shows "local_lipschitz T X f"
proof (rule local_lipschitzI)
fix t x
assume "t \<in> T" "x \<in> X"
from open_contains_cball[THEN iffD1, OF \<open>open X\<close>, rule_format, OF \<open>x \<in> X\<close>]
obtain u where u: "u > 0" "cball x u \<subseteq> X" by auto
moreover
from open_contains_cball[THEN iffD1, OF \<open>open T\<close>, rule_format, OF \<open>t \<in> T\<close>]
obtain v where v: "v > 0" "cball t v \<subseteq> T" by auto
ultimately
have "compact (cball t v \<times> cball x u)" "cball t v \<times> cball x u \<subseteq> T \<times> X"
by (auto intro!: compact_Times)
then have "compact (f' ` (cball t v \<times> cball x u))"
by (auto intro!: compact_continuous_image continuous_on_subset[OF cont_f'])
then obtain B where B: "B > 0" "\<And>s y. s \<in> cball t v \<Longrightarrow> y \<in> cball x u \<Longrightarrow> norm (f' (s, y)) \<le> B"
by (auto dest!: compact_imp_bounded simp: bounded_pos)
have lipschitz: "B-lipschitz_on (cball x (min u v) \<inter> X) (f s)" if s: "s \<in> cball t v" for s
proof -
note s
also note \<open>cball t v \<subseteq> T\<close>
finally
have deriv: "\<And>y. y \<in> cball x u \<Longrightarrow> (f s has_derivative blinfun_apply (f' (s, y))) (at y within cball x u)"
using \<open>_ \<subseteq> X\<close>
by (auto intro!: has_derivative_at_withinI[OF f'])
have "norm (f s y - f s z) \<le> B * norm (y - z)"
if "y \<in> cball x u" "z \<in> cball x u"
for y z
using s that
by (intro differentiable_bound[OF convex_cball deriv])
(auto intro!: B simp: norm_blinfun.rep_eq[symmetric])
then show ?thesis
using \<open>0 < B\<close>
by (auto intro!: lipschitz_onI simp: dist_norm)
qed
show "\<exists>u>0. \<exists>L. \<forall>t\<in>cball t u \<inter> T. L-lipschitz_on (cball x u \<inter> X) (f t)"
by (force intro: exI[where x="min u v"] exI[where x=B] intro!: lipschitz simp: u v)
qed
end
diff --git a/src/HOL/Analysis/Measure_Space.thy b/src/HOL/Analysis/Measure_Space.thy
--- a/src/HOL/Analysis/Measure_Space.thy
+++ b/src/HOL/Analysis/Measure_Space.thy
@@ -1,3883 +1,3883 @@
(* Title: HOL/Analysis/Measure_Space.thy
Author: Lawrence C Paulson
Author: Johannes Hölzl, TU München
Author: Armin Heller, TU München
*)
section \<open>Measure Spaces\<close>
theory Measure_Space
imports
Measurable "HOL-Library.Extended_Nonnegative_Real"
begin
subsection\<^marker>\<open>tag unimportant\<close> "Relate extended reals and the indicator function"
lemma suminf_cmult_indicator:
fixes f :: "nat \<Rightarrow> ennreal"
assumes "disjoint_family A" "x \<in> A i"
shows "(\<Sum>n. f n * indicator (A n) x) = f i"
proof -
have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: ennreal)"
using \<open>x \<in> A i\<close> assms unfolding disjoint_family_on_def indicator_def by auto
then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: ennreal)"
by (auto simp: sum.If_cases)
moreover have "(SUP n. if i < n then f i else 0) = (f i :: ennreal)"
proof (rule SUP_eqI)
fix y :: ennreal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y"
from this[of "Suc i"] show "f i \<le> y" by auto
qed (insert assms, simp)
ultimately show ?thesis using assms
by (subst suminf_eq_SUP) (auto simp: indicator_def)
qed
lemma suminf_indicator:
assumes "disjoint_family A"
shows "(\<Sum>n. indicator (A n) x :: ennreal) = indicator (\<Union>i. A i) x"
proof cases
assume *: "x \<in> (\<Union>i. A i)"
then obtain i where "x \<in> A i" by auto
from suminf_cmult_indicator[OF assms(1), OF \<open>x \<in> A i\<close>, of "\<lambda>k. 1"]
show ?thesis using * by simp
qed simp
lemma sum_indicator_disjoint_family:
fixes f :: "'d \<Rightarrow> 'e::semiring_1"
assumes d: "disjoint_family_on A P" and "x \<in> A j" and "finite P" and "j \<in> P"
shows "(\<Sum>i\<in>P. f i * indicator (A i) x) = f j"
proof -
have "P \<inter> {i. x \<in> A i} = {j}"
using d \<open>x \<in> A j\<close> \<open>j \<in> P\<close> unfolding disjoint_family_on_def
by auto
thus ?thesis
unfolding indicator_def
by (simp add: if_distrib sum.If_cases[OF \<open>finite P\<close>])
qed
text \<open>
The type for emeasure spaces is already defined in \<^theory>\<open>HOL-Analysis.Sigma_Algebra\<close>, as it
is also used to represent sigma algebras (with an arbitrary emeasure).
\<close>
subsection\<^marker>\<open>tag unimportant\<close> "Extend binary sets"
lemma LIMSEQ_binaryset:
assumes f: "f {} = 0"
shows "(\<lambda>n. \<Sum>i<n. f (binaryset A B i)) \<longlonglongrightarrow> f A + f B"
proof -
have "(\<lambda>n. \<Sum>i < Suc (Suc n). f (binaryset A B i)) = (\<lambda>n. f A + f B)"
proof
fix n
show "(\<Sum>i < Suc (Suc n). f (binaryset A B i)) = f A + f B"
by (induct n) (auto simp add: binaryset_def f)
qed
moreover
have "... \<longlonglongrightarrow> f A + f B" by (rule tendsto_const)
ultimately
have "(\<lambda>n. \<Sum>i< Suc (Suc n). f (binaryset A B i)) \<longlonglongrightarrow> f A + f B"
by metis
hence "(\<lambda>n. \<Sum>i< n+2. f (binaryset A B i)) \<longlonglongrightarrow> f A + f B"
by simp
thus ?thesis by (rule LIMSEQ_offset [where k=2])
qed
lemma binaryset_sums:
assumes f: "f {} = 0"
shows "(\<lambda>n. f (binaryset A B n)) sums (f A + f B)"
by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f] atLeast0LessThan)
lemma suminf_binaryset_eq:
fixes f :: "'a set \<Rightarrow> 'b::{comm_monoid_add, t2_space}"
shows "f {} = 0 \<Longrightarrow> (\<Sum>n. f (binaryset A B n)) = f A + f B"
by (metis binaryset_sums sums_unique)
subsection\<^marker>\<open>tag unimportant\<close> \<open>Properties of a premeasure \<^term>\<open>\<mu>\<close>\<close>
text \<open>
The definitions for \<^const>\<open>positive\<close> and \<^const>\<open>countably_additive\<close> should be here, by they are
necessary to define \<^typ>\<open>'a measure\<close> in \<^theory>\<open>HOL-Analysis.Sigma_Algebra\<close>.
\<close>
definition subadditive where
"subadditive M f \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<inter> y = {} \<longrightarrow> f (x \<union> y) \<le> f x + f y)"
lemma subadditiveD: "subadditive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x \<in> M \<Longrightarrow> y \<in> M \<Longrightarrow> f (x \<union> y) \<le> f x + f y"
by (auto simp add: subadditive_def)
definition countably_subadditive where
"countably_subadditive M f \<longleftrightarrow>
(\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (f (\<Union>i. A i) \<le> (\<Sum>i. f (A i))))"
lemma (in ring_of_sets) countably_subadditive_subadditive:
fixes f :: "'a set \<Rightarrow> ennreal"
assumes f: "positive M f" and cs: "countably_subadditive M f"
shows "subadditive M f"
proof (auto simp add: subadditive_def)
fix x y
assume x: "x \<in> M" and y: "y \<in> M" and "x \<inter> y = {}"
hence "disjoint_family (binaryset x y)"
by (auto simp add: disjoint_family_on_def binaryset_def)
hence "range (binaryset x y) \<subseteq> M \<longrightarrow>
(\<Union>i. binaryset x y i) \<in> M \<longrightarrow>
f (\<Union>i. binaryset x y i) \<le> (\<Sum> n. f (binaryset x y n))"
using cs by (auto simp add: countably_subadditive_def)
hence "{x,y,{}} \<subseteq> M \<longrightarrow> x \<union> y \<in> M \<longrightarrow>
f (x \<union> y) \<le> (\<Sum> n. f (binaryset x y n))"
by (simp add: range_binaryset_eq UN_binaryset_eq)
thus "f (x \<union> y) \<le> f x + f y" using f x y
by (auto simp add: Un o_def suminf_binaryset_eq positive_def)
qed
definition additive where
"additive M \<mu> \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<inter> y = {} \<longrightarrow> \<mu> (x \<union> y) = \<mu> x + \<mu> y)"
definition increasing where
"increasing M \<mu> \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<subseteq> y \<longrightarrow> \<mu> x \<le> \<mu> y)"
lemma positiveD1: "positive M f \<Longrightarrow> f {} = 0" by (auto simp: positive_def)
lemma positiveD_empty:
"positive M f \<Longrightarrow> f {} = 0"
by (auto simp add: positive_def)
lemma additiveD:
"additive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x \<in> M \<Longrightarrow> y \<in> M \<Longrightarrow> f (x \<union> y) = f x + f y"
by (auto simp add: additive_def)
lemma increasingD:
"increasing M f \<Longrightarrow> x \<subseteq> y \<Longrightarrow> x\<in>M \<Longrightarrow> y\<in>M \<Longrightarrow> f x \<le> f y"
by (auto simp add: increasing_def)
lemma countably_additiveI[case_names countably]:
"(\<And>A. range A \<subseteq> M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i. A i) \<in> M \<Longrightarrow> (\<Sum>i. f (A i)) = f (\<Union>i. A i))
\<Longrightarrow> countably_additive M f"
by (simp add: countably_additive_def)
lemma (in ring_of_sets) disjointed_additive:
assumes f: "positive M f" "additive M f" and A: "range A \<subseteq> M" "incseq A"
shows "(\<Sum>i\<le>n. f (disjointed A i)) = f (A n)"
proof (induct n)
case (Suc n)
then have "(\<Sum>i\<le>Suc n. f (disjointed A i)) = f (A n) + f (disjointed A (Suc n))"
by simp
also have "\<dots> = f (A n \<union> disjointed A (Suc n))"
using A by (subst f(2)[THEN additiveD]) (auto simp: disjointed_mono)
also have "A n \<union> disjointed A (Suc n) = A (Suc n)"
using \<open>incseq A\<close> by (auto dest: incseq_SucD simp: disjointed_mono)
finally show ?case .
qed simp
lemma (in ring_of_sets) additive_sum:
fixes A:: "'i \<Rightarrow> 'a set"
assumes f: "positive M f" and ad: "additive M f" and "finite S"
and A: "A`S \<subseteq> M"
and disj: "disjoint_family_on A S"
shows "(\<Sum>i\<in>S. f (A i)) = f (\<Union>i\<in>S. A i)"
using \<open>finite S\<close> disj A
proof induct
case empty show ?case using f by (simp add: positive_def)
next
case (insert s S)
then have "A s \<inter> (\<Union>i\<in>S. A i) = {}"
by (auto simp add: disjoint_family_on_def neq_iff)
moreover
have "A s \<in> M" using insert by blast
moreover have "(\<Union>i\<in>S. A i) \<in> M"
using insert \<open>finite S\<close> by auto
ultimately have "f (A s \<union> (\<Union>i\<in>S. A i)) = f (A s) + f(\<Union>i\<in>S. A i)"
using ad UNION_in_sets A by (auto simp add: additive_def)
with insert show ?case using ad disjoint_family_on_mono[of S "insert s S" A]
by (auto simp add: additive_def subset_insertI)
qed
lemma (in ring_of_sets) additive_increasing:
fixes f :: "'a set \<Rightarrow> ennreal"
assumes posf: "positive M f" and addf: "additive M f"
shows "increasing M f"
proof (auto simp add: increasing_def)
fix x y
assume xy: "x \<in> M" "y \<in> M" "x \<subseteq> y"
then have "y - x \<in> M" by auto
then have "f x + 0 \<le> f x + f (y-x)" by (intro add_left_mono zero_le)
also have "... = f (x \<union> (y-x))" using addf
by (auto simp add: additive_def) (metis Diff_disjoint Un_Diff_cancel Diff xy(1,2))
also have "... = f y"
by (metis Un_Diff_cancel Un_absorb1 xy(3))
finally show "f x \<le> f y" by simp
qed
lemma (in ring_of_sets) subadditive:
fixes f :: "'a set \<Rightarrow> ennreal"
assumes f: "positive M f" "additive M f" and A: "A`S \<subseteq> M" and S: "finite S"
shows "f (\<Union>i\<in>S. A i) \<le> (\<Sum>i\<in>S. f (A i))"
using S A
proof (induct S)
case empty thus ?case using f by (auto simp: positive_def)
next
case (insert x F)
hence in_M: "A x \<in> M" "(\<Union>i\<in>F. A i) \<in> M" "(\<Union>i\<in>F. A i) - A x \<in> M" using A by force+
have subs: "(\<Union>i\<in>F. A i) - A x \<subseteq> (\<Union>i\<in>F. A i)" by auto
have "(\<Union>i\<in>(insert x F). A i) = A x \<union> ((\<Union>i\<in>F. A i) - A x)" by auto
hence "f (\<Union>i\<in>(insert x F). A i) = f (A x \<union> ((\<Union>i\<in>F. A i) - A x))"
by simp
also have "\<dots> = f (A x) + f ((\<Union>i\<in>F. A i) - A x)"
using f(2) by (rule additiveD) (insert in_M, auto)
also have "\<dots> \<le> f (A x) + f (\<Union>i\<in>F. A i)"
using additive_increasing[OF f] in_M subs by (auto simp: increasing_def intro: add_left_mono)
also have "\<dots> \<le> f (A x) + (\<Sum>i\<in>F. f (A i))" using insert by (auto intro: add_left_mono)
finally show "f (\<Union>i\<in>(insert x F). A i) \<le> (\<Sum>i\<in>(insert x F). f (A i))" using insert by simp
qed
lemma (in ring_of_sets) countably_additive_additive:
fixes f :: "'a set \<Rightarrow> ennreal"
assumes posf: "positive M f" and ca: "countably_additive M f"
shows "additive M f"
proof (auto simp add: additive_def)
fix x y
assume x: "x \<in> M" and y: "y \<in> M" and "x \<inter> y = {}"
hence "disjoint_family (binaryset x y)"
by (auto simp add: disjoint_family_on_def binaryset_def)
hence "range (binaryset x y) \<subseteq> M \<longrightarrow>
(\<Union>i. binaryset x y i) \<in> M \<longrightarrow>
f (\<Union>i. binaryset x y i) = (\<Sum> n. f (binaryset x y n))"
using ca
by (simp add: countably_additive_def)
hence "{x,y,{}} \<subseteq> M \<longrightarrow> x \<union> y \<in> M \<longrightarrow>
f (x \<union> y) = (\<Sum>n. f (binaryset x y n))"
by (simp add: range_binaryset_eq UN_binaryset_eq)
thus "f (x \<union> y) = f x + f y" using posf x y
by (auto simp add: Un suminf_binaryset_eq positive_def)
qed
lemma (in algebra) increasing_additive_bound:
fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> ennreal"
assumes f: "positive M f" and ad: "additive M f"
and inc: "increasing M f"
and A: "range A \<subseteq> M"
and disj: "disjoint_family A"
shows "(\<Sum>i. f (A i)) \<le> f \<Omega>"
proof (safe intro!: suminf_le_const)
fix N
note disj_N = disjoint_family_on_mono[OF _ disj, of "{..<N}"]
have "(\<Sum>i<N. f (A i)) = f (\<Union>i\<in>{..<N}. A i)"
using A by (intro additive_sum [OF f ad _ _]) (auto simp: disj_N)
also have "... \<le> f \<Omega>" using space_closed A
by (intro increasingD[OF inc] finite_UN) auto
finally show "(\<Sum>i<N. f (A i)) \<le> f \<Omega>" by simp
qed (insert f A, auto simp: positive_def)
lemma (in ring_of_sets) countably_additiveI_finite:
fixes \<mu> :: "'a set \<Rightarrow> ennreal"
assumes "finite \<Omega>" "positive M \<mu>" "additive M \<mu>"
shows "countably_additive M \<mu>"
proof (rule countably_additiveI)
fix F :: "nat \<Rightarrow> 'a set" assume F: "range F \<subseteq> M" "(\<Union>i. F i) \<in> M" and disj: "disjoint_family F"
have "\<forall>i\<in>{i. F i \<noteq> {}}. \<exists>x. x \<in> F i" by auto
from bchoice[OF this] obtain f where f: "\<And>i. F i \<noteq> {} \<Longrightarrow> f i \<in> F i" by auto
have inj_f: "inj_on f {i. F i \<noteq> {}}"
proof (rule inj_onI, simp)
fix i j a b assume *: "f i = f j" "F i \<noteq> {}" "F j \<noteq> {}"
then have "f i \<in> F i" "f j \<in> F j" using f by force+
with disj * show "i = j" by (auto simp: disjoint_family_on_def)
qed
have "finite (\<Union>i. F i)"
by (metis F(2) assms(1) infinite_super sets_into_space)
have F_subset: "{i. \<mu> (F i) \<noteq> 0} \<subseteq> {i. F i \<noteq> {}}"
by (auto simp: positiveD_empty[OF \<open>positive M \<mu>\<close>])
moreover have fin_not_empty: "finite {i. F i \<noteq> {}}"
proof (rule finite_imageD)
from f have "f`{i. F i \<noteq> {}} \<subseteq> (\<Union>i. F i)" by auto
then show "finite (f`{i. F i \<noteq> {}})"
by (rule finite_subset) fact
qed fact
ultimately have fin_not_0: "finite {i. \<mu> (F i) \<noteq> 0}"
by (rule finite_subset)
have disj_not_empty: "disjoint_family_on F {i. F i \<noteq> {}}"
using disj by (auto simp: disjoint_family_on_def)
from fin_not_0 have "(\<Sum>i. \<mu> (F i)) = (\<Sum>i | \<mu> (F i) \<noteq> 0. \<mu> (F i))"
by (rule suminf_finite) auto
also have "\<dots> = (\<Sum>i | F i \<noteq> {}. \<mu> (F i))"
using fin_not_empty F_subset by (rule sum.mono_neutral_left) auto
also have "\<dots> = \<mu> (\<Union>i\<in>{i. F i \<noteq> {}}. F i)"
using \<open>positive M \<mu>\<close> \<open>additive M \<mu>\<close> fin_not_empty disj_not_empty F by (intro additive_sum) auto
also have "\<dots> = \<mu> (\<Union>i. F i)"
by (rule arg_cong[where f=\<mu>]) auto
finally show "(\<Sum>i. \<mu> (F i)) = \<mu> (\<Union>i. F i)" .
qed
lemma (in ring_of_sets) countably_additive_iff_continuous_from_below:
fixes f :: "'a set \<Rightarrow> ennreal"
assumes f: "positive M f" "additive M f"
shows "countably_additive M f \<longleftrightarrow>
(\<forall>A. range A \<subseteq> M \<longrightarrow> incseq A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i))"
unfolding countably_additive_def
proof safe
assume count_sum: "\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> \<Union>(A ` UNIV) \<in> M \<longrightarrow> (\<Sum>i. f (A i)) = f (\<Union>(A ` UNIV))"
fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "incseq A" "(\<Union>i. A i) \<in> M"
then have dA: "range (disjointed A) \<subseteq> M" by (auto simp: range_disjointed_sets)
with count_sum[THEN spec, of "disjointed A"] A(3)
have f_UN: "(\<Sum>i. f (disjointed A i)) = f (\<Union>i. A i)"
by (auto simp: UN_disjointed_eq disjoint_family_disjointed)
moreover have "(\<lambda>n. (\<Sum>i<n. f (disjointed A i))) \<longlonglongrightarrow> (\<Sum>i. f (disjointed A i))"
using f(1)[unfolded positive_def] dA
by (auto intro!: summable_LIMSEQ)
from LIMSEQ_Suc[OF this]
have "(\<lambda>n. (\<Sum>i\<le>n. f (disjointed A i))) \<longlonglongrightarrow> (\<Sum>i. f (disjointed A i))"
unfolding lessThan_Suc_atMost .
moreover have "\<And>n. (\<Sum>i\<le>n. f (disjointed A i)) = f (A n)"
using disjointed_additive[OF f A(1,2)] .
ultimately show "(\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i)" by simp
next
assume cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> incseq A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i)"
fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "disjoint_family A" "(\<Union>i. A i) \<in> M"
have *: "(\<Union>n. (\<Union>i<n. A i)) = (\<Union>i. A i)" by auto
have "(\<lambda>n. f (\<Union>i<n. A i)) \<longlonglongrightarrow> f (\<Union>i. A i)"
proof (unfold *[symmetric], intro cont[rule_format])
show "range (\<lambda>i. \<Union>i<i. A i) \<subseteq> M" "(\<Union>i. \<Union>i<i. A i) \<in> M"
using A * by auto
qed (force intro!: incseq_SucI)
moreover have "\<And>n. f (\<Union>i<n. A i) = (\<Sum>i<n. f (A i))"
using A
by (intro additive_sum[OF f, of _ A, symmetric])
(auto intro: disjoint_family_on_mono[where B=UNIV])
ultimately
have "(\<lambda>i. f (A i)) sums f (\<Union>i. A i)"
unfolding sums_def by simp
from sums_unique[OF this]
show "(\<Sum>i. f (A i)) = f (\<Union>i. A i)" by simp
qed
lemma (in ring_of_sets) continuous_from_above_iff_empty_continuous:
fixes f :: "'a set \<Rightarrow> ennreal"
assumes f: "positive M f" "additive M f"
shows "(\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) \<in> M \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Inter>i. A i))
\<longleftrightarrow> (\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> 0)"
proof safe
assume cont: "(\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) \<in> M \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Inter>i. A i))"
fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "decseq A" "(\<Inter>i. A i) = {}" "\<forall>i. f (A i) \<noteq> \<infinity>"
with cont[THEN spec, of A] show "(\<lambda>i. f (A i)) \<longlonglongrightarrow> 0"
using \<open>positive M f\<close>[unfolded positive_def] by auto
next
assume cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> 0"
fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "decseq A" "(\<Inter>i. A i) \<in> M" "\<forall>i. f (A i) \<noteq> \<infinity>"
have f_mono: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<subseteq> b \<Longrightarrow> f a \<le> f b"
using additive_increasing[OF f] unfolding increasing_def by simp
have decseq_fA: "decseq (\<lambda>i. f (A i))"
using A by (auto simp: decseq_def intro!: f_mono)
have decseq: "decseq (\<lambda>i. A i - (\<Inter>i. A i))"
using A by (auto simp: decseq_def)
then have decseq_f: "decseq (\<lambda>i. f (A i - (\<Inter>i. A i)))"
using A unfolding decseq_def by (auto intro!: f_mono Diff)
have "f (\<Inter>x. A x) \<le> f (A 0)"
using A by (auto intro!: f_mono)
then have f_Int_fin: "f (\<Inter>x. A x) \<noteq> \<infinity>"
using A by (auto simp: top_unique)
{ fix i
have "f (A i - (\<Inter>i. A i)) \<le> f (A i)" using A by (auto intro!: f_mono)
then have "f (A i - (\<Inter>i. A i)) \<noteq> \<infinity>"
using A by (auto simp: top_unique) }
note f_fin = this
have "(\<lambda>i. f (A i - (\<Inter>i. A i))) \<longlonglongrightarrow> 0"
proof (intro cont[rule_format, OF _ decseq _ f_fin])
show "range (\<lambda>i. A i - (\<Inter>i. A i)) \<subseteq> M" "(\<Inter>i. A i - (\<Inter>i. A i)) = {}"
using A by auto
qed
from INF_Lim[OF decseq_f this]
have "(INF n. f (A n - (\<Inter>i. A i))) = 0" .
moreover have "(INF n. f (\<Inter>i. A i)) = f (\<Inter>i. A i)"
by auto
ultimately have "(INF n. f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i)) = 0 + f (\<Inter>i. A i)"
using A(4) f_fin f_Int_fin
by (subst INF_ennreal_add_const) (auto simp: decseq_f)
moreover {
fix n
have "f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i) = f ((A n - (\<Inter>i. A i)) \<union> (\<Inter>i. A i))"
using A by (subst f(2)[THEN additiveD]) auto
also have "(A n - (\<Inter>i. A i)) \<union> (\<Inter>i. A i) = A n"
by auto
finally have "f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i) = f (A n)" . }
ultimately have "(INF n. f (A n)) = f (\<Inter>i. A i)"
by simp
with LIMSEQ_INF[OF decseq_fA]
show "(\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Inter>i. A i)" by simp
qed
lemma (in ring_of_sets) empty_continuous_imp_continuous_from_below:
fixes f :: "'a set \<Rightarrow> ennreal"
assumes f: "positive M f" "additive M f" "\<forall>A\<in>M. f A \<noteq> \<infinity>"
assumes cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> 0"
assumes A: "range A \<subseteq> M" "incseq A" "(\<Union>i. A i) \<in> M"
shows "(\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i)"
proof -
from A have "(\<lambda>i. f ((\<Union>i. A i) - A i)) \<longlonglongrightarrow> 0"
by (intro cont[rule_format]) (auto simp: decseq_def incseq_def)
moreover
{ fix i
have "f ((\<Union>i. A i) - A i \<union> A i) = f ((\<Union>i. A i) - A i) + f (A i)"
using A by (intro f(2)[THEN additiveD]) auto
also have "((\<Union>i. A i) - A i) \<union> A i = (\<Union>i. A i)"
by auto
finally have "f ((\<Union>i. A i) - A i) = f (\<Union>i. A i) - f (A i)"
using f(3)[rule_format, of "A i"] A by (auto simp: ennreal_add_diff_cancel subset_eq) }
moreover have "\<forall>\<^sub>F i in sequentially. f (A i) \<le> f (\<Union>i. A i)"
using increasingD[OF additive_increasing[OF f(1, 2)], of "A _" "\<Union>i. A i"] A
by (auto intro!: always_eventually simp: subset_eq)
ultimately show "(\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i)"
by (auto intro: ennreal_tendsto_const_minus)
qed
lemma (in ring_of_sets) empty_continuous_imp_countably_additive:
fixes f :: "'a set \<Rightarrow> ennreal"
assumes f: "positive M f" "additive M f" and fin: "\<forall>A\<in>M. 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 "countably_additive M f"
using countably_additive_iff_continuous_from_below[OF f]
using empty_continuous_imp_continuous_from_below[OF f fin] cont
by blast
subsection\<^marker>\<open>tag unimportant\<close> \<open>Properties of \<^const>\<open>emeasure\<close>\<close>
lemma emeasure_positive: "positive (sets M) (emeasure M)"
by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def)
lemma emeasure_empty[simp, intro]: "emeasure M {} = 0"
using emeasure_positive[of M] by (simp add: positive_def)
lemma emeasure_single_in_space: "emeasure M {x} \<noteq> 0 \<Longrightarrow> x \<in> space M"
using emeasure_notin_sets[of "{x}" M] by (auto dest: sets.sets_into_space zero_less_iff_neq_zero[THEN iffD2])
lemma emeasure_countably_additive: "countably_additive (sets M) (emeasure M)"
by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def)
lemma suminf_emeasure:
"range A \<subseteq> sets M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Sum>i. emeasure M (A i)) = emeasure M (\<Union>i. A i)"
using sets.countable_UN[of A UNIV M] emeasure_countably_additive[of M]
by (simp add: countably_additive_def)
lemma sums_emeasure:
"disjoint_family F \<Longrightarrow> (\<And>i. F i \<in> sets M) \<Longrightarrow> (\<lambda>i. emeasure M (F i)) sums emeasure M (\<Union>i. F i)"
unfolding sums_iff by (intro conjI suminf_emeasure) auto
lemma emeasure_additive: "additive (sets M) (emeasure M)"
by (metis sets.countably_additive_additive emeasure_positive emeasure_countably_additive)
lemma plus_emeasure:
"a \<in> sets M \<Longrightarrow> b \<in> sets M \<Longrightarrow> a \<inter> b = {} \<Longrightarrow> emeasure M a + emeasure M b = emeasure M (a \<union> b)"
using additiveD[OF emeasure_additive] ..
lemma emeasure_Un:
"A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> emeasure M (A \<union> B) = emeasure M A + emeasure M (B - A)"
using plus_emeasure[of A M "B - A"] by auto
lemma emeasure_Un_Int:
assumes "A \<in> sets M" "B \<in> sets M"
shows "emeasure M A + emeasure M B = emeasure M (A \<union> B) + emeasure M (A \<inter> B)"
proof -
have "A = (A-B) \<union> (A \<inter> B)" by auto
then have "emeasure M A = emeasure M (A-B) + emeasure M (A \<inter> B)"
by (metis Diff_Diff_Int Diff_disjoint assms plus_emeasure sets.Diff)
moreover have "A \<union> B = (A-B) \<union> B" by auto
then have "emeasure M (A \<union> B) = emeasure M (A-B) + emeasure M B"
by (metis Diff_disjoint Int_commute assms plus_emeasure sets.Diff)
ultimately show ?thesis by (metis add.assoc add.commute)
qed
lemma sum_emeasure:
"F`I \<subseteq> sets M \<Longrightarrow> disjoint_family_on F I \<Longrightarrow> finite I \<Longrightarrow>
(\<Sum>i\<in>I. emeasure M (F i)) = emeasure M (\<Union>i\<in>I. F i)"
by (metis sets.additive_sum emeasure_positive emeasure_additive)
lemma emeasure_mono:
"a \<subseteq> b \<Longrightarrow> b \<in> sets M \<Longrightarrow> emeasure M a \<le> emeasure M b"
by (metis zero_le sets.additive_increasing emeasure_additive emeasure_notin_sets emeasure_positive increasingD)
lemma emeasure_space:
"emeasure M A \<le> emeasure M (space M)"
by (metis emeasure_mono emeasure_notin_sets sets.sets_into_space sets.top zero_le)
lemma emeasure_Diff:
assumes finite: "emeasure M B \<noteq> \<infinity>"
and [measurable]: "A \<in> sets M" "B \<in> sets M" and "B \<subseteq> A"
shows "emeasure M (A - B) = emeasure M A - emeasure M B"
proof -
have "(A - B) \<union> B = A" using \<open>B \<subseteq> A\<close> by auto
then have "emeasure M A = emeasure M ((A - B) \<union> B)" by simp
also have "\<dots> = emeasure M (A - B) + emeasure M B"
by (subst plus_emeasure[symmetric]) auto
finally show "emeasure M (A - B) = emeasure M A - emeasure M B"
using finite by simp
qed
lemma emeasure_compl:
"s \<in> sets M \<Longrightarrow> emeasure M s \<noteq> \<infinity> \<Longrightarrow> emeasure M (space M - s) = emeasure M (space M) - emeasure M s"
by (rule emeasure_Diff) (auto dest: sets.sets_into_space)
lemma Lim_emeasure_incseq:
"range A \<subseteq> sets M \<Longrightarrow> incseq A \<Longrightarrow> (\<lambda>i. (emeasure M (A i))) \<longlonglongrightarrow> emeasure M (\<Union>i. A i)"
using emeasure_countably_additive
by (auto simp add: sets.countably_additive_iff_continuous_from_below emeasure_positive
emeasure_additive)
lemma incseq_emeasure:
assumes "range B \<subseteq> sets M" "incseq B"
shows "incseq (\<lambda>i. emeasure M (B i))"
using assms by (auto simp: incseq_def intro!: emeasure_mono)
lemma SUP_emeasure_incseq:
assumes A: "range A \<subseteq> sets M" "incseq A"
shows "(SUP n. emeasure M (A n)) = emeasure M (\<Union>i. A i)"
using LIMSEQ_SUP[OF incseq_emeasure, OF A] Lim_emeasure_incseq[OF A]
by (simp add: LIMSEQ_unique)
lemma decseq_emeasure:
assumes "range B \<subseteq> sets M" "decseq B"
shows "decseq (\<lambda>i. emeasure M (B i))"
using assms by (auto simp: decseq_def intro!: emeasure_mono)
lemma INF_emeasure_decseq:
assumes A: "range A \<subseteq> sets M" and "decseq A"
and finite: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
shows "(INF n. emeasure M (A n)) = emeasure M (\<Inter>i. A i)"
proof -
have le_MI: "emeasure M (\<Inter>i. A i) \<le> emeasure M (A 0)"
using A by (auto intro!: emeasure_mono)
hence *: "emeasure M (\<Inter>i. A i) \<noteq> \<infinity>" using finite[of 0] by (auto simp: top_unique)
have "emeasure M (A 0) - (INF n. emeasure M (A n)) = (SUP n. emeasure M (A 0) - emeasure M (A n))"
by (simp add: ennreal_INF_const_minus)
also have "\<dots> = (SUP n. emeasure M (A 0 - A n))"
using A finite \<open>decseq A\<close>[unfolded decseq_def] by (subst emeasure_Diff) auto
also have "\<dots> = emeasure M (\<Union>i. A 0 - A i)"
proof (rule SUP_emeasure_incseq)
show "range (\<lambda>n. A 0 - A n) \<subseteq> sets M"
using A by auto
show "incseq (\<lambda>n. A 0 - A n)"
using \<open>decseq A\<close> by (auto simp add: incseq_def decseq_def)
qed
also have "\<dots> = emeasure M (A 0) - emeasure M (\<Inter>i. A i)"
using A finite * by (simp, subst emeasure_Diff) auto
finally show ?thesis
by (rule ennreal_minus_cancel[rotated 3])
(insert finite A, auto intro: INF_lower emeasure_mono)
qed
lemma INF_emeasure_decseq':
assumes A: "\<And>i. A i \<in> sets M" and "decseq A"
and finite: "\<exists>i. emeasure M (A i) \<noteq> \<infinity>"
shows "(INF n. emeasure M (A n)) = emeasure M (\<Inter>i. A i)"
proof -
from finite obtain i where i: "emeasure M (A i) < \<infinity>"
by (auto simp: less_top)
have fin: "i \<le> j \<Longrightarrow> emeasure M (A j) < \<infinity>" for j
by (rule le_less_trans[OF emeasure_mono i]) (auto intro!: decseqD[OF \<open>decseq A\<close>] A)
have "(INF n. emeasure M (A n)) = (INF n. emeasure M (A (n + i)))"
proof (rule INF_eq)
show "\<exists>j\<in>UNIV. emeasure M (A (j + i)) \<le> emeasure M (A i')" for i'
by (intro bexI[of _ i'] emeasure_mono decseqD[OF \<open>decseq A\<close>] A) auto
qed auto
also have "\<dots> = emeasure M (INF n. (A (n + i)))"
using A \<open>decseq A\<close> fin by (intro INF_emeasure_decseq) (auto simp: decseq_def less_top)
also have "(INF n. (A (n + i))) = (INF n. A n)"
by (meson INF_eq UNIV_I assms(2) decseqD le_add1)
finally show ?thesis .
qed
lemma emeasure_INT_decseq_subset:
fixes F :: "nat \<Rightarrow> 'a set"
assumes I: "I \<noteq> {}" and F: "\<And>i j. i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> i \<le> j \<Longrightarrow> F j \<subseteq> F i"
assumes F_sets[measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> sets M"
and fin: "\<And>i. i \<in> I \<Longrightarrow> emeasure M (F i) \<noteq> \<infinity>"
shows "emeasure M (\<Inter>i\<in>I. F i) = (INF i\<in>I. emeasure M (F i))"
proof cases
assume "finite I"
have "(\<Inter>i\<in>I. F i) = F (Max I)"
using I \<open>finite I\<close> by (intro antisym INF_lower INF_greatest F) auto
moreover have "(INF i\<in>I. emeasure M (F i)) = emeasure M (F (Max I))"
using I \<open>finite I\<close> by (intro antisym INF_lower INF_greatest F emeasure_mono) auto
ultimately show ?thesis
by simp
next
assume "infinite I"
define L where "L n = (LEAST i. i \<in> I \<and> i \<ge> n)" for n
have L: "L n \<in> I \<and> n \<le> L n" for n
unfolding L_def
proof (rule LeastI_ex)
show "\<exists>x. x \<in> I \<and> n \<le> x"
using \<open>infinite I\<close> finite_subset[of I "{..< n}"]
by (rule_tac ccontr) (auto simp: not_le)
qed
have L_eq[simp]: "i \<in> I \<Longrightarrow> L i = i" for i
unfolding L_def by (intro Least_equality) auto
have L_mono: "i \<le> j \<Longrightarrow> L i \<le> L j" for i j
using L[of j] unfolding L_def by (intro Least_le) (auto simp: L_def)
have "emeasure M (\<Inter>i. F (L i)) = (INF i. emeasure M (F (L i)))"
proof (intro INF_emeasure_decseq[symmetric])
show "decseq (\<lambda>i. F (L i))"
using L by (intro antimonoI F L_mono) auto
qed (insert L fin, auto)
also have "\<dots> = (INF i\<in>I. emeasure M (F i))"
proof (intro antisym INF_greatest)
show "i \<in> I \<Longrightarrow> (INF i. emeasure M (F (L i))) \<le> emeasure M (F i)" for i
by (intro INF_lower2[of i]) auto
qed (insert L, auto intro: INF_lower)
also have "(\<Inter>i. F (L i)) = (\<Inter>i\<in>I. F i)"
proof (intro antisym INF_greatest)
show "i \<in> I \<Longrightarrow> (\<Inter>i. F (L i)) \<subseteq> F i" for i
by (intro INF_lower2[of i]) auto
qed (insert L, auto)
finally show ?thesis .
qed
lemma Lim_emeasure_decseq:
assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
shows "(\<lambda>i. emeasure M (A i)) \<longlonglongrightarrow> emeasure M (\<Inter>i. A i)"
using LIMSEQ_INF[OF decseq_emeasure, OF A]
using INF_emeasure_decseq[OF A fin] by simp
lemma emeasure_lfp'[consumes 1, case_names cont measurable]:
assumes "P M"
assumes cont: "sup_continuous F"
assumes *: "\<And>M A. P M \<Longrightarrow> (\<And>N. P N \<Longrightarrow> Measurable.pred N A) \<Longrightarrow> Measurable.pred M (F A)"
shows "emeasure M {x\<in>space M. lfp F x} = (SUP i. emeasure M {x\<in>space M. (F ^^ i) (\<lambda>x. False) x})"
proof -
have "emeasure M {x\<in>space M. lfp F x} = emeasure M (\<Union>i. {x\<in>space M. (F ^^ i) (\<lambda>x. False) x})"
using sup_continuous_lfp[OF cont] by (auto simp add: bot_fun_def intro!: arg_cong2[where f=emeasure])
moreover { fix i from \<open>P M\<close> have "{x\<in>space M. (F ^^ i) (\<lambda>x. False) x} \<in> sets M"
by (induct i arbitrary: M) (auto simp add: pred_def[symmetric] intro: *) }
moreover have "incseq (\<lambda>i. {x\<in>space M. (F ^^ i) (\<lambda>x. False) x})"
proof (rule incseq_SucI)
fix i
have "(F ^^ i) (\<lambda>x. False) \<le> (F ^^ (Suc i)) (\<lambda>x. False)"
proof (induct i)
case 0 show ?case by (simp add: le_fun_def)
next
case Suc thus ?case using monoD[OF sup_continuous_mono[OF cont] Suc] by auto
qed
then show "{x \<in> space M. (F ^^ i) (\<lambda>x. False) x} \<subseteq> {x \<in> space M. (F ^^ Suc i) (\<lambda>x. False) x}"
by auto
qed
ultimately show ?thesis
by (subst SUP_emeasure_incseq) auto
qed
lemma emeasure_lfp:
assumes [simp]: "\<And>s. sets (M s) = sets N"
assumes cont: "sup_continuous F" "sup_continuous f"
assumes meas: "\<And>P. Measurable.pred N P \<Longrightarrow> Measurable.pred N (F P)"
assumes iter: "\<And>P s. Measurable.pred N P \<Longrightarrow> P \<le> lfp F \<Longrightarrow> emeasure (M s) {x\<in>space N. F P x} = f (\<lambda>s. emeasure (M s) {x\<in>space N. P x}) s"
shows "emeasure (M s) {x\<in>space N. lfp F x} = lfp f s"
proof (subst lfp_transfer_bounded[where \<alpha>="\<lambda>F s. emeasure (M s) {x\<in>space N. F x}" and g=f and f=F and P="Measurable.pred N", symmetric])
fix C assume "incseq C" "\<And>i. Measurable.pred N (C i)"
then show "(\<lambda>s. emeasure (M s) {x \<in> space N. (SUP i. C i) x}) = (SUP i. (\<lambda>s. emeasure (M s) {x \<in> space N. C i x}))"
unfolding SUP_apply[abs_def]
by (subst SUP_emeasure_incseq) (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure])
qed (auto simp add: iter le_fun_def SUP_apply[abs_def] intro!: meas cont)
lemma emeasure_subadditive_finite:
"finite I \<Longrightarrow> A ` I \<subseteq> sets M \<Longrightarrow> emeasure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. emeasure M (A i))"
by (rule sets.subadditive[OF emeasure_positive emeasure_additive]) auto
lemma emeasure_subadditive:
"A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> emeasure M (A \<union> B) \<le> emeasure M A + emeasure M B"
using emeasure_subadditive_finite[of "{True, False}" "\<lambda>True \<Rightarrow> A | False \<Rightarrow> B" M] by simp
lemma emeasure_subadditive_countably:
assumes "range f \<subseteq> sets M"
shows "emeasure M (\<Union>i. f i) \<le> (\<Sum>i. emeasure M (f i))"
proof -
have "emeasure M (\<Union>i. f i) = emeasure M (\<Union>i. disjointed f i)"
unfolding UN_disjointed_eq ..
also have "\<dots> = (\<Sum>i. emeasure M (disjointed f i))"
using sets.range_disjointed_sets[OF assms] suminf_emeasure[of "disjointed f"]
by (simp add: disjoint_family_disjointed comp_def)
also have "\<dots> \<le> (\<Sum>i. emeasure M (f i))"
using sets.range_disjointed_sets[OF assms] assms
by (auto intro!: suminf_le emeasure_mono disjointed_subset)
finally show ?thesis .
qed
lemma emeasure_insert:
assumes sets: "{x} \<in> sets M" "A \<in> sets M" and "x \<notin> A"
shows "emeasure M (insert x A) = emeasure M {x} + emeasure M A"
proof -
have "{x} \<inter> A = {}" using \<open>x \<notin> A\<close> by auto
from plus_emeasure[OF sets this] show ?thesis by simp
qed
lemma emeasure_insert_ne:
"A \<noteq> {} \<Longrightarrow> {x} \<in> sets M \<Longrightarrow> A \<in> sets M \<Longrightarrow> x \<notin> A \<Longrightarrow> emeasure M (insert x A) = emeasure M {x} + emeasure M A"
by (rule emeasure_insert)
lemma emeasure_eq_sum_singleton:
assumes "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
shows "emeasure M S = (\<Sum>x\<in>S. emeasure M {x})"
using sum_emeasure[of "\<lambda>x. {x}" S M] assms
by (auto simp: disjoint_family_on_def subset_eq)
lemma sum_emeasure_cover:
assumes "finite S" and "A \<in> sets M" and br_in_M: "B ` S \<subseteq> sets M"
assumes A: "A \<subseteq> (\<Union>i\<in>S. B i)"
assumes disj: "disjoint_family_on B S"
shows "emeasure M A = (\<Sum>i\<in>S. emeasure M (A \<inter> (B i)))"
proof -
have "(\<Sum>i\<in>S. emeasure M (A \<inter> (B i))) = emeasure M (\<Union>i\<in>S. A \<inter> (B i))"
proof (rule sum_emeasure)
show "disjoint_family_on (\<lambda>i. A \<inter> B i) S"
using \<open>disjoint_family_on B S\<close>
unfolding disjoint_family_on_def by auto
qed (insert assms, auto)
also have "(\<Union>i\<in>S. A \<inter> (B i)) = A"
using A by auto
finally show ?thesis by simp
qed
lemma emeasure_eq_0:
"N \<in> sets M \<Longrightarrow> emeasure M N = 0 \<Longrightarrow> K \<subseteq> N \<Longrightarrow> emeasure M K = 0"
by (metis emeasure_mono order_eq_iff zero_le)
lemma emeasure_UN_eq_0:
assumes "\<And>i::nat. emeasure M (N i) = 0" and "range N \<subseteq> sets M"
shows "emeasure M (\<Union>i. N i) = 0"
proof -
have "emeasure M (\<Union>i. N i) \<le> 0"
using emeasure_subadditive_countably[OF assms(2)] assms(1) by simp
then show ?thesis
by (auto intro: antisym zero_le)
qed
lemma measure_eqI_finite:
assumes [simp]: "sets M = Pow A" "sets N = Pow A" and "finite A"
assumes eq: "\<And>a. a \<in> A \<Longrightarrow> emeasure M {a} = emeasure N {a}"
shows "M = N"
proof (rule measure_eqI)
fix X assume "X \<in> sets M"
then have X: "X \<subseteq> A" by auto
then have "emeasure M X = (\<Sum>a\<in>X. emeasure M {a})"
using \<open>finite A\<close> by (subst emeasure_eq_sum_singleton) (auto dest: finite_subset)
also have "\<dots> = (\<Sum>a\<in>X. emeasure N {a})"
using X eq by (auto intro!: sum.cong)
also have "\<dots> = emeasure N X"
using X \<open>finite A\<close> by (subst emeasure_eq_sum_singleton) (auto dest: finite_subset)
finally show "emeasure M X = emeasure N X" .
qed simp
lemma measure_eqI_generator_eq:
fixes M N :: "'a measure" and E :: "'a set set" and A :: "nat \<Rightarrow> 'a set"
assumes "Int_stable E" "E \<subseteq> Pow \<Omega>"
and eq: "\<And>X. X \<in> E \<Longrightarrow> emeasure M X = emeasure N X"
and M: "sets M = sigma_sets \<Omega> E"
and N: "sets N = sigma_sets \<Omega> E"
and A: "range A \<subseteq> E" "(\<Union>i. A i) = \<Omega>" "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
shows "M = N"
proof -
let ?\<mu> = "emeasure M" and ?\<nu> = "emeasure N"
interpret S: sigma_algebra \<Omega> "sigma_sets \<Omega> E" by (rule sigma_algebra_sigma_sets) fact
have "space M = \<Omega>"
using sets.top[of M] sets.space_closed[of M] S.top S.space_closed \<open>sets M = sigma_sets \<Omega> E\<close>
by blast
{ fix F D assume "F \<in> E" and "?\<mu> F \<noteq> \<infinity>"
then have [intro]: "F \<in> sigma_sets \<Omega> E" by auto
have "?\<nu> F \<noteq> \<infinity>" using \<open>?\<mu> F \<noteq> \<infinity>\<close> \<open>F \<in> E\<close> eq by simp
assume "D \<in> sets M"
with \<open>Int_stable E\<close> \<open>E \<subseteq> Pow \<Omega>\<close> have "emeasure M (F \<inter> D) = emeasure N (F \<inter> D)"
unfolding M
proof (induct rule: sigma_sets_induct_disjoint)
case (basic A)
then have "F \<inter> A \<in> E" using \<open>Int_stable E\<close> \<open>F \<in> E\<close> by (auto simp: Int_stable_def)
then show ?case using eq by auto
next
case empty then show ?case by simp
next
case (compl A)
then have **: "F \<inter> (\<Omega> - A) = F - (F \<inter> A)"
and [intro]: "F \<inter> A \<in> sigma_sets \<Omega> E"
using \<open>F \<in> E\<close> S.sets_into_space by (auto simp: M)
have "?\<nu> (F \<inter> A) \<le> ?\<nu> F" by (auto intro!: emeasure_mono simp: M N)
then have "?\<nu> (F \<inter> A) \<noteq> \<infinity>" using \<open>?\<nu> F \<noteq> \<infinity>\<close> by (auto simp: top_unique)
have "?\<mu> (F \<inter> A) \<le> ?\<mu> F" by (auto intro!: emeasure_mono simp: M N)
then have "?\<mu> (F \<inter> A) \<noteq> \<infinity>" using \<open>?\<mu> F \<noteq> \<infinity>\<close> by (auto simp: top_unique)
then have "?\<mu> (F \<inter> (\<Omega> - A)) = ?\<mu> F - ?\<mu> (F \<inter> A)" unfolding **
using \<open>F \<inter> A \<in> sigma_sets \<Omega> E\<close> by (auto intro!: emeasure_Diff simp: M N)
also have "\<dots> = ?\<nu> F - ?\<nu> (F \<inter> A)" using eq \<open>F \<in> E\<close> compl by simp
also have "\<dots> = ?\<nu> (F \<inter> (\<Omega> - A))" unfolding **
using \<open>F \<inter> A \<in> sigma_sets \<Omega> E\<close> \<open>?\<nu> (F \<inter> A) \<noteq> \<infinity>\<close>
by (auto intro!: emeasure_Diff[symmetric] simp: M N)
finally show ?case
using \<open>space M = \<Omega>\<close> by auto
next
case (union A)
then have "?\<mu> (\<Union>x. F \<inter> A x) = ?\<nu> (\<Union>x. F \<inter> A x)"
by (subst (1 2) suminf_emeasure[symmetric]) (auto simp: disjoint_family_on_def subset_eq M N)
with A show ?case
by auto
qed }
note * = this
show "M = N"
proof (rule measure_eqI)
show "sets M = sets N"
using M N by simp
have [simp, intro]: "\<And>i. A i \<in> sets M"
using A(1) by (auto simp: subset_eq M)
fix F assume "F \<in> sets M"
let ?D = "disjointed (\<lambda>i. F \<inter> A i)"
from \<open>space M = \<Omega>\<close> have F_eq: "F = (\<Union>i. ?D i)"
using \<open>F \<in> sets M\<close>[THEN sets.sets_into_space] A(2)[symmetric] by (auto simp: UN_disjointed_eq)
have [simp, intro]: "\<And>i. ?D i \<in> sets M"
using sets.range_disjointed_sets[of "\<lambda>i. F \<inter> A i" M] \<open>F \<in> sets M\<close>
by (auto simp: subset_eq)
have "disjoint_family ?D"
by (auto simp: disjoint_family_disjointed)
moreover
have "(\<Sum>i. emeasure M (?D i)) = (\<Sum>i. emeasure N (?D i))"
proof (intro arg_cong[where f=suminf] ext)
fix i
have "A i \<inter> ?D i = ?D i"
by (auto simp: disjointed_def)
then show "emeasure M (?D i) = emeasure N (?D i)"
using *[of "A i" "?D i", OF _ A(3)] A(1) by auto
qed
ultimately show "emeasure M F = emeasure N F"
by (simp add: image_subset_iff \<open>sets M = sets N\<close>[symmetric] F_eq[symmetric] suminf_emeasure)
qed
qed
lemma space_empty: "space M = {} \<Longrightarrow> M = count_space {}"
by (rule measure_eqI) (simp_all add: space_empty_iff)
lemma measure_eqI_generator_eq_countable:
fixes M N :: "'a measure" and E :: "'a set set" and A :: "'a set set"
assumes E: "Int_stable E" "E \<subseteq> Pow \<Omega>" "\<And>X. X \<in> E \<Longrightarrow> emeasure M X = emeasure N X"
and sets: "sets M = sigma_sets \<Omega> E" "sets N = sigma_sets \<Omega> E"
and A: "A \<subseteq> E" "(\<Union>A) = \<Omega>" "countable A" "\<And>a. a \<in> A \<Longrightarrow> emeasure M a \<noteq> \<infinity>"
shows "M = N"
proof cases
assume "\<Omega> = {}"
have *: "sigma_sets \<Omega> E = sets (sigma \<Omega> E)"
using E(2) by simp
have "space M = \<Omega>" "space N = \<Omega>"
using sets E(2) unfolding * by (auto dest: sets_eq_imp_space_eq simp del: sets_measure_of)
then show "M = N"
unfolding \<open>\<Omega> = {}\<close> by (auto dest: space_empty)
next
assume "\<Omega> \<noteq> {}" with \<open>\<Union>A = \<Omega>\<close> have "A \<noteq> {}" by auto
from this \<open>countable A\<close> have rng: "range (from_nat_into A) = A"
by (rule range_from_nat_into)
show "M = N"
proof (rule measure_eqI_generator_eq[OF E sets])
show "range (from_nat_into A) \<subseteq> E"
unfolding rng using \<open>A \<subseteq> E\<close> .
show "(\<Union>i. from_nat_into A i) = \<Omega>"
unfolding rng using \<open>\<Union>A = \<Omega>\<close> .
show "emeasure M (from_nat_into A i) \<noteq> \<infinity>" for i
using rng by (intro A) auto
qed
qed
lemma measure_of_of_measure: "measure_of (space M) (sets M) (emeasure M) = M"
proof (intro measure_eqI emeasure_measure_of_sigma)
show "sigma_algebra (space M) (sets M)" ..
show "positive (sets M) (emeasure M)"
by (simp add: positive_def)
show "countably_additive (sets M) (emeasure M)"
by (simp add: emeasure_countably_additive)
qed simp_all
subsection \<open>\<open>\<mu>\<close>-null sets\<close>
definition\<^marker>\<open>tag important\<close> null_sets :: "'a measure \<Rightarrow> 'a set set" where
"null_sets M = {N\<in>sets M. emeasure M N = 0}"
lemma null_setsD1[dest]: "A \<in> null_sets M \<Longrightarrow> emeasure M A = 0"
by (simp add: null_sets_def)
lemma null_setsD2[dest]: "A \<in> null_sets M \<Longrightarrow> A \<in> sets M"
unfolding null_sets_def by simp
lemma null_setsI[intro]: "emeasure M A = 0 \<Longrightarrow> A \<in> sets M \<Longrightarrow> A \<in> null_sets M"
unfolding null_sets_def by simp
interpretation null_sets: ring_of_sets "space M" "null_sets M" for M
proof (rule ring_of_setsI)
show "null_sets M \<subseteq> Pow (space M)"
using sets.sets_into_space by auto
show "{} \<in> null_sets M"
by auto
fix A B assume null_sets: "A \<in> null_sets M" "B \<in> null_sets M"
then have sets: "A \<in> sets M" "B \<in> sets M"
by auto
then have *: "emeasure M (A \<union> B) \<le> emeasure M A + emeasure M B"
"emeasure M (A - B) \<le> emeasure M A"
by (auto intro!: emeasure_subadditive emeasure_mono)
then have "emeasure M B = 0" "emeasure M A = 0"
using null_sets by auto
with sets * show "A - B \<in> null_sets M" "A \<union> B \<in> null_sets M"
by (auto intro!: antisym zero_le)
qed
lemma UN_from_nat_into:
assumes I: "countable I" "I \<noteq> {}"
shows "(\<Union>i\<in>I. N i) = (\<Union>i. N (from_nat_into I i))"
proof -
have "(\<Union>i\<in>I. N i) = \<Union>(N ` range (from_nat_into I))"
using I by simp
also have "\<dots> = (\<Union>i. (N \<circ> from_nat_into I) i)"
by simp
finally show ?thesis by simp
qed
lemma null_sets_UN':
assumes "countable I"
assumes "\<And>i. i \<in> I \<Longrightarrow> N i \<in> null_sets M"
shows "(\<Union>i\<in>I. N i) \<in> null_sets M"
proof cases
assume "I = {}" then show ?thesis by simp
next
assume "I \<noteq> {}"
show ?thesis
proof (intro conjI CollectI null_setsI)
show "(\<Union>i\<in>I. N i) \<in> sets M"
using assms by (intro sets.countable_UN') auto
have "emeasure M (\<Union>i\<in>I. N i) \<le> (\<Sum>n. emeasure M (N (from_nat_into I n)))"
unfolding UN_from_nat_into[OF \<open>countable I\<close> \<open>I \<noteq> {}\<close>]
using assms \<open>I \<noteq> {}\<close> by (intro emeasure_subadditive_countably) (auto intro: from_nat_into)
also have "(\<lambda>n. emeasure M (N (from_nat_into I n))) = (\<lambda>_. 0)"
using assms \<open>I \<noteq> {}\<close> by (auto intro: from_nat_into)
finally show "emeasure M (\<Union>i\<in>I. N i) = 0"
by (intro antisym zero_le) simp
qed
qed
lemma null_sets_UN[intro]:
"(\<And>i::'i::countable. N i \<in> null_sets M) \<Longrightarrow> (\<Union>i. N i) \<in> null_sets M"
by (rule null_sets_UN') auto
lemma null_set_Int1:
assumes "B \<in> null_sets M" "A \<in> sets M" shows "A \<inter> B \<in> null_sets M"
proof (intro CollectI conjI null_setsI)
show "emeasure M (A \<inter> B) = 0" using assms
by (intro emeasure_eq_0[of B _ "A \<inter> B"]) auto
qed (insert assms, auto)
lemma null_set_Int2:
assumes "B \<in> null_sets M" "A \<in> sets M" shows "B \<inter> A \<in> null_sets M"
using assms by (subst Int_commute) (rule null_set_Int1)
lemma emeasure_Diff_null_set:
assumes "B \<in> null_sets M" "A \<in> sets M"
shows "emeasure M (A - B) = emeasure M A"
proof -
have *: "A - B = (A - (A \<inter> B))" by auto
have "A \<inter> B \<in> null_sets M" using assms by (rule null_set_Int1)
then show ?thesis
unfolding * using assms
by (subst emeasure_Diff) auto
qed
lemma null_set_Diff:
assumes "B \<in> null_sets M" "A \<in> sets M" shows "B - A \<in> null_sets M"
proof (intro CollectI conjI null_setsI)
show "emeasure M (B - A) = 0" using assms by (intro emeasure_eq_0[of B _ "B - A"]) auto
qed (insert assms, auto)
lemma emeasure_Un_null_set:
assumes "A \<in> sets M" "B \<in> null_sets M"
shows "emeasure M (A \<union> B) = emeasure M A"
proof -
have *: "A \<union> B = A \<union> (B - A)" by auto
have "B - A \<in> null_sets M" using assms(2,1) by (rule null_set_Diff)
then show ?thesis
unfolding * using assms
by (subst plus_emeasure[symmetric]) auto
qed
lemma emeasure_Un':
assumes "A \<in> sets M" "B \<in> sets M" "A \<inter> B \<in> null_sets M"
shows "emeasure M (A \<union> B) = emeasure M A + emeasure M B"
proof -
have "A \<union> B = A \<union> (B - A \<inter> B)" by blast
also have "emeasure M \<dots> = emeasure M A + emeasure M (B - A \<inter> B)"
using assms by (subst plus_emeasure) auto
also have "emeasure M (B - A \<inter> B) = emeasure M B"
using assms by (intro emeasure_Diff_null_set) auto
finally show ?thesis .
qed
subsection \<open>The almost everywhere filter (i.e.\ quantifier)\<close>
definition\<^marker>\<open>tag important\<close> ae_filter :: "'a measure \<Rightarrow> 'a filter" where
"ae_filter M = (INF N\<in>null_sets M. principal (space M - N))"
abbreviation almost_everywhere :: "'a measure \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
"almost_everywhere M P \<equiv> eventually P (ae_filter M)"
syntax
"_almost_everywhere" :: "pttrn \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> bool" ("AE _ in _. _" [0,0,10] 10)
translations
"AE x in M. P" \<rightleftharpoons> "CONST almost_everywhere M (\<lambda>x. P)"
abbreviation
"set_almost_everywhere A M P \<equiv> AE x in M. x \<in> A \<longrightarrow> P x"
syntax
"_set_almost_everywhere" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> bool"
("AE _\<in>_ in _./ _" [0,0,0,10] 10)
translations
"AE x\<in>A in M. P" \<rightleftharpoons> "CONST set_almost_everywhere A M (\<lambda>x. P)"
lemma eventually_ae_filter: "eventually P (ae_filter M) \<longleftrightarrow> (\<exists>N\<in>null_sets M. {x \<in> space M. \<not> P x} \<subseteq> N)"
unfolding ae_filter_def by (subst eventually_INF_base) (auto simp: eventually_principal subset_eq)
lemma AE_I':
"N \<in> null_sets M \<Longrightarrow> {x\<in>space M. \<not> P x} \<subseteq> N \<Longrightarrow> (AE x in M. P x)"
unfolding eventually_ae_filter by auto
lemma AE_iff_null:
assumes "{x\<in>space M. \<not> P x} \<in> sets M" (is "?P \<in> sets M")
shows "(AE x in M. P x) \<longleftrightarrow> {x\<in>space M. \<not> P x} \<in> null_sets M"
proof
assume "AE x in M. P x" then obtain N where N: "N \<in> sets M" "?P \<subseteq> N" "emeasure M N = 0"
unfolding eventually_ae_filter by auto
have "emeasure M ?P \<le> emeasure M N"
using assms N(1,2) by (auto intro: emeasure_mono)
then have "emeasure M ?P = 0"
unfolding \<open>emeasure M N = 0\<close> by auto
then show "?P \<in> null_sets M" using assms by auto
next
assume "?P \<in> null_sets M" with assms show "AE x in M. P x" by (auto intro: AE_I')
qed
lemma AE_iff_null_sets:
"N \<in> sets M \<Longrightarrow> N \<in> null_sets M \<longleftrightarrow> (AE x in M. x \<notin> N)"
using Int_absorb1[OF sets.sets_into_space, of N M]
by (subst AE_iff_null) (auto simp: Int_def[symmetric])
lemma AE_not_in:
"N \<in> null_sets M \<Longrightarrow> AE x in M. x \<notin> N"
by (metis AE_iff_null_sets null_setsD2)
lemma AE_iff_measurable:
"N \<in> sets M \<Longrightarrow> {x\<in>space M. \<not> P x} = N \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> emeasure M N = 0"
using AE_iff_null[of _ P] by auto
lemma AE_E[consumes 1]:
assumes "AE x in M. P x"
obtains N where "{x \<in> space M. \<not> P x} \<subseteq> N" "emeasure M N = 0" "N \<in> sets M"
using assms unfolding eventually_ae_filter by auto
lemma AE_E2:
assumes "AE x in M. P x" "{x\<in>space M. P x} \<in> sets M"
shows "emeasure M {x\<in>space M. \<not> P x} = 0" (is "emeasure M ?P = 0")
proof -
have "{x\<in>space M. \<not> P x} = space M - {x\<in>space M. P x}" by auto
with AE_iff_null[of M P] assms show ?thesis by auto
qed
lemma AE_E3:
assumes "AE x in M. P x"
obtains N where "\<And>x. x \<in> space M - N \<Longrightarrow> P x" "N \<in> null_sets M"
using assms unfolding eventually_ae_filter by auto
lemma AE_I:
assumes "{x \<in> space M. \<not> P x} \<subseteq> N" "emeasure M N = 0" "N \<in> sets M"
shows "AE x in M. P x"
using assms unfolding eventually_ae_filter by auto
lemma AE_mp[elim!]:
assumes AE_P: "AE x in M. P x" and AE_imp: "AE x in M. P x \<longrightarrow> Q x"
shows "AE x in M. Q x"
proof -
from AE_P obtain A where P: "{x\<in>space M. \<not> P x} \<subseteq> A"
and A: "A \<in> sets M" "emeasure M A = 0"
by (auto elim!: AE_E)
from AE_imp obtain B where imp: "{x\<in>space M. P x \<and> \<not> Q x} \<subseteq> B"
and B: "B \<in> sets M" "emeasure M B = 0"
by (auto elim!: AE_E)
show ?thesis
proof (intro AE_I)
have "emeasure M (A \<union> B) \<le> 0"
using emeasure_subadditive[of A M B] A B by auto
then show "A \<union> B \<in> sets M" "emeasure M (A \<union> B) = 0"
using A B by auto
show "{x\<in>space M. \<not> Q x} \<subseteq> A \<union> B"
using P imp by auto
qed
qed
text \<open>The next lemma is convenient to combine with a lemma whose conclusion is of the
form \<open>AE x in M. P x = Q x\<close>: for such a lemma, there is no \<open>[symmetric]\<close> variant,
but using \<open>AE_symmetric[OF...]\<close> will replace it.\<close>
(* depricated replace by laws about eventually *)
lemma
shows AE_iffI: "AE x in M. P x \<Longrightarrow> AE x in M. P x \<longleftrightarrow> Q x \<Longrightarrow> AE x in M. Q x"
and AE_disjI1: "AE x in M. P x \<Longrightarrow> AE x in M. P x \<or> Q x"
and AE_disjI2: "AE x in M. Q x \<Longrightarrow> AE x in M. P x \<or> Q x"
and AE_conjI: "AE x in M. P x \<Longrightarrow> AE x in M. Q x \<Longrightarrow> AE x in M. P x \<and> Q x"
and AE_conj_iff[simp]: "(AE x in M. P x \<and> Q x) \<longleftrightarrow> (AE x in M. P x) \<and> (AE x in M. Q x)"
by auto
lemma AE_symmetric:
assumes "AE x in M. P x = Q x"
shows "AE x in M. Q x = P x"
using assms by auto
lemma AE_impI:
"(P \<Longrightarrow> AE x in M. Q x) \<Longrightarrow> AE x in M. P \<longrightarrow> Q x"
by fastforce
lemma AE_measure:
assumes AE: "AE x in M. P x" and sets: "{x\<in>space M. P x} \<in> sets M" (is "?P \<in> sets M")
shows "emeasure M {x\<in>space M. P x} = emeasure M (space M)"
proof -
from AE_E[OF AE] guess N . note N = this
with sets have "emeasure M (space M) \<le> emeasure M (?P \<union> N)"
by (intro emeasure_mono) auto
also have "\<dots> \<le> emeasure M ?P + emeasure M N"
using sets N by (intro emeasure_subadditive) auto
also have "\<dots> = emeasure M ?P" using N by simp
finally show "emeasure M ?P = emeasure M (space M)"
using emeasure_space[of M "?P"] by auto
qed
lemma AE_space: "AE x in M. x \<in> space M"
by (rule AE_I[where N="{}"]) auto
lemma AE_I2[simp, intro]:
"(\<And>x. x \<in> space M \<Longrightarrow> P x) \<Longrightarrow> AE x in M. P x"
using AE_space by force
lemma AE_Ball_mp:
"\<forall>x\<in>space M. P x \<Longrightarrow> AE x in M. P x \<longrightarrow> Q x \<Longrightarrow> AE x in M. Q x"
by auto
lemma AE_cong[cong]:
"(\<And>x. x \<in> space M \<Longrightarrow> P x \<longleftrightarrow> Q x) \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> (AE x in M. Q x)"
by auto
lemma AE_cong_simp: "M = N \<Longrightarrow> (\<And>x. x \<in> space N =simp=> P x = Q x) \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> (AE x in N. Q x)"
by (auto simp: simp_implies_def)
lemma AE_all_countable:
"(AE x in M. \<forall>i. P i x) \<longleftrightarrow> (\<forall>i::'i::countable. AE x in M. P i x)"
proof
assume "\<forall>i. AE x in M. P i x"
from this[unfolded eventually_ae_filter Bex_def, THEN choice]
obtain N where N: "\<And>i. N i \<in> null_sets M" "\<And>i. {x\<in>space M. \<not> P i x} \<subseteq> N i" by auto
have "{x\<in>space M. \<not> (\<forall>i. P i x)} \<subseteq> (\<Union>i. {x\<in>space M. \<not> P i x})" by auto
also have "\<dots> \<subseteq> (\<Union>i. N i)" using N by auto
finally have "{x\<in>space M. \<not> (\<forall>i. P i x)} \<subseteq> (\<Union>i. N i)" .
moreover from N have "(\<Union>i. N i) \<in> null_sets M"
by (intro null_sets_UN) auto
ultimately show "AE x in M. \<forall>i. P i x"
unfolding eventually_ae_filter by auto
qed auto
lemma AE_ball_countable:
assumes [intro]: "countable X"
shows "(AE x in M. \<forall>y\<in>X. P x y) \<longleftrightarrow> (\<forall>y\<in>X. AE x in M. P x y)"
proof
assume "\<forall>y\<in>X. AE x in M. P x y"
from this[unfolded eventually_ae_filter Bex_def, THEN bchoice]
obtain N where N: "\<And>y. y \<in> X \<Longrightarrow> N y \<in> null_sets M" "\<And>y. y \<in> X \<Longrightarrow> {x\<in>space M. \<not> P x y} \<subseteq> N y"
by auto
have "{x\<in>space M. \<not> (\<forall>y\<in>X. P x y)} \<subseteq> (\<Union>y\<in>X. {x\<in>space M. \<not> P x y})"
by auto
also have "\<dots> \<subseteq> (\<Union>y\<in>X. N y)"
using N by auto
finally have "{x\<in>space M. \<not> (\<forall>y\<in>X. P x y)} \<subseteq> (\<Union>y\<in>X. N y)" .
moreover from N have "(\<Union>y\<in>X. N y) \<in> null_sets M"
by (intro null_sets_UN') auto
ultimately show "AE x in M. \<forall>y\<in>X. P x y"
unfolding eventually_ae_filter by auto
qed auto
lemma AE_ball_countable':
"(\<And>N. N \<in> I \<Longrightarrow> AE x in M. P N x) \<Longrightarrow> countable I \<Longrightarrow> AE x in M. \<forall>N \<in> I. P N x"
unfolding AE_ball_countable by simp
lemma AE_pairwise: "countable F \<Longrightarrow> pairwise (\<lambda>A B. AE x in M. R x A B) F \<longleftrightarrow> (AE x in M. pairwise (R x) F)"
unfolding pairwise_alt by (simp add: AE_ball_countable)
lemma AE_discrete_difference:
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"
shows "AE x in M. x \<notin> X"
proof -
have "(\<Union>x\<in>X. {x}) \<in> null_sets M"
using assms by (intro null_sets_UN') auto
from AE_not_in[OF this] show "AE x in M. x \<notin> X"
by auto
qed
lemma AE_finite_all:
assumes f: "finite S" shows "(AE x in M. \<forall>i\<in>S. P i x) \<longleftrightarrow> (\<forall>i\<in>S. AE x in M. P i x)"
using f by induct auto
lemma AE_finite_allI:
assumes "finite S"
shows "(\<And>s. s \<in> S \<Longrightarrow> AE x in M. Q s x) \<Longrightarrow> AE x in M. \<forall>s\<in>S. Q s x"
using AE_finite_all[OF \<open>finite S\<close>] by auto
lemma emeasure_mono_AE:
assumes imp: "AE x in M. x \<in> A \<longrightarrow> x \<in> B"
and B: "B \<in> sets M"
shows "emeasure M A \<le> emeasure M B"
proof cases
assume A: "A \<in> sets M"
from imp obtain N where N: "{x\<in>space M. \<not> (x \<in> A \<longrightarrow> x \<in> B)} \<subseteq> N" "N \<in> null_sets M"
by (auto simp: eventually_ae_filter)
have "emeasure M A = emeasure M (A - N)"
using N A by (subst emeasure_Diff_null_set) auto
also have "emeasure M (A - N) \<le> emeasure M (B - N)"
using N A B sets.sets_into_space by (auto intro!: emeasure_mono)
also have "emeasure M (B - N) = emeasure M B"
using N B by (subst emeasure_Diff_null_set) auto
finally show ?thesis .
qed (simp add: emeasure_notin_sets)
lemma emeasure_eq_AE:
assumes iff: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B"
assumes A: "A \<in> sets M" and B: "B \<in> sets M"
shows "emeasure M A = emeasure M B"
using assms by (safe intro!: antisym emeasure_mono_AE) auto
lemma emeasure_Collect_eq_AE:
"AE x in M. P x \<longleftrightarrow> Q x \<Longrightarrow> Measurable.pred M Q \<Longrightarrow> Measurable.pred M P \<Longrightarrow>
emeasure M {x\<in>space M. P x} = emeasure M {x\<in>space M. Q x}"
by (intro emeasure_eq_AE) auto
lemma emeasure_eq_0_AE: "AE x in M. \<not> P x \<Longrightarrow> emeasure M {x\<in>space M. P x} = 0"
using AE_iff_measurable[OF _ refl, of M "\<lambda>x. \<not> P x"]
by (cases "{x\<in>space M. P x} \<in> sets M") (simp_all add: emeasure_notin_sets)
lemma emeasure_0_AE:
assumes "emeasure M (space M) = 0"
shows "AE x in M. P x"
using eventually_ae_filter assms by blast
lemma emeasure_add_AE:
assumes [measurable]: "A \<in> sets M" "B \<in> sets M" "C \<in> sets M"
assumes 1: "AE x in M. x \<in> C \<longleftrightarrow> x \<in> A \<or> x \<in> B"
assumes 2: "AE x in M. \<not> (x \<in> A \<and> x \<in> B)"
shows "emeasure M C = emeasure M A + emeasure M B"
proof -
have "emeasure M C = emeasure M (A \<union> B)"
by (rule emeasure_eq_AE) (insert 1, auto)
also have "\<dots> = emeasure M A + emeasure M (B - A)"
by (subst plus_emeasure) auto
also have "emeasure M (B - A) = emeasure M B"
by (rule emeasure_eq_AE) (insert 2, auto)
finally show ?thesis .
qed
subsection \<open>\<open>\<sigma>\<close>-finite Measures\<close>
locale\<^marker>\<open>tag important\<close> sigma_finite_measure =
fixes M :: "'a measure"
assumes sigma_finite_countable:
"\<exists>A::'a set set. countable A \<and> A \<subseteq> sets M \<and> (\<Union>A) = space M \<and> (\<forall>a\<in>A. emeasure M a \<noteq> \<infinity>)"
lemma (in sigma_finite_measure) sigma_finite:
obtains A :: "nat \<Rightarrow> 'a set"
where "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
proof -
obtain A :: "'a set set" where
[simp]: "countable A" and
A: "A \<subseteq> sets M" "(\<Union>A) = space M" "\<And>a. a \<in> A \<Longrightarrow> emeasure M a \<noteq> \<infinity>"
using sigma_finite_countable by metis
show thesis
proof cases
assume "A = {}" with \<open>(\<Union>A) = space M\<close> show thesis
by (intro that[of "\<lambda>_. {}"]) auto
next
assume "A \<noteq> {}"
show thesis
proof
show "range (from_nat_into A) \<subseteq> sets M"
using \<open>A \<noteq> {}\<close> A by auto
have "(\<Union>i. from_nat_into A i) = \<Union>A"
using range_from_nat_into[OF \<open>A \<noteq> {}\<close> \<open>countable A\<close>] by auto
with A show "(\<Union>i. from_nat_into A i) = space M"
by auto
qed (intro A from_nat_into \<open>A \<noteq> {}\<close>)
qed
qed
lemma (in sigma_finite_measure) sigma_finite_disjoint:
obtains A :: "nat \<Rightarrow> 'a set"
where "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>" "disjoint_family A"
proof -
obtain A :: "nat \<Rightarrow> 'a set" where
range: "range A \<subseteq> sets M" and
space: "(\<Union>i. A i) = space M" and
measure: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
using sigma_finite by blast
show thesis
proof (rule that[of "disjointed A"])
show "range (disjointed A) \<subseteq> sets M"
by (rule sets.range_disjointed_sets[OF range])
show "(\<Union>i. disjointed A i) = space M"
and "disjoint_family (disjointed A)"
using disjoint_family_disjointed UN_disjointed_eq[of A] space range
by auto
show "emeasure M (disjointed A i) \<noteq> \<infinity>" for i
proof -
have "emeasure M (disjointed A i) \<le> emeasure M (A i)"
using range disjointed_subset[of A i] by (auto intro!: emeasure_mono)
then show ?thesis using measure[of i] by (auto simp: top_unique)
qed
qed
qed
lemma (in sigma_finite_measure) sigma_finite_incseq:
obtains A :: "nat \<Rightarrow> 'a set"
where "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>" "incseq A"
proof -
obtain F :: "nat \<Rightarrow> 'a set" where
F: "range F \<subseteq> sets M" "(\<Union>i. F i) = space M" "\<And>i. emeasure M (F i) \<noteq> \<infinity>"
using sigma_finite by blast
show thesis
proof (rule that[of "\<lambda>n. \<Union>i\<le>n. F i"])
show "range (\<lambda>n. \<Union>i\<le>n. F i) \<subseteq> sets M"
using F by (force simp: incseq_def)
show "(\<Union>n. \<Union>i\<le>n. F i) = space M"
proof -
from F have "\<And>x. x \<in> space M \<Longrightarrow> \<exists>i. x \<in> F i" by auto
with F show ?thesis by fastforce
qed
show "emeasure M (\<Union>i\<le>n. F i) \<noteq> \<infinity>" for n
proof -
have "emeasure M (\<Union>i\<le>n. F i) \<le> (\<Sum>i\<le>n. emeasure M (F i))"
using F by (auto intro!: emeasure_subadditive_finite)
also have "\<dots> < \<infinity>"
using F by (auto simp: sum_Pinfty less_top)
finally show ?thesis by simp
qed
show "incseq (\<lambda>n. \<Union>i\<le>n. F i)"
by (force simp: incseq_def)
qed
qed
lemma (in sigma_finite_measure) approx_PInf_emeasure_with_finite:
fixes C::real
assumes W_meas: "W \<in> sets M"
and W_inf: "emeasure M W = \<infinity>"
obtains Z where "Z \<in> sets M" "Z \<subseteq> W" "emeasure M Z < \<infinity>" "emeasure M Z > C"
proof -
obtain A :: "nat \<Rightarrow> 'a set"
where A: "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>" "incseq A"
using sigma_finite_incseq by blast
define B where "B = (\<lambda>i. W \<inter> A i)"
have B_meas: "\<And>i. B i \<in> sets M" using W_meas \<open>range A \<subseteq> sets M\<close> B_def by blast
have b: "\<And>i. B i \<subseteq> W" using B_def by blast
{ fix i
have "emeasure M (B i) \<le> emeasure M (A i)"
using A by (intro emeasure_mono) (auto simp: B_def)
also have "emeasure M (A i) < \<infinity>"
using \<open>\<And>i. emeasure M (A i) \<noteq> \<infinity>\<close> by (simp add: less_top)
finally have "emeasure M (B i) < \<infinity>" . }
note c = this
have "W = (\<Union>i. B i)" using B_def \<open>(\<Union>i. A i) = space M\<close> W_meas by auto
moreover have "incseq B" using B_def \<open>incseq A\<close> by (simp add: incseq_def subset_eq)
ultimately have "(\<lambda>i. emeasure M (B i)) \<longlonglongrightarrow> emeasure M W" using W_meas B_meas
by (simp add: B_meas Lim_emeasure_incseq image_subset_iff)
then have "(\<lambda>i. emeasure M (B i)) \<longlonglongrightarrow> \<infinity>" using W_inf by simp
from order_tendstoD(1)[OF this, of C]
obtain i where d: "emeasure M (B i) > C"
by (auto simp: eventually_sequentially)
have "B i \<in> sets M" "B i \<subseteq> W" "emeasure M (B i) < \<infinity>" "emeasure M (B i) > C"
using B_meas b c d by auto
then show ?thesis using that by blast
qed
subsection \<open>Measure space induced by distribution of \<^const>\<open>measurable\<close>-functions\<close>
definition\<^marker>\<open>tag important\<close> distr :: "'a measure \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b measure" where
"distr M N f =
measure_of (space N) (sets N) (\<lambda>A. emeasure M (f -` A \<inter> space M))"
lemma
shows sets_distr[simp, measurable_cong]: "sets (distr M N f) = sets N"
and space_distr[simp]: "space (distr M N f) = space N"
by (auto simp: distr_def)
lemma
shows measurable_distr_eq1[simp]: "measurable (distr Mf Nf f) Mf' = measurable Nf Mf'"
and measurable_distr_eq2[simp]: "measurable Mg' (distr Mg Ng g) = measurable Mg' Ng"
by (auto simp: measurable_def)
lemma distr_cong:
"M = K \<Longrightarrow> sets N = sets L \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> distr M N f = distr K L g"
using sets_eq_imp_space_eq[of N L] by (simp add: distr_def Int_def cong: rev_conj_cong)
lemma emeasure_distr:
fixes f :: "'a \<Rightarrow> 'b"
assumes f: "f \<in> measurable M N" and A: "A \<in> sets N"
shows "emeasure (distr M N f) A = emeasure M (f -` A \<inter> space M)" (is "_ = ?\<mu> A")
unfolding distr_def
proof (rule emeasure_measure_of_sigma)
show "positive (sets N) ?\<mu>"
by (auto simp: positive_def)
show "countably_additive (sets N) ?\<mu>"
proof (intro countably_additiveI)
fix A :: "nat \<Rightarrow> 'b set" assume "range A \<subseteq> sets N" "disjoint_family A"
then have A: "\<And>i. A i \<in> sets N" "(\<Union>i. A i) \<in> sets N" by auto
then have *: "range (\<lambda>i. f -` (A i) \<inter> space M) \<subseteq> sets M"
using f by (auto simp: measurable_def)
moreover have "(\<Union>i. f -` A i \<inter> space M) \<in> sets M"
using * by blast
moreover have **: "disjoint_family (\<lambda>i. f -` A i \<inter> space M)"
using \<open>disjoint_family A\<close> by (auto simp: disjoint_family_on_def)
ultimately show "(\<Sum>i. ?\<mu> (A i)) = ?\<mu> (\<Union>i. A i)"
using suminf_emeasure[OF _ **] A f
by (auto simp: comp_def vimage_UN)
qed
show "sigma_algebra (space N) (sets N)" ..
qed fact
lemma emeasure_Collect_distr:
assumes X[measurable]: "X \<in> measurable M N" "Measurable.pred N P"
shows "emeasure (distr M N X) {x\<in>space N. P x} = emeasure M {x\<in>space M. P (X x)}"
by (subst emeasure_distr)
(auto intro!: arg_cong2[where f=emeasure] X(1)[THEN measurable_space])
lemma emeasure_lfp2[consumes 1, case_names cont f measurable]:
assumes "P M"
assumes cont: "sup_continuous F"
assumes f: "\<And>M. P M \<Longrightarrow> f \<in> measurable M' M"
assumes *: "\<And>M A. P M \<Longrightarrow> (\<And>N. P N \<Longrightarrow> Measurable.pred N A) \<Longrightarrow> Measurable.pred M (F A)"
shows "emeasure M' {x\<in>space M'. lfp F (f x)} = (SUP i. emeasure M' {x\<in>space M'. (F ^^ i) (\<lambda>x. False) (f x)})"
proof (subst (1 2) emeasure_Collect_distr[symmetric, where X=f])
show "f \<in> measurable M' M" "f \<in> measurable M' M"
using f[OF \<open>P M\<close>] by auto
{ fix i show "Measurable.pred M ((F ^^ i) (\<lambda>x. False))"
using \<open>P M\<close> by (induction i arbitrary: M) (auto intro!: *) }
show "Measurable.pred M (lfp F)"
using \<open>P M\<close> cont * by (rule measurable_lfp_coinduct[of P])
have "emeasure (distr M' M f) {x \<in> space (distr M' M f). lfp F x} =
(SUP i. emeasure (distr M' M f) {x \<in> space (distr M' M f). (F ^^ i) (\<lambda>x. False) x})"
using \<open>P M\<close>
proof (coinduction arbitrary: M rule: emeasure_lfp')
case (measurable A N) then have "\<And>N. P N \<Longrightarrow> Measurable.pred (distr M' N f) A"
by metis
then have "\<And>N. P N \<Longrightarrow> Measurable.pred N A"
by simp
with \<open>P N\<close>[THEN *] show ?case
by auto
qed fact
then show "emeasure (distr M' M f) {x \<in> space M. lfp F x} =
(SUP i. emeasure (distr M' M f) {x \<in> space M. (F ^^ i) (\<lambda>x. False) x})"
by simp
qed
lemma distr_id[simp]: "distr N N (\<lambda>x. x) = N"
by (rule measure_eqI) (auto simp: emeasure_distr)
lemma distr_id2: "sets M = sets N \<Longrightarrow> distr N M (\<lambda>x. x) = N"
by (rule measure_eqI) (auto simp: emeasure_distr)
lemma measure_distr:
"f \<in> measurable M N \<Longrightarrow> S \<in> sets N \<Longrightarrow> measure (distr M N f) S = measure M (f -` S \<inter> space M)"
by (simp add: emeasure_distr measure_def)
lemma distr_cong_AE:
assumes 1: "M = K" "sets N = sets L" and
2: "(AE x in M. f x = g x)" and "f \<in> measurable M N" and "g \<in> measurable K L"
shows "distr M N f = distr K L g"
proof (rule measure_eqI)
fix A assume "A \<in> sets (distr M N f)"
with assms show "emeasure (distr M N f) A = emeasure (distr K L g) A"
by (auto simp add: emeasure_distr intro!: emeasure_eq_AE measurable_sets)
qed (insert 1, simp)
lemma AE_distrD:
assumes f: "f \<in> measurable M M'"
and AE: "AE x in distr M M' f. P x"
shows "AE x in M. P (f x)"
proof -
from AE[THEN AE_E] guess N .
with f show ?thesis
unfolding eventually_ae_filter
by (intro bexI[of _ "f -` N \<inter> space M"])
(auto simp: emeasure_distr measurable_def)
qed
lemma AE_distr_iff:
assumes f[measurable]: "f \<in> measurable M N" and P[measurable]: "{x \<in> space N. P x} \<in> sets N"
shows "(AE x in distr M N f. P x) \<longleftrightarrow> (AE x in M. P (f x))"
proof (subst (1 2) AE_iff_measurable[OF _ refl])
have "f -` {x\<in>space N. \<not> P x} \<inter> space M = {x \<in> space M. \<not> P (f x)}"
using f[THEN measurable_space] by auto
then show "(emeasure (distr M N f) {x \<in> space (distr M N f). \<not> P x} = 0) =
(emeasure M {x \<in> space M. \<not> P (f x)} = 0)"
by (simp add: emeasure_distr)
qed auto
lemma null_sets_distr_iff:
"f \<in> measurable M N \<Longrightarrow> A \<in> null_sets (distr M N f) \<longleftrightarrow> f -` A \<inter> space M \<in> null_sets M \<and> A \<in> sets N"
by (auto simp add: null_sets_def emeasure_distr)
proposition distr_distr:
"g \<in> measurable N L \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> distr (distr M N f) L g = distr M L (g \<circ> f)"
by (auto simp add: emeasure_distr measurable_space
intro!: arg_cong[where f="emeasure M"] measure_eqI)
subsection\<^marker>\<open>tag unimportant\<close> \<open>Real measure values\<close>
lemma ring_of_finite_sets: "ring_of_sets (space M) {A\<in>sets M. emeasure M A \<noteq> top}"
proof (rule ring_of_setsI)
show "a \<in> {A \<in> sets M. emeasure M A \<noteq> top} \<Longrightarrow> b \<in> {A \<in> sets M. emeasure M A \<noteq> top} \<Longrightarrow>
a \<union> b \<in> {A \<in> sets M. emeasure M A \<noteq> top}" for a b
using emeasure_subadditive[of a M b] by (auto simp: top_unique)
show "a \<in> {A \<in> sets M. emeasure M A \<noteq> top} \<Longrightarrow> b \<in> {A \<in> sets M. emeasure M A \<noteq> top} \<Longrightarrow>
a - b \<in> {A \<in> sets M. emeasure M A \<noteq> top}" for a b
using emeasure_mono[of "a - b" a M] by (auto simp: top_unique)
qed (auto dest: sets.sets_into_space)
lemma measure_nonneg[simp]: "0 \<le> measure M A"
unfolding measure_def by auto
lemma measure_nonneg' [simp]: "\<not> measure M A < 0"
using measure_nonneg not_le by blast
lemma zero_less_measure_iff: "0 < measure M A \<longleftrightarrow> measure M A \<noteq> 0"
using measure_nonneg[of M A] by (auto simp add: le_less)
lemma measure_le_0_iff: "measure M X \<le> 0 \<longleftrightarrow> measure M X = 0"
using measure_nonneg[of M X] by linarith
lemma measure_empty[simp]: "measure M {} = 0"
unfolding measure_def by (simp add: zero_ennreal.rep_eq)
lemma emeasure_eq_ennreal_measure:
"emeasure M A \<noteq> top \<Longrightarrow> emeasure M A = ennreal (measure M A)"
by (cases "emeasure M A" rule: ennreal_cases) (auto simp: measure_def)
lemma measure_zero_top: "emeasure M A = top \<Longrightarrow> measure M A = 0"
- by (simp add: measure_def enn2ereal_top)
+ by (simp add: measure_def)
lemma measure_eq_emeasure_eq_ennreal: "0 \<le> x \<Longrightarrow> emeasure M A = ennreal x \<Longrightarrow> measure M A = x"
using emeasure_eq_ennreal_measure[of M A]
by (cases "A \<in> M") (auto simp: measure_notin_sets emeasure_notin_sets)
lemma enn2real_plus:"a < top \<Longrightarrow> b < top \<Longrightarrow> enn2real (a + b) = enn2real a + enn2real b"
by (simp add: enn2real_def plus_ennreal.rep_eq real_of_ereal_add less_top
del: real_of_ereal_enn2ereal)
lemma enn2real_sum:"(\<And>i. i \<in> I \<Longrightarrow> f i < top) \<Longrightarrow> enn2real (sum f I) = sum (enn2real \<circ> f) I"
by (induction I rule: infinite_finite_induct) (auto simp: enn2real_plus)
lemma measure_eq_AE:
assumes iff: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B"
assumes A: "A \<in> sets M" and B: "B \<in> sets M"
shows "measure M A = measure M B"
using assms emeasure_eq_AE[OF assms] by (simp add: measure_def)
lemma measure_Union:
"emeasure M A \<noteq> \<infinity> \<Longrightarrow> emeasure M B \<noteq> \<infinity> \<Longrightarrow> A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> A \<inter> B = {} \<Longrightarrow>
measure M (A \<union> B) = measure M A + measure M B"
by (simp add: measure_def plus_emeasure[symmetric] enn2real_plus less_top)
lemma disjoint_family_on_insert:
"i \<notin> I \<Longrightarrow> disjoint_family_on A (insert i I) \<longleftrightarrow> A i \<inter> (\<Union>i\<in>I. A i) = {} \<and> disjoint_family_on A I"
by (fastforce simp: disjoint_family_on_def)
lemma measure_finite_Union:
"finite S \<Longrightarrow> A`S \<subseteq> sets M \<Longrightarrow> disjoint_family_on A S \<Longrightarrow> (\<And>i. i \<in> S \<Longrightarrow> emeasure M (A i) \<noteq> \<infinity>) \<Longrightarrow>
measure M (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. measure M (A i))"
by (induction S rule: finite_induct)
(auto simp: disjoint_family_on_insert measure_Union sum_emeasure[symmetric] sets.countable_UN'[OF countable_finite])
lemma measure_Diff:
assumes finite: "emeasure M A \<noteq> \<infinity>"
and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
shows "measure M (A - B) = measure M A - measure M B"
proof -
have "emeasure M (A - B) \<le> emeasure M A" "emeasure M B \<le> emeasure M A"
using measurable by (auto intro!: emeasure_mono)
hence "measure M ((A - B) \<union> B) = measure M (A - B) + measure M B"
using measurable finite by (rule_tac measure_Union) (auto simp: top_unique)
thus ?thesis using \<open>B \<subseteq> A\<close> by (auto simp: Un_absorb2)
qed
lemma measure_UNION:
assumes measurable: "range A \<subseteq> sets M" "disjoint_family A"
assumes finite: "emeasure M (\<Union>i. A i) \<noteq> \<infinity>"
shows "(\<lambda>i. measure M (A i)) sums (measure M (\<Union>i. A i))"
proof -
have "(\<lambda>i. emeasure M (A i)) sums (emeasure M (\<Union>i. A i))"
unfolding suminf_emeasure[OF measurable, symmetric] by (simp add: summable_sums)
moreover
{ fix i
have "emeasure M (A i) \<le> emeasure M (\<Union>i. A i)"
using measurable by (auto intro!: emeasure_mono)
then have "emeasure M (A i) = ennreal ((measure M (A i)))"
using finite by (intro emeasure_eq_ennreal_measure) (auto simp: top_unique) }
ultimately show ?thesis using finite
by (subst (asm) (2) emeasure_eq_ennreal_measure) simp_all
qed
lemma measure_subadditive:
assumes measurable: "A \<in> sets M" "B \<in> sets M"
and fin: "emeasure M A \<noteq> \<infinity>" "emeasure M B \<noteq> \<infinity>"
shows "measure M (A \<union> B) \<le> measure M A + measure M B"
proof -
have "emeasure M (A \<union> B) \<noteq> \<infinity>"
using emeasure_subadditive[OF measurable] fin by (auto simp: top_unique)
then show "(measure M (A \<union> B)) \<le> (measure M A) + (measure M B)"
using emeasure_subadditive[OF measurable] fin
apply simp
apply (subst (asm) (2 3 4) emeasure_eq_ennreal_measure)
apply (auto simp flip: ennreal_plus)
done
qed
lemma measure_subadditive_finite:
assumes A: "finite I" "A`I \<subseteq> sets M" and fin: "\<And>i. i \<in> I \<Longrightarrow> emeasure M (A i) \<noteq> \<infinity>"
shows "measure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. measure M (A i))"
proof -
{ have "emeasure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. emeasure M (A i))"
using emeasure_subadditive_finite[OF A] .
also have "\<dots> < \<infinity>"
using fin by (simp add: less_top A)
finally have "emeasure M (\<Union>i\<in>I. A i) \<noteq> top" by simp }
note * = this
show ?thesis
using emeasure_subadditive_finite[OF A] fin
unfolding emeasure_eq_ennreal_measure[OF *]
by (simp_all add: sum_nonneg emeasure_eq_ennreal_measure)
qed
lemma measure_subadditive_countably:
assumes A: "range A \<subseteq> sets M" and fin: "(\<Sum>i. emeasure M (A i)) \<noteq> \<infinity>"
shows "measure M (\<Union>i. A i) \<le> (\<Sum>i. measure M (A i))"
proof -
from fin have **: "\<And>i. emeasure M (A i) \<noteq> top"
using ennreal_suminf_lessD[of "\<lambda>i. emeasure M (A i)"] by (simp add: less_top)
{ have "emeasure M (\<Union>i. A i) \<le> (\<Sum>i. emeasure M (A i))"
using emeasure_subadditive_countably[OF A] .
also have "\<dots> < \<infinity>"
using fin by (simp add: less_top)
finally have "emeasure M (\<Union>i. A i) \<noteq> top" by simp }
then have "ennreal (measure M (\<Union>i. A i)) = emeasure M (\<Union>i. A i)"
by (rule emeasure_eq_ennreal_measure[symmetric])
also have "\<dots> \<le> (\<Sum>i. emeasure M (A i))"
using emeasure_subadditive_countably[OF A] .
also have "\<dots> = ennreal (\<Sum>i. measure M (A i))"
using fin unfolding emeasure_eq_ennreal_measure[OF **]
by (subst suminf_ennreal) (auto simp: **)
finally show ?thesis
apply (rule ennreal_le_iff[THEN iffD1, rotated])
apply (intro suminf_nonneg allI measure_nonneg summable_suminf_not_top)
using fin
apply (simp add: emeasure_eq_ennreal_measure[OF **])
done
qed
lemma measure_Un_null_set: "A \<in> sets M \<Longrightarrow> B \<in> null_sets M \<Longrightarrow> measure M (A \<union> B) = measure M A"
by (simp add: measure_def emeasure_Un_null_set)
lemma measure_Diff_null_set: "A \<in> sets M \<Longrightarrow> B \<in> null_sets M \<Longrightarrow> measure M (A - B) = measure M A"
by (simp add: measure_def emeasure_Diff_null_set)
lemma measure_eq_sum_singleton:
"finite S \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M) \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> emeasure M {x} \<noteq> \<infinity>) \<Longrightarrow>
measure M S = (\<Sum>x\<in>S. measure M {x})"
using emeasure_eq_sum_singleton[of S M]
by (intro measure_eq_emeasure_eq_ennreal) (auto simp: sum_nonneg emeasure_eq_ennreal_measure)
lemma Lim_measure_incseq:
assumes A: "range A \<subseteq> sets M" "incseq A" and fin: "emeasure M (\<Union>i. A i) \<noteq> \<infinity>"
shows "(\<lambda>i. measure M (A i)) \<longlonglongrightarrow> measure M (\<Union>i. A i)"
proof (rule tendsto_ennrealD)
have "ennreal (measure M (\<Union>i. A i)) = emeasure M (\<Union>i. A i)"
using fin by (auto simp: emeasure_eq_ennreal_measure)
moreover have "ennreal (measure M (A i)) = emeasure M (A i)" for i
using assms emeasure_mono[of "A _" "\<Union>i. A i" M]
by (intro emeasure_eq_ennreal_measure[symmetric]) (auto simp: less_top UN_upper intro: le_less_trans)
ultimately show "(\<lambda>x. ennreal (measure M (A x))) \<longlonglongrightarrow> ennreal (measure M (\<Union>i. A i))"
using A by (auto intro!: Lim_emeasure_incseq)
qed auto
lemma Lim_measure_decseq:
assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
shows "(\<lambda>n. measure M (A n)) \<longlonglongrightarrow> measure M (\<Inter>i. A i)"
proof (rule tendsto_ennrealD)
have "ennreal (measure M (\<Inter>i. A i)) = emeasure M (\<Inter>i. A i)"
using fin[of 0] A emeasure_mono[of "\<Inter>i. A i" "A 0" M]
by (auto intro!: emeasure_eq_ennreal_measure[symmetric] simp: INT_lower less_top intro: le_less_trans)
moreover have "ennreal (measure M (A i)) = emeasure M (A i)" for i
using A fin[of i] by (intro emeasure_eq_ennreal_measure[symmetric]) auto
ultimately show "(\<lambda>x. ennreal (measure M (A x))) \<longlonglongrightarrow> ennreal (measure M (\<Inter>i. A i))"
using fin A by (auto intro!: Lim_emeasure_decseq)
qed auto
subsection \<open>Set of measurable sets with finite measure\<close>
definition\<^marker>\<open>tag important\<close> fmeasurable :: "'a measure \<Rightarrow> 'a set set" where
"fmeasurable M = {A\<in>sets M. emeasure M A < \<infinity>}"
lemma fmeasurableD[dest, measurable_dest]: "A \<in> fmeasurable M \<Longrightarrow> A \<in> sets M"
by (auto simp: fmeasurable_def)
lemma fmeasurableD2: "A \<in> fmeasurable M \<Longrightarrow> emeasure M A \<noteq> top"
by (auto simp: fmeasurable_def)
lemma fmeasurableI: "A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow> A \<in> fmeasurable M"
by (auto simp: fmeasurable_def)
lemma fmeasurableI_null_sets: "A \<in> null_sets M \<Longrightarrow> A \<in> fmeasurable M"
by (auto simp: fmeasurable_def)
lemma fmeasurableI2: "A \<in> fmeasurable M \<Longrightarrow> B \<subseteq> A \<Longrightarrow> B \<in> sets M \<Longrightarrow> B \<in> fmeasurable M"
using emeasure_mono[of B A M] by (auto simp: fmeasurable_def)
lemma measure_mono_fmeasurable:
"A \<subseteq> B \<Longrightarrow> A \<in> sets M \<Longrightarrow> B \<in> fmeasurable M \<Longrightarrow> measure M A \<le> measure M B"
by (auto simp: measure_def fmeasurable_def intro!: emeasure_mono enn2real_mono)
lemma emeasure_eq_measure2: "A \<in> fmeasurable M \<Longrightarrow> emeasure M A = measure M A"
by (simp add: emeasure_eq_ennreal_measure fmeasurable_def less_top)
interpretation fmeasurable: ring_of_sets "space M" "fmeasurable M"
proof (rule ring_of_setsI)
show "fmeasurable M \<subseteq> Pow (space M)" "{} \<in> fmeasurable M"
by (auto simp: fmeasurable_def dest: sets.sets_into_space)
fix a b assume *: "a \<in> fmeasurable M" "b \<in> fmeasurable M"
then have "emeasure M (a \<union> b) \<le> emeasure M a + emeasure M b"
by (intro emeasure_subadditive) auto
also have "\<dots> < top"
using * by (auto simp: fmeasurable_def)
finally show "a \<union> b \<in> fmeasurable M"
using * by (auto intro: fmeasurableI)
show "a - b \<in> fmeasurable M"
using emeasure_mono[of "a - b" a M] * by (auto simp: fmeasurable_def)
qed
subsection\<^marker>\<open>tag unimportant\<close>\<open>Measurable sets formed by unions and intersections\<close>
lemma fmeasurable_Diff: "A \<in> fmeasurable M \<Longrightarrow> B \<in> sets M \<Longrightarrow> A - B \<in> fmeasurable M"
using fmeasurableI2[of A M "A - B"] by auto
lemma fmeasurable_Int_fmeasurable:
"\<lbrakk>S \<in> fmeasurable M; T \<in> sets M\<rbrakk> \<Longrightarrow> (S \<inter> T) \<in> fmeasurable M"
by (meson fmeasurableD fmeasurableI2 inf_le1 sets.Int)
lemma fmeasurable_UN:
assumes "countable I" "\<And>i. i \<in> I \<Longrightarrow> F i \<subseteq> A" "\<And>i. i \<in> I \<Longrightarrow> F i \<in> sets M" "A \<in> fmeasurable M"
shows "(\<Union>i\<in>I. F i) \<in> fmeasurable M"
proof (rule fmeasurableI2)
show "A \<in> fmeasurable M" "(\<Union>i\<in>I. F i) \<subseteq> A" using assms by auto
show "(\<Union>i\<in>I. F i) \<in> sets M"
using assms by (intro sets.countable_UN') auto
qed
lemma fmeasurable_INT:
assumes "countable I" "i \<in> I" "\<And>i. i \<in> I \<Longrightarrow> F i \<in> sets M" "F i \<in> fmeasurable M"
shows "(\<Inter>i\<in>I. F i) \<in> fmeasurable M"
proof (rule fmeasurableI2)
show "F i \<in> fmeasurable M" "(\<Inter>i\<in>I. F i) \<subseteq> F i"
using assms by auto
show "(\<Inter>i\<in>I. F i) \<in> sets M"
using assms by (intro sets.countable_INT') auto
qed
lemma measurable_measure_Diff:
assumes "A \<in> fmeasurable M" "B \<in> sets M" "B \<subseteq> A"
shows "measure M (A - B) = measure M A - measure M B"
by (simp add: assms fmeasurableD fmeasurableD2 measure_Diff)
lemma measurable_Un_null_set:
assumes "B \<in> null_sets M"
shows "(A \<union> B \<in> fmeasurable M \<and> A \<in> sets M) \<longleftrightarrow> A \<in> fmeasurable M"
using assms by (fastforce simp add: fmeasurable.Un fmeasurableI_null_sets intro: fmeasurableI2)
lemma measurable_Diff_null_set:
assumes "B \<in> null_sets M"
shows "(A - B) \<in> fmeasurable M \<and> A \<in> sets M \<longleftrightarrow> A \<in> fmeasurable M"
using assms
by (metis Un_Diff_cancel2 fmeasurable.Diff fmeasurableD fmeasurableI_null_sets measurable_Un_null_set)
lemma fmeasurable_Diff_D:
assumes m: "T - S \<in> fmeasurable M" "S \<in> fmeasurable M" and sub: "S \<subseteq> T"
shows "T \<in> fmeasurable M"
proof -
have "T = S \<union> (T - S)"
using assms by blast
then show ?thesis
by (metis m fmeasurable.Un)
qed
lemma measure_Un2:
"A \<in> fmeasurable M \<Longrightarrow> B \<in> fmeasurable M \<Longrightarrow> measure M (A \<union> B) = measure M A + measure M (B - A)"
using measure_Union[of M A "B - A"] by (auto simp: fmeasurableD2 fmeasurable.Diff)
lemma measure_Un3:
assumes "A \<in> fmeasurable M" "B \<in> fmeasurable M"
shows "measure M (A \<union> B) = measure M A + measure M B - measure M (A \<inter> B)"
proof -
have "measure M (A \<union> B) = measure M A + measure M (B - A)"
using assms by (rule measure_Un2)
also have "B - A = B - (A \<inter> B)"
by auto
also have "measure M (B - (A \<inter> B)) = measure M B - measure M (A \<inter> B)"
using assms by (intro measure_Diff) (auto simp: fmeasurable_def)
finally show ?thesis
by simp
qed
lemma measure_Un_AE:
"AE x in M. x \<notin> A \<or> x \<notin> B \<Longrightarrow> A \<in> fmeasurable M \<Longrightarrow> B \<in> fmeasurable M \<Longrightarrow>
measure M (A \<union> B) = measure M A + measure M B"
by (subst measure_Un2) (auto intro!: measure_eq_AE)
lemma measure_UNION_AE:
assumes I: "finite I"
shows "(\<And>i. i \<in> I \<Longrightarrow> F i \<in> fmeasurable M) \<Longrightarrow> pairwise (\<lambda>i j. AE x in M. x \<notin> F i \<or> x \<notin> F j) I \<Longrightarrow>
measure M (\<Union>i\<in>I. F i) = (\<Sum>i\<in>I. measure M (F i))"
unfolding AE_pairwise[OF countable_finite, OF I]
using I
proof (induction I rule: finite_induct)
case (insert x I)
have "measure M (F x \<union> \<Union>(F ` I)) = measure M (F x) + measure M (\<Union>(F ` I))"
by (rule measure_Un_AE) (use insert in \<open>auto simp: pairwise_insert\<close>)
with insert show ?case
by (simp add: pairwise_insert )
qed simp
lemma measure_UNION':
"finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> F i \<in> fmeasurable M) \<Longrightarrow> pairwise (\<lambda>i j. disjnt (F i) (F j)) I \<Longrightarrow>
measure M (\<Union>i\<in>I. F i) = (\<Sum>i\<in>I. measure M (F i))"
by (intro measure_UNION_AE) (auto simp: disjnt_def elim!: pairwise_mono intro!: always_eventually)
lemma measure_Union_AE:
"finite F \<Longrightarrow> (\<And>S. S \<in> F \<Longrightarrow> S \<in> fmeasurable M) \<Longrightarrow> pairwise (\<lambda>S T. AE x in M. x \<notin> S \<or> x \<notin> T) F \<Longrightarrow>
measure M (\<Union>F) = (\<Sum>S\<in>F. measure M S)"
using measure_UNION_AE[of F "\<lambda>x. x" M] by simp
lemma measure_Union':
"finite F \<Longrightarrow> (\<And>S. S \<in> F \<Longrightarrow> S \<in> fmeasurable M) \<Longrightarrow> pairwise disjnt F \<Longrightarrow> measure M (\<Union>F) = (\<Sum>S\<in>F. measure M S)"
using measure_UNION'[of F "\<lambda>x. x" M] by simp
lemma measure_Un_le:
assumes "A \<in> sets M" "B \<in> sets M" shows "measure M (A \<union> B) \<le> measure M A + measure M B"
proof cases
assume "A \<in> fmeasurable M \<and> B \<in> fmeasurable M"
with measure_subadditive[of A M B] assms show ?thesis
by (auto simp: fmeasurableD2)
next
assume "\<not> (A \<in> fmeasurable M \<and> B \<in> fmeasurable M)"
then have "A \<union> B \<notin> fmeasurable M"
using fmeasurableI2[of "A \<union> B" M A] fmeasurableI2[of "A \<union> B" M B] assms by auto
with assms show ?thesis
by (auto simp: fmeasurable_def measure_def less_top[symmetric])
qed
lemma measure_UNION_le:
"finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> F i \<in> sets M) \<Longrightarrow> measure M (\<Union>i\<in>I. F i) \<le> (\<Sum>i\<in>I. measure M (F i))"
proof (induction I rule: finite_induct)
case (insert i I)
then have "measure M (\<Union>i\<in>insert i I. F i) \<le> measure M (F i) + measure M (\<Union>i\<in>I. F i)"
by (auto intro!: measure_Un_le)
also have "measure M (\<Union>i\<in>I. F i) \<le> (\<Sum>i\<in>I. measure M (F i))"
using insert by auto
finally show ?case
using insert by simp
qed simp
lemma measure_Union_le:
"finite F \<Longrightarrow> (\<And>S. S \<in> F \<Longrightarrow> S \<in> sets M) \<Longrightarrow> measure M (\<Union>F) \<le> (\<Sum>S\<in>F. measure M S)"
using measure_UNION_le[of F "\<lambda>x. x" M] by simp
text\<open>Version for indexed union over a countable set\<close>
lemma
assumes "countable I" and I: "\<And>i. i \<in> I \<Longrightarrow> A i \<in> fmeasurable M"
and bound: "\<And>I'. I' \<subseteq> I \<Longrightarrow> finite I' \<Longrightarrow> measure M (\<Union>i\<in>I'. A i) \<le> B"
shows fmeasurable_UN_bound: "(\<Union>i\<in>I. A i) \<in> fmeasurable M" (is ?fm)
and measure_UN_bound: "measure M (\<Union>i\<in>I. A i) \<le> B" (is ?m)
proof -
have "B \<ge> 0"
using bound by force
have "?fm \<and> ?m"
proof cases
assume "I = {}"
with \<open>B \<ge> 0\<close> show ?thesis
by simp
next
assume "I \<noteq> {}"
have "(\<Union>i\<in>I. A i) = (\<Union>i. (\<Union>n\<le>i. A (from_nat_into I n)))"
by (subst range_from_nat_into[symmetric, OF \<open>I \<noteq> {}\<close> \<open>countable I\<close>]) auto
then have "emeasure M (\<Union>i\<in>I. A i) = emeasure M (\<Union>i. (\<Union>n\<le>i. A (from_nat_into I n)))" by simp
also have "\<dots> = (SUP i. emeasure M (\<Union>n\<le>i. A (from_nat_into I n)))"
using I \<open>I \<noteq> {}\<close>[THEN from_nat_into] by (intro SUP_emeasure_incseq[symmetric]) (fastforce simp: incseq_Suc_iff)+
also have "\<dots> \<le> B"
proof (intro SUP_least)
fix i :: nat
have "emeasure M (\<Union>n\<le>i. A (from_nat_into I n)) = measure M (\<Union>n\<le>i. A (from_nat_into I n))"
using I \<open>I \<noteq> {}\<close>[THEN from_nat_into] by (intro emeasure_eq_measure2 fmeasurable.finite_UN) auto
also have "\<dots> = measure M (\<Union>n\<in>from_nat_into I ` {..i}. A n)"
by simp
also have "\<dots> \<le> B"
by (intro ennreal_leI bound) (auto intro: from_nat_into[OF \<open>I \<noteq> {}\<close>])
finally show "emeasure M (\<Union>n\<le>i. A (from_nat_into I n)) \<le> ennreal B" .
qed
finally have *: "emeasure M (\<Union>i\<in>I. A i) \<le> B" .
then have ?fm
using I \<open>countable I\<close> by (intro fmeasurableI conjI) (auto simp: less_top[symmetric] top_unique)
with * \<open>0\<le>B\<close> show ?thesis
by (simp add: emeasure_eq_measure2)
qed
then show ?fm ?m by auto
qed
text\<open>Version for big union of a countable set\<close>
lemma
assumes "countable \<D>"
and meas: "\<And>D. D \<in> \<D> \<Longrightarrow> D \<in> fmeasurable M"
and bound: "\<And>\<E>. \<lbrakk>\<E> \<subseteq> \<D>; finite \<E>\<rbrakk> \<Longrightarrow> measure M (\<Union>\<E>) \<le> B"
shows fmeasurable_Union_bound: "\<Union>\<D> \<in> fmeasurable M" (is ?fm)
and measure_Union_bound: "measure M (\<Union>\<D>) \<le> B" (is ?m)
proof -
have "B \<ge> 0"
using bound by force
have "?fm \<and> ?m"
proof (cases "\<D> = {}")
case True
with \<open>B \<ge> 0\<close> show ?thesis
by auto
next
case False
then obtain D :: "nat \<Rightarrow> 'a set" where D: "\<D> = range D"
using \<open>countable \<D>\<close> uncountable_def by force
have 1: "\<And>i. D i \<in> fmeasurable M"
by (simp add: D meas)
have 2: "\<And>I'. finite I' \<Longrightarrow> measure M (\<Union>x\<in>I'. D x) \<le> B"
by (simp add: D bound image_subset_iff)
show ?thesis
unfolding D
by (intro conjI fmeasurable_UN_bound [OF _ 1 2] measure_UN_bound [OF _ 1 2]) auto
qed
then show ?fm ?m by auto
qed
text\<open>Version for indexed union over the type of naturals\<close>
lemma
fixes S :: "nat \<Rightarrow> 'a set"
assumes S: "\<And>i. S i \<in> fmeasurable M" and B: "\<And>n. measure M (\<Union>i\<le>n. S i) \<le> B"
shows fmeasurable_countable_Union: "(\<Union>i. S i) \<in> fmeasurable M"
and measure_countable_Union_le: "measure M (\<Union>i. S i) \<le> B"
proof -
have mB: "measure M (\<Union>i\<in>I. S i) \<le> B" if "finite I" for I
proof -
have "(\<Union>i\<in>I. S i) \<subseteq> (\<Union>i\<le>Max I. S i)"
using Max_ge that by force
then have "measure M (\<Union>i\<in>I. S i) \<le> measure M (\<Union>i \<le> Max I. S i)"
by (rule measure_mono_fmeasurable) (use S in \<open>blast+\<close>)
then show ?thesis
using B order_trans by blast
qed
show "(\<Union>i. S i) \<in> fmeasurable M"
by (auto intro: fmeasurable_UN_bound [OF _ S mB])
show "measure M (\<Union>n. S n) \<le> B"
by (auto intro: measure_UN_bound [OF _ S mB])
qed
lemma measure_diff_le_measure_setdiff:
assumes "S \<in> fmeasurable M" "T \<in> fmeasurable M"
shows "measure M S - measure M T \<le> measure M (S - T)"
proof -
have "measure M S \<le> measure M ((S - T) \<union> T)"
by (simp add: assms fmeasurable.Un fmeasurableD measure_mono_fmeasurable)
also have "\<dots> \<le> measure M (S - T) + measure M T"
using assms by (blast intro: measure_Un_le)
finally show ?thesis
by (simp add: algebra_simps)
qed
lemma suminf_exist_split2:
fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
assumes "summable f"
shows "(\<lambda>n. (\<Sum>k. f(k+n))) \<longlonglongrightarrow> 0"
by (subst lim_sequentially, auto simp add: dist_norm suminf_exist_split[OF _ assms])
lemma emeasure_union_summable:
assumes [measurable]: "\<And>n. A n \<in> sets M"
and "\<And>n. emeasure M (A n) < \<infinity>" "summable (\<lambda>n. measure M (A n))"
shows "emeasure M (\<Union>n. A n) < \<infinity>" "emeasure M (\<Union>n. A n) \<le> (\<Sum>n. measure M (A n))"
proof -
define B where "B = (\<lambda>N. (\<Union>n\<in>{..<N}. A n))"
have [measurable]: "B N \<in> sets M" for N unfolding B_def by auto
have "(\<lambda>N. emeasure M (B N)) \<longlonglongrightarrow> emeasure M (\<Union>N. B N)"
apply (rule Lim_emeasure_incseq) unfolding B_def by (auto simp add: SUP_subset_mono incseq_def)
moreover have "emeasure M (B N) \<le> ennreal (\<Sum>n. measure M (A n))" for N
proof -
have *: "(\<Sum>n\<in>{..<N}. measure M (A n)) \<le> (\<Sum>n. measure M (A n))"
using assms(3) measure_nonneg sum_le_suminf by blast
have "emeasure M (B N) \<le> (\<Sum>n\<in>{..<N}. emeasure M (A n))"
unfolding B_def by (rule emeasure_subadditive_finite, auto)
also have "... = (\<Sum>n\<in>{..<N}. ennreal(measure M (A n)))"
using assms(2) by (simp add: emeasure_eq_ennreal_measure less_top)
also have "... = ennreal (\<Sum>n\<in>{..<N}. measure M (A n))"
by auto
also have "... \<le> ennreal (\<Sum>n. measure M (A n))"
using * by (auto simp: ennreal_leI)
finally show ?thesis by simp
qed
ultimately have "emeasure M (\<Union>N. B N) \<le> ennreal (\<Sum>n. measure M (A n))"
by (simp add: Lim_bounded)
then show "emeasure M (\<Union>n. A n) \<le> (\<Sum>n. measure M (A n))"
unfolding B_def by (metis UN_UN_flatten UN_lessThan_UNIV)
then show "emeasure M (\<Union>n. A n) < \<infinity>"
by (auto simp: less_top[symmetric] top_unique)
qed
lemma borel_cantelli_limsup1:
assumes [measurable]: "\<And>n. A n \<in> sets M"
and "\<And>n. emeasure M (A n) < \<infinity>" "summable (\<lambda>n. measure M (A n))"
shows "limsup A \<in> null_sets M"
proof -
have "emeasure M (limsup A) \<le> 0"
proof (rule LIMSEQ_le_const)
have "(\<lambda>n. (\<Sum>k. measure M (A (k+n)))) \<longlonglongrightarrow> 0" by (rule suminf_exist_split2[OF assms(3)])
then show "(\<lambda>n. ennreal (\<Sum>k. measure M (A (k+n)))) \<longlonglongrightarrow> 0"
unfolding ennreal_0[symmetric] by (intro tendsto_ennrealI)
have "emeasure M (limsup A) \<le> (\<Sum>k. measure M (A (k+n)))" for n
proof -
have I: "(\<Union>k\<in>{n..}. A k) = (\<Union>k. A (k+n))" by (auto, metis le_add_diff_inverse2, fastforce)
have "emeasure M (limsup A) \<le> emeasure M (\<Union>k\<in>{n..}. A k)"
by (rule emeasure_mono, auto simp add: limsup_INF_SUP)
also have "... = emeasure M (\<Union>k. A (k+n))"
using I by auto
also have "... \<le> (\<Sum>k. measure M (A (k+n)))"
apply (rule emeasure_union_summable)
using assms summable_ignore_initial_segment[OF assms(3), of n] by auto
finally show ?thesis by simp
qed
then show "\<exists>N. \<forall>n\<ge>N. emeasure M (limsup A) \<le> (\<Sum>k. measure M (A (k+n)))"
by auto
qed
then show ?thesis using assms(1) measurable_limsup by auto
qed
lemma borel_cantelli_AE1:
assumes [measurable]: "\<And>n. A n \<in> sets M"
and "\<And>n. emeasure M (A n) < \<infinity>" "summable (\<lambda>n. measure M (A n))"
shows "AE x in M. eventually (\<lambda>n. x \<in> space M - A n) sequentially"
proof -
have "AE x in M. x \<notin> limsup A"
using borel_cantelli_limsup1[OF assms] unfolding eventually_ae_filter by auto
moreover
{
fix x assume "x \<notin> limsup A"
then obtain N where "x \<notin> (\<Union>n\<in>{N..}. A n)" unfolding limsup_INF_SUP by blast
then have "eventually (\<lambda>n. x \<notin> A n) sequentially" using eventually_sequentially by auto
}
ultimately show ?thesis by auto
qed
subsection \<open>Measure spaces with \<^term>\<open>emeasure M (space M) < \<infinity>\<close>\<close>
locale\<^marker>\<open>tag important\<close> finite_measure = sigma_finite_measure M for M +
assumes finite_emeasure_space: "emeasure M (space M) \<noteq> top"
lemma finite_measureI[Pure.intro!]:
"emeasure M (space M) \<noteq> \<infinity> \<Longrightarrow> finite_measure M"
proof qed (auto intro!: exI[of _ "{space M}"])
lemma (in finite_measure) emeasure_finite[simp, intro]: "emeasure M A \<noteq> top"
using finite_emeasure_space emeasure_space[of M A] by (auto simp: top_unique)
lemma (in finite_measure) fmeasurable_eq_sets: "fmeasurable M = sets M"
by (auto simp: fmeasurable_def less_top[symmetric])
lemma (in finite_measure) emeasure_eq_measure: "emeasure M A = ennreal (measure M A)"
by (intro emeasure_eq_ennreal_measure) simp
lemma (in finite_measure) emeasure_real: "\<exists>r. 0 \<le> r \<and> emeasure M A = ennreal r"
using emeasure_finite[of A] by (cases "emeasure M A" rule: ennreal_cases) auto
lemma (in finite_measure) bounded_measure: "measure M A \<le> measure M (space M)"
using emeasure_space[of M A] emeasure_real[of A] emeasure_real[of "space M"] by (auto simp: measure_def)
lemma (in finite_measure) finite_measure_Diff:
assumes sets: "A \<in> sets M" "B \<in> sets M" and "B \<subseteq> A"
shows "measure M (A - B) = measure M A - measure M B"
using measure_Diff[OF _ assms] by simp
lemma (in finite_measure) finite_measure_Union:
assumes sets: "A \<in> sets M" "B \<in> sets M" and "A \<inter> B = {}"
shows "measure M (A \<union> B) = measure M A + measure M B"
using measure_Union[OF _ _ assms] by simp
lemma (in finite_measure) finite_measure_finite_Union:
assumes measurable: "finite S" "A`S \<subseteq> sets M" "disjoint_family_on A S"
shows "measure M (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. measure M (A i))"
using measure_finite_Union[OF assms] by simp
lemma (in finite_measure) finite_measure_UNION:
assumes A: "range A \<subseteq> sets M" "disjoint_family A"
shows "(\<lambda>i. measure M (A i)) sums (measure M (\<Union>i. A i))"
using measure_UNION[OF A] by simp
lemma (in finite_measure) finite_measure_mono:
assumes "A \<subseteq> B" "B \<in> sets M" shows "measure M A \<le> measure M B"
using emeasure_mono[OF assms] emeasure_real[of A] emeasure_real[of B] by (auto simp: measure_def)
lemma (in finite_measure) finite_measure_subadditive:
assumes m: "A \<in> sets M" "B \<in> sets M"
shows "measure M (A \<union> B) \<le> measure M A + measure M B"
using measure_subadditive[OF m] by simp
lemma (in finite_measure) finite_measure_subadditive_finite:
assumes "finite I" "A`I \<subseteq> sets M" shows "measure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. measure M (A i))"
using measure_subadditive_finite[OF assms] by simp
lemma (in finite_measure) finite_measure_subadditive_countably:
"range A \<subseteq> sets M \<Longrightarrow> summable (\<lambda>i. measure M (A i)) \<Longrightarrow> measure M (\<Union>i. A i) \<le> (\<Sum>i. measure M (A i))"
by (rule measure_subadditive_countably)
(simp_all add: ennreal_suminf_neq_top emeasure_eq_measure)
lemma (in finite_measure) finite_measure_eq_sum_singleton:
assumes "finite S" and *: "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
shows "measure M S = (\<Sum>x\<in>S. measure M {x})"
using measure_eq_sum_singleton[OF assms] by simp
lemma (in finite_measure) finite_Lim_measure_incseq:
assumes A: "range A \<subseteq> sets M" "incseq A"
shows "(\<lambda>i. measure M (A i)) \<longlonglongrightarrow> measure M (\<Union>i. A i)"
using Lim_measure_incseq[OF A] by simp
lemma (in finite_measure) finite_Lim_measure_decseq:
assumes A: "range A \<subseteq> sets M" "decseq A"
shows "(\<lambda>n. measure M (A n)) \<longlonglongrightarrow> measure M (\<Inter>i. A i)"
using Lim_measure_decseq[OF A] by simp
lemma (in finite_measure) finite_measure_compl:
assumes S: "S \<in> sets M"
shows "measure M (space M - S) = measure M (space M) - measure M S"
using measure_Diff[OF _ sets.top S sets.sets_into_space] S by simp
lemma (in finite_measure) finite_measure_mono_AE:
assumes imp: "AE x in M. x \<in> A \<longrightarrow> x \<in> B" and B: "B \<in> sets M"
shows "measure M A \<le> measure M B"
using assms emeasure_mono_AE[OF imp B]
by (simp add: emeasure_eq_measure)
lemma (in finite_measure) finite_measure_eq_AE:
assumes iff: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B"
assumes A: "A \<in> sets M" and B: "B \<in> sets M"
shows "measure M A = measure M B"
using assms emeasure_eq_AE[OF assms] by (simp add: emeasure_eq_measure)
lemma (in finite_measure) measure_increasing: "increasing M (measure M)"
by (auto intro!: finite_measure_mono simp: increasing_def)
lemma (in finite_measure) measure_zero_union:
assumes "s \<in> sets M" "t \<in> sets M" "measure M t = 0"
shows "measure M (s \<union> t) = measure M s"
using assms
proof -
have "measure M (s \<union> t) \<le> measure M s"
using finite_measure_subadditive[of s t] assms by auto
moreover have "measure M (s \<union> t) \<ge> measure M s"
using assms by (blast intro: finite_measure_mono)
ultimately show ?thesis by simp
qed
lemma (in finite_measure) measure_eq_compl:
assumes "s \<in> sets M" "t \<in> sets M"
assumes "measure M (space M - s) = measure M (space M - t)"
shows "measure M s = measure M t"
using assms finite_measure_compl by auto
lemma (in finite_measure) measure_eq_bigunion_image:
assumes "range f \<subseteq> sets M" "range g \<subseteq> sets M"
assumes "disjoint_family f" "disjoint_family g"
assumes "\<And> n :: nat. measure M (f n) = measure M (g n)"
shows "measure M (\<Union>i. f i) = measure M (\<Union>i. g i)"
using assms
proof -
have a: "(\<lambda> i. measure M (f i)) sums (measure M (\<Union>i. f i))"
by (rule finite_measure_UNION[OF assms(1,3)])
have b: "(\<lambda> i. measure M (g i)) sums (measure M (\<Union>i. g i))"
by (rule finite_measure_UNION[OF assms(2,4)])
show ?thesis using sums_unique[OF b] sums_unique[OF a] assms by simp
qed
lemma (in finite_measure) measure_countably_zero:
assumes "range c \<subseteq> sets M"
assumes "\<And> i. measure M (c i) = 0"
shows "measure M (\<Union>i :: nat. c i) = 0"
proof (rule antisym)
show "measure M (\<Union>i :: nat. c i) \<le> 0"
using finite_measure_subadditive_countably[OF assms(1)] by (simp add: assms(2))
qed simp
lemma (in finite_measure) measure_space_inter:
assumes events:"s \<in> sets M" "t \<in> sets M"
assumes "measure M t = measure M (space M)"
shows "measure M (s \<inter> t) = measure M s"
proof -
have "measure M ((space M - s) \<union> (space M - t)) = measure M (space M - s)"
using events assms finite_measure_compl[of "t"] by (auto intro!: measure_zero_union)
also have "(space M - s) \<union> (space M - t) = space M - (s \<inter> t)"
by blast
finally show "measure M (s \<inter> t) = measure M s"
using events by (auto intro!: measure_eq_compl[of "s \<inter> t" s])
qed
lemma (in finite_measure) measure_equiprobable_finite_unions:
assumes s: "finite s" "\<And>x. x \<in> s \<Longrightarrow> {x} \<in> sets M"
assumes "\<And> x y. \<lbrakk>x \<in> s; y \<in> s\<rbrakk> \<Longrightarrow> measure M {x} = measure M {y}"
shows "measure M s = real (card s) * measure M {SOME x. x \<in> s}"
proof cases
assume "s \<noteq> {}"
then have "\<exists> x. x \<in> s" by blast
from someI_ex[OF this] assms
have prob_some: "\<And> x. x \<in> s \<Longrightarrow> measure M {x} = measure M {SOME y. y \<in> s}" by blast
have "measure M s = (\<Sum> x \<in> s. measure M {x})"
using finite_measure_eq_sum_singleton[OF s] by simp
also have "\<dots> = (\<Sum> x \<in> s. measure M {SOME y. y \<in> s})" using prob_some by auto
also have "\<dots> = real (card s) * measure M {(SOME x. x \<in> s)}"
using sum_constant assms by simp
finally show ?thesis by simp
qed simp
lemma (in finite_measure) measure_real_sum_image_fn:
assumes "e \<in> sets M"
assumes "\<And> x. x \<in> s \<Longrightarrow> e \<inter> f x \<in> sets M"
assumes "finite s"
assumes disjoint: "\<And> x y. \<lbrakk>x \<in> s ; y \<in> s ; x \<noteq> y\<rbrakk> \<Longrightarrow> f x \<inter> f y = {}"
assumes upper: "space M \<subseteq> (\<Union>i \<in> s. f i)"
shows "measure M e = (\<Sum> x \<in> s. measure M (e \<inter> f x))"
proof -
have "e \<subseteq> (\<Union>i\<in>s. f i)"
using \<open>e \<in> sets M\<close> sets.sets_into_space upper by blast
then have e: "e = (\<Union>i \<in> s. e \<inter> f i)"
by auto
hence "measure M e = measure M (\<Union>i \<in> s. e \<inter> f i)" by simp
also have "\<dots> = (\<Sum> x \<in> s. measure M (e \<inter> f x))"
proof (rule finite_measure_finite_Union)
show "finite s" by fact
show "(\<lambda>i. e \<inter> f i)`s \<subseteq> sets M" using assms(2) by auto
show "disjoint_family_on (\<lambda>i. e \<inter> f i) s"
using disjoint by (auto simp: disjoint_family_on_def)
qed
finally show ?thesis .
qed
lemma (in finite_measure) measure_exclude:
assumes "A \<in> sets M" "B \<in> sets M"
assumes "measure M A = measure M (space M)" "A \<inter> B = {}"
shows "measure M B = 0"
using measure_space_inter[of B A] assms by (auto simp: ac_simps)
lemma (in finite_measure) finite_measure_distr:
assumes f: "f \<in> measurable M M'"
shows "finite_measure (distr M M' f)"
proof (rule finite_measureI)
have "f -` space M' \<inter> space M = space M" using f by (auto dest: measurable_space)
with f show "emeasure (distr M M' f) (space (distr M M' f)) \<noteq> \<infinity>" by (auto simp: emeasure_distr)
qed
lemma emeasure_gfp[consumes 1, case_names cont measurable]:
assumes sets[simp]: "\<And>s. sets (M s) = sets N"
assumes "\<And>s. finite_measure (M s)"
assumes cont: "inf_continuous F" "inf_continuous f"
assumes meas: "\<And>P. Measurable.pred N P \<Longrightarrow> Measurable.pred N (F P)"
assumes iter: "\<And>P s. Measurable.pred N P \<Longrightarrow> emeasure (M s) {x\<in>space N. F P x} = f (\<lambda>s. emeasure (M s) {x\<in>space N. P x}) s"
assumes bound: "\<And>P. f P \<le> f (\<lambda>s. emeasure (M s) (space (M s)))"
shows "emeasure (M s) {x\<in>space N. gfp F x} = gfp f s"
proof (subst gfp_transfer_bounded[where \<alpha>="\<lambda>F s. emeasure (M s) {x\<in>space N. F x}" and g=f and f=F and
P="Measurable.pred N", symmetric])
interpret finite_measure "M s" for s by fact
fix C assume "decseq C" "\<And>i. Measurable.pred N (C i)"
then show "(\<lambda>s. emeasure (M s) {x \<in> space N. (INF i. C i) x}) = (INF i. (\<lambda>s. emeasure (M s) {x \<in> space N. C i x}))"
unfolding INF_apply[abs_def]
by (subst INF_emeasure_decseq) (auto simp: antimono_def fun_eq_iff intro!: arg_cong2[where f=emeasure])
next
show "f x \<le> (\<lambda>s. emeasure (M s) {x \<in> space N. F top x})" for x
using bound[of x] sets_eq_imp_space_eq[OF sets] by (simp add: iter)
qed (auto simp add: iter le_fun_def INF_apply[abs_def] intro!: meas cont)
subsection\<^marker>\<open>tag unimportant\<close> \<open>Counting space\<close>
lemma strict_monoI_Suc:
assumes ord [simp]: "(\<And>n. f n < f (Suc n))" shows "strict_mono f"
unfolding strict_mono_def
proof safe
fix n m :: nat assume "n < m" then show "f n < f m"
by (induct m) (auto simp: less_Suc_eq intro: less_trans ord)
qed
lemma emeasure_count_space:
assumes "X \<subseteq> A" shows "emeasure (count_space A) X = (if finite X then of_nat (card X) else \<infinity>)"
(is "_ = ?M X")
unfolding count_space_def
proof (rule emeasure_measure_of_sigma)
show "X \<in> Pow A" using \<open>X \<subseteq> A\<close> by auto
show "sigma_algebra A (Pow A)" by (rule sigma_algebra_Pow)
show positive: "positive (Pow A) ?M"
by (auto simp: positive_def)
have additive: "additive (Pow A) ?M"
by (auto simp: card_Un_disjoint additive_def)
interpret ring_of_sets A "Pow A"
by (rule ring_of_setsI) auto
show "countably_additive (Pow A) ?M"
unfolding countably_additive_iff_continuous_from_below[OF positive additive]
proof safe
fix F :: "nat \<Rightarrow> 'a set" assume "incseq F"
show "(\<lambda>i. ?M (F i)) \<longlonglongrightarrow> ?M (\<Union>i. F i)"
proof cases
assume "\<exists>i. \<forall>j\<ge>i. F i = F j"
then guess i .. note i = this
{ fix j from i \<open>incseq F\<close> have "F j \<subseteq> F i"
by (cases "i \<le> j") (auto simp: incseq_def) }
then have eq: "(\<Union>i. F i) = F i"
by auto
with i show ?thesis
by (auto intro!: Lim_transform_eventually[OF tendsto_const] eventually_sequentiallyI[where c=i])
next
assume "\<not> (\<exists>i. \<forall>j\<ge>i. F i = F j)"
then obtain f where f: "\<And>i. i \<le> f i" "\<And>i. F i \<noteq> F (f i)" by metis
then have "\<And>i. F i \<subseteq> F (f i)" using \<open>incseq F\<close> by (auto simp: incseq_def)
with f have *: "\<And>i. F i \<subset> F (f i)" by auto
have "incseq (\<lambda>i. ?M (F i))"
using \<open>incseq F\<close> unfolding incseq_def by (auto simp: card_mono dest: finite_subset)
then have "(\<lambda>i. ?M (F i)) \<longlonglongrightarrow> (SUP n. ?M (F n))"
by (rule LIMSEQ_SUP)
moreover have "(SUP n. ?M (F n)) = top"
proof (rule ennreal_SUP_eq_top)
fix n :: nat show "\<exists>k::nat\<in>UNIV. of_nat n \<le> ?M (F k)"
proof (induct n)
case (Suc n)
then guess k .. note k = this
moreover have "finite (F k) \<Longrightarrow> finite (F (f k)) \<Longrightarrow> card (F k) < card (F (f k))"
using \<open>F k \<subset> F (f k)\<close> by (simp add: psubset_card_mono)
moreover have "finite (F (f k)) \<Longrightarrow> finite (F k)"
using \<open>k \<le> f k\<close> \<open>incseq F\<close> by (auto simp: incseq_def dest: finite_subset)
ultimately show ?case
by (auto intro!: exI[of _ "f k"] simp del: of_nat_Suc)
qed auto
qed
moreover
have "inj (\<lambda>n. F ((f ^^ n) 0))"
by (intro strict_mono_imp_inj_on strict_monoI_Suc) (simp add: *)
then have 1: "infinite (range (\<lambda>i. F ((f ^^ i) 0)))"
by (rule range_inj_infinite)
have "infinite (Pow (\<Union>i. F i))"
by (rule infinite_super[OF _ 1]) auto
then have "infinite (\<Union>i. F i)"
by auto
ultimately show ?thesis by (simp only:) simp
qed
qed
qed
lemma distr_bij_count_space:
assumes f: "bij_betw f A B"
shows "distr (count_space A) (count_space B) f = count_space B"
proof (rule measure_eqI)
have f': "f \<in> measurable (count_space A) (count_space B)"
using f unfolding Pi_def bij_betw_def by auto
fix X assume "X \<in> sets (distr (count_space A) (count_space B) f)"
then have X: "X \<in> sets (count_space B)" by auto
moreover from X have "f -` X \<inter> A = the_inv_into A f ` X"
using f by (auto simp: bij_betw_def subset_image_iff image_iff the_inv_into_f_f intro: the_inv_into_f_f[symmetric])
moreover have "inj_on (the_inv_into A f) B"
using X f by (auto simp: bij_betw_def inj_on_the_inv_into)
with X have "inj_on (the_inv_into A f) X"
by (auto intro: subset_inj_on)
ultimately show "emeasure (distr (count_space A) (count_space B) f) X = emeasure (count_space B) X"
using f unfolding emeasure_distr[OF f' X]
by (subst (1 2) emeasure_count_space) (auto simp: card_image dest: finite_imageD)
qed simp
lemma emeasure_count_space_finite[simp]:
"X \<subseteq> A \<Longrightarrow> finite X \<Longrightarrow> emeasure (count_space A) X = of_nat (card X)"
using emeasure_count_space[of X A] by simp
lemma emeasure_count_space_infinite[simp]:
"X \<subseteq> A \<Longrightarrow> infinite X \<Longrightarrow> emeasure (count_space A) X = \<infinity>"
using emeasure_count_space[of X A] by simp
lemma measure_count_space: "measure (count_space A) X = (if X \<subseteq> A then of_nat (card X) else 0)"
by (cases "finite X") (auto simp: measure_notin_sets ennreal_of_nat_eq_real_of_nat
measure_zero_top measure_eq_emeasure_eq_ennreal)
lemma emeasure_count_space_eq_0:
"emeasure (count_space A) X = 0 \<longleftrightarrow> (X \<subseteq> A \<longrightarrow> X = {})"
proof cases
assume X: "X \<subseteq> A"
then show ?thesis
proof (intro iffI impI)
assume "emeasure (count_space A) X = 0"
with X show "X = {}"
by (subst (asm) emeasure_count_space) (auto split: if_split_asm)
qed simp
qed (simp add: emeasure_notin_sets)
lemma null_sets_count_space: "null_sets (count_space A) = { {} }"
unfolding null_sets_def by (auto simp add: emeasure_count_space_eq_0)
lemma AE_count_space: "(AE x in count_space A. P x) \<longleftrightarrow> (\<forall>x\<in>A. P x)"
unfolding eventually_ae_filter by (auto simp add: null_sets_count_space)
lemma sigma_finite_measure_count_space_countable:
assumes A: "countable A"
shows "sigma_finite_measure (count_space A)"
proof qed (insert A, auto intro!: exI[of _ "(\<lambda>a. {a}) ` A"])
lemma sigma_finite_measure_count_space:
fixes A :: "'a::countable set" shows "sigma_finite_measure (count_space A)"
by (rule sigma_finite_measure_count_space_countable) auto
lemma finite_measure_count_space:
assumes [simp]: "finite A"
shows "finite_measure (count_space A)"
by rule simp
lemma sigma_finite_measure_count_space_finite:
assumes A: "finite A" shows "sigma_finite_measure (count_space A)"
proof -
interpret finite_measure "count_space A" using A by (rule finite_measure_count_space)
show "sigma_finite_measure (count_space A)" ..
qed
subsection\<^marker>\<open>tag unimportant\<close> \<open>Measure restricted to space\<close>
lemma emeasure_restrict_space:
assumes "\<Omega> \<inter> space M \<in> sets M" "A \<subseteq> \<Omega>"
shows "emeasure (restrict_space M \<Omega>) A = emeasure M A"
proof (cases "A \<in> sets M")
case True
show ?thesis
proof (rule emeasure_measure_of[OF restrict_space_def])
show "(\<inter>) \<Omega> ` sets M \<subseteq> Pow (\<Omega> \<inter> space M)" "A \<in> sets (restrict_space M \<Omega>)"
using \<open>A \<subseteq> \<Omega>\<close> \<open>A \<in> sets M\<close> sets.space_closed by (auto simp: sets_restrict_space)
show "positive (sets (restrict_space M \<Omega>)) (emeasure M)"
by (auto simp: positive_def)
show "countably_additive (sets (restrict_space M \<Omega>)) (emeasure M)"
proof (rule countably_additiveI)
fix A :: "nat \<Rightarrow> _" assume "range A \<subseteq> sets (restrict_space M \<Omega>)" "disjoint_family A"
with assms have "\<And>i. A i \<in> sets M" "\<And>i. A i \<subseteq> space M" "disjoint_family A"
by (fastforce simp: sets_restrict_space_iff[OF assms(1)] image_subset_iff
dest: sets.sets_into_space)+
then show "(\<Sum>i. emeasure M (A i)) = emeasure M (\<Union>i. A i)"
by (subst suminf_emeasure) (auto simp: disjoint_family_subset)
qed
qed
next
case False
with assms have "A \<notin> sets (restrict_space M \<Omega>)"
by (simp add: sets_restrict_space_iff)
with False show ?thesis
by (simp add: emeasure_notin_sets)
qed
lemma measure_restrict_space:
assumes "\<Omega> \<inter> space M \<in> sets M" "A \<subseteq> \<Omega>"
shows "measure (restrict_space M \<Omega>) A = measure M A"
using emeasure_restrict_space[OF assms] by (simp add: measure_def)
lemma AE_restrict_space_iff:
assumes "\<Omega> \<inter> space M \<in> sets M"
shows "(AE x in restrict_space M \<Omega>. P x) \<longleftrightarrow> (AE x in M. x \<in> \<Omega> \<longrightarrow> P x)"
proof -
have ex_cong: "\<And>P Q f. (\<And>x. P x \<Longrightarrow> Q x) \<Longrightarrow> (\<And>x. Q x \<Longrightarrow> P (f x)) \<Longrightarrow> (\<exists>x. P x) \<longleftrightarrow> (\<exists>x. Q x)"
by auto
{ fix X assume X: "X \<in> sets M" "emeasure M X = 0"
then have "emeasure M (\<Omega> \<inter> space M \<inter> X) \<le> emeasure M X"
by (intro emeasure_mono) auto
then have "emeasure M (\<Omega> \<inter> space M \<inter> X) = 0"
using X by (auto intro!: antisym) }
with assms show ?thesis
unfolding eventually_ae_filter
by (auto simp add: space_restrict_space null_sets_def sets_restrict_space_iff
emeasure_restrict_space cong: conj_cong
intro!: ex_cong[where f="\<lambda>X. (\<Omega> \<inter> space M) \<inter> X"])
qed
lemma restrict_restrict_space:
assumes "A \<inter> space M \<in> sets M" "B \<inter> space M \<in> sets M"
shows "restrict_space (restrict_space M A) B = restrict_space M (A \<inter> B)" (is "?l = ?r")
proof (rule measure_eqI[symmetric])
show "sets ?r = sets ?l"
unfolding sets_restrict_space image_comp by (intro image_cong) auto
next
fix X assume "X \<in> sets (restrict_space M (A \<inter> B))"
then obtain Y where "Y \<in> sets M" "X = Y \<inter> A \<inter> B"
by (auto simp: sets_restrict_space)
with assms sets.Int[OF assms] show "emeasure ?r X = emeasure ?l X"
by (subst (1 2) emeasure_restrict_space)
(auto simp: space_restrict_space sets_restrict_space_iff emeasure_restrict_space ac_simps)
qed
lemma restrict_count_space: "restrict_space (count_space B) A = count_space (A \<inter> B)"
proof (rule measure_eqI)
show "sets (restrict_space (count_space B) A) = sets (count_space (A \<inter> B))"
by (subst sets_restrict_space) auto
moreover fix X assume "X \<in> sets (restrict_space (count_space B) A)"
ultimately have "X \<subseteq> A \<inter> B" by auto
then show "emeasure (restrict_space (count_space B) A) X = emeasure (count_space (A \<inter> B)) X"
by (cases "finite X") (auto simp add: emeasure_restrict_space)
qed
lemma sigma_finite_measure_restrict_space:
assumes "sigma_finite_measure M"
and A: "A \<in> sets M"
shows "sigma_finite_measure (restrict_space M A)"
proof -
interpret sigma_finite_measure M by fact
from sigma_finite_countable obtain C
where C: "countable C" "C \<subseteq> sets M" "(\<Union>C) = space M" "\<forall>a\<in>C. emeasure M a \<noteq> \<infinity>"
by blast
let ?C = "(\<inter>) A ` C"
from C have "countable ?C" "?C \<subseteq> sets (restrict_space M A)" "(\<Union>?C) = space (restrict_space M A)"
by(auto simp add: sets_restrict_space space_restrict_space)
moreover {
fix a
assume "a \<in> ?C"
then obtain a' where "a = A \<inter> a'" "a' \<in> C" ..
then have "emeasure (restrict_space M A) a \<le> emeasure M a'"
using A C by(auto simp add: emeasure_restrict_space intro: emeasure_mono)
also have "\<dots> < \<infinity>" using C(4)[rule_format, of a'] \<open>a' \<in> C\<close> by (simp add: less_top)
finally have "emeasure (restrict_space M A) a \<noteq> \<infinity>" by simp }
ultimately show ?thesis
by unfold_locales (rule exI conjI|assumption|blast)+
qed
lemma finite_measure_restrict_space:
assumes "finite_measure M"
and A: "A \<in> sets M"
shows "finite_measure (restrict_space M A)"
proof -
interpret finite_measure M by fact
show ?thesis
by(rule finite_measureI)(simp add: emeasure_restrict_space A space_restrict_space)
qed
lemma restrict_distr:
assumes [measurable]: "f \<in> measurable M N"
assumes [simp]: "\<Omega> \<inter> space N \<in> sets N" and restrict: "f \<in> space M \<rightarrow> \<Omega>"
shows "restrict_space (distr M N f) \<Omega> = distr M (restrict_space N \<Omega>) f"
(is "?l = ?r")
proof (rule measure_eqI)
fix A assume "A \<in> sets (restrict_space (distr M N f) \<Omega>)"
with restrict show "emeasure ?l A = emeasure ?r A"
by (subst emeasure_distr)
(auto simp: sets_restrict_space_iff emeasure_restrict_space emeasure_distr
intro!: measurable_restrict_space2)
qed (simp add: sets_restrict_space)
lemma measure_eqI_restrict_generator:
assumes E: "Int_stable E" "E \<subseteq> Pow \<Omega>" "\<And>X. X \<in> E \<Longrightarrow> emeasure M X = emeasure N X"
assumes sets_eq: "sets M = sets N" and \<Omega>: "\<Omega> \<in> sets M"
assumes "sets (restrict_space M \<Omega>) = sigma_sets \<Omega> E"
assumes "sets (restrict_space N \<Omega>) = sigma_sets \<Omega> E"
assumes ae: "AE x in M. x \<in> \<Omega>" "AE x in N. x \<in> \<Omega>"
assumes A: "countable A" "A \<noteq> {}" "A \<subseteq> E" "\<Union>A = \<Omega>" "\<And>a. a \<in> A \<Longrightarrow> emeasure M a \<noteq> \<infinity>"
shows "M = N"
proof (rule measure_eqI)
fix X assume X: "X \<in> sets M"
then have "emeasure M X = emeasure (restrict_space M \<Omega>) (X \<inter> \<Omega>)"
using ae \<Omega> by (auto simp add: emeasure_restrict_space intro!: emeasure_eq_AE)
also have "restrict_space M \<Omega> = restrict_space N \<Omega>"
proof (rule measure_eqI_generator_eq)
fix X assume "X \<in> E"
then show "emeasure (restrict_space M \<Omega>) X = emeasure (restrict_space N \<Omega>) X"
using E \<Omega> by (subst (1 2) emeasure_restrict_space) (auto simp: sets_eq sets_eq[THEN sets_eq_imp_space_eq])
next
show "range (from_nat_into A) \<subseteq> E" "(\<Union>i. from_nat_into A i) = \<Omega>"
using A by (auto cong del: SUP_cong_simp)
next
fix i
have "emeasure (restrict_space M \<Omega>) (from_nat_into A i) = emeasure M (from_nat_into A i)"
using A \<Omega> by (subst emeasure_restrict_space)
(auto simp: sets_eq sets_eq[THEN sets_eq_imp_space_eq] intro: from_nat_into)
with A show "emeasure (restrict_space M \<Omega>) (from_nat_into A i) \<noteq> \<infinity>"
by (auto intro: from_nat_into)
qed fact+
also have "emeasure (restrict_space N \<Omega>) (X \<inter> \<Omega>) = emeasure N X"
using X ae \<Omega> by (auto simp add: emeasure_restrict_space sets_eq intro!: emeasure_eq_AE)
finally show "emeasure M X = emeasure N X" .
qed fact
subsection\<^marker>\<open>tag unimportant\<close> \<open>Null measure\<close>
definition null_measure :: "'a measure \<Rightarrow> 'a measure" where
"null_measure M = sigma (space M) (sets M)"
lemma space_null_measure[simp]: "space (null_measure M) = space M"
by (simp add: null_measure_def)
lemma sets_null_measure[simp, measurable_cong]: "sets (null_measure M) = sets M"
by (simp add: null_measure_def)
lemma emeasure_null_measure[simp]: "emeasure (null_measure M) X = 0"
by (cases "X \<in> sets M", rule emeasure_measure_of)
(auto simp: positive_def countably_additive_def emeasure_notin_sets null_measure_def
dest: sets.sets_into_space)
lemma measure_null_measure[simp]: "measure (null_measure M) X = 0"
by (intro measure_eq_emeasure_eq_ennreal) auto
lemma null_measure_idem [simp]: "null_measure (null_measure M) = null_measure M"
by(rule measure_eqI) simp_all
subsection \<open>Scaling a measure\<close>
definition\<^marker>\<open>tag important\<close> scale_measure :: "ennreal \<Rightarrow> 'a measure \<Rightarrow> 'a measure" where
"scale_measure r M = measure_of (space M) (sets M) (\<lambda>A. r * emeasure M A)"
lemma space_scale_measure: "space (scale_measure r M) = space M"
by (simp add: scale_measure_def)
lemma sets_scale_measure [simp, measurable_cong]: "sets (scale_measure r M) = sets M"
by (simp add: scale_measure_def)
lemma emeasure_scale_measure [simp]:
"emeasure (scale_measure r M) A = r * emeasure M A"
(is "_ = ?\<mu> A")
proof(cases "A \<in> sets M")
case True
show ?thesis unfolding scale_measure_def
proof(rule emeasure_measure_of_sigma)
show "sigma_algebra (space M) (sets M)" ..
show "positive (sets M) ?\<mu>" by (simp add: positive_def)
show "countably_additive (sets M) ?\<mu>"
proof (rule countably_additiveI)
fix A :: "nat \<Rightarrow> _" assume *: "range A \<subseteq> sets M" "disjoint_family A"
have "(\<Sum>i. ?\<mu> (A i)) = r * (\<Sum>i. emeasure M (A i))"
by simp
also have "\<dots> = ?\<mu> (\<Union>i. A i)" using * by(simp add: suminf_emeasure)
finally show "(\<Sum>i. ?\<mu> (A i)) = ?\<mu> (\<Union>i. A i)" .
qed
qed(fact True)
qed(simp add: emeasure_notin_sets)
lemma scale_measure_1 [simp]: "scale_measure 1 M = M"
by(rule measure_eqI) simp_all
lemma scale_measure_0[simp]: "scale_measure 0 M = null_measure M"
by(rule measure_eqI) simp_all
lemma measure_scale_measure [simp]: "0 \<le> r \<Longrightarrow> measure (scale_measure r M) A = r * measure M A"
using emeasure_scale_measure[of r M A]
emeasure_eq_ennreal_measure[of M A]
measure_eq_emeasure_eq_ennreal[of _ "scale_measure r M" A]
by (cases "emeasure (scale_measure r M) A = top")
(auto simp del: emeasure_scale_measure
simp: ennreal_top_eq_mult_iff ennreal_mult_eq_top_iff measure_zero_top ennreal_mult[symmetric])
lemma scale_scale_measure [simp]:
"scale_measure r (scale_measure r' M) = scale_measure (r * r') M"
by (rule measure_eqI) (simp_all add: max_def mult.assoc)
lemma scale_null_measure [simp]: "scale_measure r (null_measure M) = null_measure M"
by (rule measure_eqI) simp_all
subsection \<open>Complete lattice structure on measures\<close>
lemma (in finite_measure) finite_measure_Diff':
"A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> measure M (A - B) = measure M A - measure M (A \<inter> B)"
using finite_measure_Diff[of A "A \<inter> B"] by (auto simp: Diff_Int)
lemma (in finite_measure) finite_measure_Union':
"A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> measure M (A \<union> B) = measure M A + measure M (B - A)"
using finite_measure_Union[of A "B - A"] by auto
lemma finite_unsigned_Hahn_decomposition:
assumes "finite_measure M" "finite_measure N" and [simp]: "sets N = sets M"
shows "\<exists>Y\<in>sets M. (\<forall>X\<in>sets M. X \<subseteq> Y \<longrightarrow> N X \<le> M X) \<and> (\<forall>X\<in>sets M. X \<inter> Y = {} \<longrightarrow> M X \<le> N X)"
proof -
interpret M: finite_measure M by fact
interpret N: finite_measure N by fact
define d where "d X = measure M X - measure N X" for X
have [intro]: "bdd_above (d`sets M)"
using sets.sets_into_space[of _ M]
by (intro bdd_aboveI[where M="measure M (space M)"])
(auto simp: d_def field_simps subset_eq intro!: add_increasing M.finite_measure_mono)
define \<gamma> where "\<gamma> = (SUP X\<in>sets M. d X)"
have le_\<gamma>[intro]: "X \<in> sets M \<Longrightarrow> d X \<le> \<gamma>" for X
by (auto simp: \<gamma>_def intro!: cSUP_upper)
have "\<exists>f. \<forall>n. f n \<in> sets M \<and> d (f n) > \<gamma> - 1 / 2^n"
proof (intro choice_iff[THEN iffD1] allI)
fix n
have "\<exists>X\<in>sets M. \<gamma> - 1 / 2^n < d X"
unfolding \<gamma>_def by (intro less_cSUP_iff[THEN iffD1]) auto
then show "\<exists>y. y \<in> sets M \<and> \<gamma> - 1 / 2 ^ n < d y"
by auto
qed
then obtain E where [measurable]: "E n \<in> sets M" and E: "d (E n) > \<gamma> - 1 / 2^n" for n
by auto
define F where "F m n = (if m \<le> n then \<Inter>i\<in>{m..n}. E i else space M)" for m n
have [measurable]: "m \<le> n \<Longrightarrow> F m n \<in> sets M" for m n
by (auto simp: F_def)
have 1: "\<gamma> - 2 / 2 ^ m + 1 / 2 ^ n \<le> d (F m n)" if "m \<le> n" for m n
using that
proof (induct rule: dec_induct)
case base with E[of m] show ?case
by (simp add: F_def field_simps)
next
case (step i)
have F_Suc: "F m (Suc i) = F m i \<inter> E (Suc i)"
using \<open>m \<le> i\<close> by (auto simp: F_def le_Suc_eq)
have "\<gamma> + (\<gamma> - 2 / 2^m + 1 / 2 ^ Suc i) \<le> (\<gamma> - 1 / 2^Suc i) + (\<gamma> - 2 / 2^m + 1 / 2^i)"
by (simp add: field_simps)
also have "\<dots> \<le> d (E (Suc i)) + d (F m i)"
using E[of "Suc i"] by (intro add_mono step) auto
also have "\<dots> = d (E (Suc i)) + d (F m i - E (Suc i)) + d (F m (Suc i))"
using \<open>m \<le> i\<close> by (simp add: d_def field_simps F_Suc M.finite_measure_Diff' N.finite_measure_Diff')
also have "\<dots> = d (E (Suc i) \<union> F m i) + d (F m (Suc i))"
using \<open>m \<le> i\<close> by (simp add: d_def field_simps M.finite_measure_Union' N.finite_measure_Union')
also have "\<dots> \<le> \<gamma> + d (F m (Suc i))"
using \<open>m \<le> i\<close> by auto
finally show ?case
by auto
qed
define F' where "F' m = (\<Inter>i\<in>{m..}. E i)" for m
have F'_eq: "F' m = (\<Inter>i. F m (i + m))" for m
by (fastforce simp: le_iff_add[of m] F'_def F_def)
have [measurable]: "F' m \<in> sets M" for m
by (auto simp: F'_def)
have \<gamma>_le: "\<gamma> - 0 \<le> d (\<Union>m. F' m)"
proof (rule LIMSEQ_le)
show "(\<lambda>n. \<gamma> - 2 / 2 ^ n) \<longlonglongrightarrow> \<gamma> - 0"
by (intro tendsto_intros LIMSEQ_divide_realpow_zero) auto
have "incseq F'"
by (auto simp: incseq_def F'_def)
then show "(\<lambda>m. d (F' m)) \<longlonglongrightarrow> d (\<Union>m. F' m)"
unfolding d_def
by (intro tendsto_diff M.finite_Lim_measure_incseq N.finite_Lim_measure_incseq) auto
have "\<gamma> - 2 / 2 ^ m + 0 \<le> d (F' m)" for m
proof (rule LIMSEQ_le)
have *: "decseq (\<lambda>n. F m (n + m))"
by (auto simp: decseq_def F_def)
show "(\<lambda>n. d (F m n)) \<longlonglongrightarrow> d (F' m)"
unfolding d_def F'_eq
by (rule LIMSEQ_offset[where k=m])
(auto intro!: tendsto_diff M.finite_Lim_measure_decseq N.finite_Lim_measure_decseq *)
show "(\<lambda>n. \<gamma> - 2 / 2 ^ m + 1 / 2 ^ n) \<longlonglongrightarrow> \<gamma> - 2 / 2 ^ m + 0"
by (intro tendsto_add LIMSEQ_divide_realpow_zero tendsto_const) auto
show "\<exists>N. \<forall>n\<ge>N. \<gamma> - 2 / 2 ^ m + 1 / 2 ^ n \<le> d (F m n)"
using 1[of m] by (intro exI[of _ m]) auto
qed
then show "\<exists>N. \<forall>n\<ge>N. \<gamma> - 2 / 2 ^ n \<le> d (F' n)"
by auto
qed
show ?thesis
proof (safe intro!: bexI[of _ "\<Union>m. F' m"])
fix X assume [measurable]: "X \<in> sets M" and X: "X \<subseteq> (\<Union>m. F' m)"
have "d (\<Union>m. F' m) - d X = d ((\<Union>m. F' m) - X)"
using X by (auto simp: d_def M.finite_measure_Diff N.finite_measure_Diff)
also have "\<dots> \<le> \<gamma>"
by auto
finally have "0 \<le> d X"
using \<gamma>_le by auto
then show "emeasure N X \<le> emeasure M X"
by (auto simp: d_def M.emeasure_eq_measure N.emeasure_eq_measure)
next
fix X assume [measurable]: "X \<in> sets M" and X: "X \<inter> (\<Union>m. F' m) = {}"
then have "d (\<Union>m. F' m) + d X = d (X \<union> (\<Union>m. F' m))"
by (auto simp: d_def M.finite_measure_Union N.finite_measure_Union)
also have "\<dots> \<le> \<gamma>"
by auto
finally have "d X \<le> 0"
using \<gamma>_le by auto
then show "emeasure M X \<le> emeasure N X"
by (auto simp: d_def M.emeasure_eq_measure N.emeasure_eq_measure)
qed auto
qed
proposition unsigned_Hahn_decomposition:
assumes [simp]: "sets N = sets M" and [measurable]: "A \<in> sets M"
and [simp]: "emeasure M A \<noteq> top" "emeasure N A \<noteq> top"
shows "\<exists>Y\<in>sets M. Y \<subseteq> A \<and> (\<forall>X\<in>sets M. X \<subseteq> Y \<longrightarrow> N X \<le> M X) \<and> (\<forall>X\<in>sets M. X \<subseteq> A \<longrightarrow> X \<inter> Y = {} \<longrightarrow> M X \<le> N X)"
proof -
have "\<exists>Y\<in>sets (restrict_space M A).
(\<forall>X\<in>sets (restrict_space M A). X \<subseteq> Y \<longrightarrow> (restrict_space N A) X \<le> (restrict_space M A) X) \<and>
(\<forall>X\<in>sets (restrict_space M A). X \<inter> Y = {} \<longrightarrow> (restrict_space M A) X \<le> (restrict_space N A) X)"
proof (rule finite_unsigned_Hahn_decomposition)
show "finite_measure (restrict_space M A)" "finite_measure (restrict_space N A)"
by (auto simp: space_restrict_space emeasure_restrict_space less_top intro!: finite_measureI)
qed (simp add: sets_restrict_space)
then guess Y ..
then show ?thesis
apply (intro bexI[of _ Y] conjI ballI conjI)
apply (simp_all add: sets_restrict_space emeasure_restrict_space)
apply safe
subgoal for X Z
by (erule ballE[of _ _ X]) (auto simp add: Int_absorb1)
subgoal for X Z
by (erule ballE[of _ _ X]) (auto simp add: Int_absorb1 ac_simps)
apply auto
done
qed
text\<^marker>\<open>tag important\<close> \<open>
Define a lexicographical order on \<^type>\<open>measure\<close>, in the order space, sets and measure. The parts
of the lexicographical order are point-wise ordered.
\<close>
instantiation measure :: (type) order_bot
begin
inductive less_eq_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> bool" where
"space M \<subset> space N \<Longrightarrow> less_eq_measure M N"
| "space M = space N \<Longrightarrow> sets M \<subset> sets N \<Longrightarrow> less_eq_measure M N"
| "space M = space N \<Longrightarrow> sets M = sets N \<Longrightarrow> emeasure M \<le> emeasure N \<Longrightarrow> less_eq_measure M N"
lemma le_measure_iff:
"M \<le> N \<longleftrightarrow> (if space M = space N then
if sets M = sets N then emeasure M \<le> emeasure N else sets M \<subseteq> sets N else space M \<subseteq> space N)"
by (auto elim: less_eq_measure.cases intro: less_eq_measure.intros)
definition\<^marker>\<open>tag important\<close> less_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> bool" where
"less_measure M N \<longleftrightarrow> (M \<le> N \<and> \<not> N \<le> M)"
definition\<^marker>\<open>tag important\<close> bot_measure :: "'a measure" where
"bot_measure = sigma {} {}"
lemma
shows space_bot[simp]: "space bot = {}"
and sets_bot[simp]: "sets bot = {{}}"
and emeasure_bot[simp]: "emeasure bot X = 0"
by (auto simp: bot_measure_def sigma_sets_empty_eq emeasure_sigma)
instance
proof standard
show "bot \<le> a" for a :: "'a measure"
by (simp add: le_measure_iff bot_measure_def sigma_sets_empty_eq emeasure_sigma le_fun_def)
qed (auto simp: le_measure_iff less_measure_def split: if_split_asm intro: measure_eqI)
end
proposition le_measure: "sets M = sets N \<Longrightarrow> M \<le> N \<longleftrightarrow> (\<forall>A\<in>sets M. emeasure M A \<le> emeasure N A)"
apply -
apply (auto simp: le_measure_iff le_fun_def dest: sets_eq_imp_space_eq)
subgoal for X
by (cases "X \<in> sets M") (auto simp: emeasure_notin_sets)
done
definition\<^marker>\<open>tag important\<close> sup_measure' :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure" where
"sup_measure' A B =
measure_of (space A) (sets A)
(\<lambda>X. SUP Y\<in>sets A. emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y))"
lemma assumes [simp]: "sets B = sets A"
shows space_sup_measure'[simp]: "space (sup_measure' A B) = space A"
and sets_sup_measure'[simp]: "sets (sup_measure' A B) = sets A"
using sets_eq_imp_space_eq[OF assms] by (simp_all add: sup_measure'_def)
lemma emeasure_sup_measure':
assumes sets_eq[simp]: "sets B = sets A" and [simp, intro]: "X \<in> sets A"
shows "emeasure (sup_measure' A B) X = (SUP Y\<in>sets A. emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y))"
(is "_ = ?S X")
proof -
note sets_eq_imp_space_eq[OF sets_eq, simp]
show ?thesis
using sup_measure'_def
proof (rule emeasure_measure_of)
let ?d = "\<lambda>X Y. emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y)"
show "countably_additive (sets (sup_measure' A B)) (\<lambda>X. SUP Y \<in> sets A. emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y))"
proof (rule countably_additiveI, goal_cases)
case (1 X)
then have [measurable]: "\<And>i. X i \<in> sets A" and "disjoint_family X"
by auto
have disjoint: "disjoint_family (\<lambda>i. X i \<inter> Y)" "disjoint_family (\<lambda>i. X i - Y)" for Y
by (auto intro: disjoint_family_on_bisimulation [OF \<open>disjoint_family X\<close>, simplified])
have "(\<Sum>i. ?S (X i)) = (SUP Y\<in>sets A. \<Sum>i. ?d (X i) Y)"
proof (rule ennreal_suminf_SUP_eq_directed)
fix J :: "nat set" and a b assume "finite J" and [measurable]: "a \<in> sets A" "b \<in> sets A"
have "\<exists>c\<in>sets A. c \<subseteq> X i \<and> (\<forall>a\<in>sets A. ?d (X i) a \<le> ?d (X i) c)" for i
proof cases
assume "emeasure A (X i) = top \<or> emeasure B (X i) = top"
then show ?thesis
proof
assume "emeasure A (X i) = top" then show ?thesis
by (intro bexI[of _ "X i"]) auto
next
assume "emeasure B (X i) = top" then show ?thesis
by (intro bexI[of _ "{}"]) auto
qed
next
assume finite: "\<not> (emeasure A (X i) = top \<or> emeasure B (X i) = top)"
then have "\<exists>Y\<in>sets A. Y \<subseteq> X i \<and> (\<forall>C\<in>sets A. C \<subseteq> Y \<longrightarrow> B C \<le> A C) \<and> (\<forall>C\<in>sets A. C \<subseteq> X i \<longrightarrow> C \<inter> Y = {} \<longrightarrow> A C \<le> B C)"
using unsigned_Hahn_decomposition[of B A "X i"] by simp
then obtain Y where [measurable]: "Y \<in> sets A" and [simp]: "Y \<subseteq> X i"
and B_le_A: "\<And>C. C \<in> sets A \<Longrightarrow> C \<subseteq> Y \<Longrightarrow> B C \<le> A C"
and A_le_B: "\<And>C. C \<in> sets A \<Longrightarrow> C \<subseteq> X i \<Longrightarrow> C \<inter> Y = {} \<Longrightarrow> A C \<le> B C"
by auto
show ?thesis
proof (intro bexI[of _ Y] ballI conjI)
fix a assume [measurable]: "a \<in> sets A"
have *: "(X i \<inter> a \<inter> Y \<union> (X i \<inter> a - Y)) = X i \<inter> a" "(X i - a) \<inter> Y \<union> (X i - a - Y) = X i \<inter> - a"
for a Y by auto
then have "?d (X i) a =
(A (X i \<inter> a \<inter> Y) + A (X i \<inter> a \<inter> - Y)) + (B (X i \<inter> - a \<inter> Y) + B (X i \<inter> - a \<inter> - Y))"
by (subst (1 2) plus_emeasure) (auto simp: Diff_eq[symmetric])
also have "\<dots> \<le> (A (X i \<inter> a \<inter> Y) + B (X i \<inter> a \<inter> - Y)) + (A (X i \<inter> - a \<inter> Y) + B (X i \<inter> - a \<inter> - Y))"
by (intro add_mono order_refl B_le_A A_le_B) (auto simp: Diff_eq[symmetric])
also have "\<dots> \<le> (A (X i \<inter> Y \<inter> a) + A (X i \<inter> Y \<inter> - a)) + (B (X i \<inter> - Y \<inter> a) + B (X i \<inter> - Y \<inter> - a))"
by (simp add: ac_simps)
also have "\<dots> \<le> A (X i \<inter> Y) + B (X i \<inter> - Y)"
by (subst (1 2) plus_emeasure) (auto simp: Diff_eq[symmetric] *)
finally show "?d (X i) a \<le> ?d (X i) Y" .
qed auto
qed
then obtain C where [measurable]: "C i \<in> sets A" and "C i \<subseteq> X i"
and C: "\<And>a. a \<in> sets A \<Longrightarrow> ?d (X i) a \<le> ?d (X i) (C i)" for i
by metis
have *: "X i \<inter> (\<Union>i. C i) = X i \<inter> C i" for i
proof safe
fix x j assume "x \<in> X i" "x \<in> C j"
moreover have "i = j \<or> X i \<inter> X j = {}"
using \<open>disjoint_family X\<close> by (auto simp: disjoint_family_on_def)
ultimately show "x \<in> C i"
using \<open>C i \<subseteq> X i\<close> \<open>C j \<subseteq> X j\<close> by auto
qed auto
have **: "X i \<inter> - (\<Union>i. C i) = X i \<inter> - C i" for i
proof safe
fix x j assume "x \<in> X i" "x \<notin> C i" "x \<in> C j"
moreover have "i = j \<or> X i \<inter> X j = {}"
using \<open>disjoint_family X\<close> by (auto simp: disjoint_family_on_def)
ultimately show False
using \<open>C i \<subseteq> X i\<close> \<open>C j \<subseteq> X j\<close> by auto
qed auto
show "\<exists>c\<in>sets A. \<forall>i\<in>J. ?d (X i) a \<le> ?d (X i) c \<and> ?d (X i) b \<le> ?d (X i) c"
apply (intro bexI[of _ "\<Union>i. C i"])
unfolding * **
apply (intro C ballI conjI)
apply auto
done
qed
also have "\<dots> = ?S (\<Union>i. X i)"
apply (simp only: UN_extend_simps(4))
apply (rule arg_cong [of _ _ Sup])
apply (rule image_cong)
apply (fact refl)
using disjoint
apply (auto simp add: suminf_add [symmetric] Diff_eq [symmetric] image_subset_iff suminf_emeasure simp del: UN_simps)
done
finally show "(\<Sum>i. ?S (X i)) = ?S (\<Union>i. X i)" .
qed
qed (auto dest: sets.sets_into_space simp: positive_def intro!: SUP_const)
qed
lemma le_emeasure_sup_measure'1:
assumes "sets B = sets A" "X \<in> sets A" shows "emeasure A X \<le> emeasure (sup_measure' A B) X"
by (subst emeasure_sup_measure'[OF assms]) (auto intro!: SUP_upper2[of "X"] assms)
lemma le_emeasure_sup_measure'2:
assumes "sets B = sets A" "X \<in> sets A" shows "emeasure B X \<le> emeasure (sup_measure' A B) X"
by (subst emeasure_sup_measure'[OF assms]) (auto intro!: SUP_upper2[of "{}"] assms)
lemma emeasure_sup_measure'_le2:
assumes [simp]: "sets B = sets C" "sets A = sets C" and [measurable]: "X \<in> sets C"
assumes A: "\<And>Y. Y \<subseteq> X \<Longrightarrow> Y \<in> sets A \<Longrightarrow> emeasure A Y \<le> emeasure C Y"
assumes B: "\<And>Y. Y \<subseteq> X \<Longrightarrow> Y \<in> sets A \<Longrightarrow> emeasure B Y \<le> emeasure C Y"
shows "emeasure (sup_measure' A B) X \<le> emeasure C X"
proof (subst emeasure_sup_measure')
show "(SUP Y\<in>sets A. emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y)) \<le> emeasure C X"
unfolding \<open>sets A = sets C\<close>
proof (intro SUP_least)
fix Y assume [measurable]: "Y \<in> sets C"
have [simp]: "X \<inter> Y \<union> (X - Y) = X"
by auto
have "emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y) \<le> emeasure C (X \<inter> Y) + emeasure C (X \<inter> - Y)"
by (intro add_mono A B) (auto simp: Diff_eq[symmetric])
also have "\<dots> = emeasure C X"
by (subst plus_emeasure) (auto simp: Diff_eq[symmetric])
finally show "emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y) \<le> emeasure C X" .
qed
qed simp_all
definition\<^marker>\<open>tag important\<close> sup_lexord :: "'a \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'b::order) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" where
"sup_lexord A B k s c =
(if k A = k B then c else
if \<not> k A \<le> k B \<and> \<not> k B \<le> k A then s else
if k B \<le> k A then A else B)"
lemma sup_lexord:
"(k A < k B \<Longrightarrow> P B) \<Longrightarrow> (k B < k A \<Longrightarrow> P A) \<Longrightarrow> (k A = k B \<Longrightarrow> P c) \<Longrightarrow>
(\<not> k B \<le> k A \<Longrightarrow> \<not> k A \<le> k B \<Longrightarrow> P s) \<Longrightarrow> P (sup_lexord A B k s c)"
by (auto simp: sup_lexord_def)
lemmas le_sup_lexord = sup_lexord[where P="\<lambda>a. c \<le> a" for c]
lemma sup_lexord1: "k A = k B \<Longrightarrow> sup_lexord A B k s c = c"
by (simp add: sup_lexord_def)
lemma sup_lexord_commute: "sup_lexord A B k s c = sup_lexord B A k s c"
by (auto simp: sup_lexord_def)
lemma sigma_sets_le_sets_iff: "(sigma_sets (space x) \<A> \<subseteq> sets x) = (\<A> \<subseteq> sets x)"
using sets.sigma_sets_subset[of \<A> x] by auto
lemma sigma_le_iff: "\<A> \<subseteq> Pow \<Omega> \<Longrightarrow> sigma \<Omega> \<A> \<le> x \<longleftrightarrow> (\<Omega> \<subseteq> space x \<and> (space x = \<Omega> \<longrightarrow> \<A> \<subseteq> sets x))"
by (cases "\<Omega> = space x")
(simp_all add: eq_commute[of _ "sets x"] le_measure_iff emeasure_sigma le_fun_def
sigma_sets_superset_generator sigma_sets_le_sets_iff)
instantiation measure :: (type) semilattice_sup
begin
definition\<^marker>\<open>tag important\<close> sup_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure" where
"sup_measure A B =
sup_lexord A B space (sigma (space A \<union> space B) {})
(sup_lexord A B sets (sigma (space A) (sets A \<union> sets B)) (sup_measure' A B))"
instance
proof
fix x y z :: "'a measure"
show "x \<le> sup x y"
unfolding sup_measure_def
proof (intro le_sup_lexord)
assume "space x = space y"
then have *: "sets x \<union> sets y \<subseteq> Pow (space x)"
using sets.space_closed by auto
assume "\<not> sets y \<subseteq> sets x" "\<not> sets x \<subseteq> sets y"
then have "sets x \<subset> sets x \<union> sets y"
by auto
also have "\<dots> \<le> sigma (space x) (sets x \<union> sets y)"
by (subst sets_measure_of[OF *]) (rule sigma_sets_superset_generator)
finally show "x \<le> sigma (space x) (sets x \<union> sets y)"
by (simp add: space_measure_of[OF *] less_eq_measure.intros(2))
next
assume "\<not> space y \<subseteq> space x" "\<not> space x \<subseteq> space y"
then show "x \<le> sigma (space x \<union> space y) {}"
by (intro less_eq_measure.intros) auto
next
assume "sets x = sets y" then show "x \<le> sup_measure' x y"
by (simp add: le_measure le_emeasure_sup_measure'1)
qed (auto intro: less_eq_measure.intros)
show "y \<le> sup x y"
unfolding sup_measure_def
proof (intro le_sup_lexord)
assume **: "space x = space y"
then have *: "sets x \<union> sets y \<subseteq> Pow (space y)"
using sets.space_closed by auto
assume "\<not> sets y \<subseteq> sets x" "\<not> sets x \<subseteq> sets y"
then have "sets y \<subset> sets x \<union> sets y"
by auto
also have "\<dots> \<le> sigma (space y) (sets x \<union> sets y)"
by (subst sets_measure_of[OF *]) (rule sigma_sets_superset_generator)
finally show "y \<le> sigma (space x) (sets x \<union> sets y)"
by (simp add: ** space_measure_of[OF *] less_eq_measure.intros(2))
next
assume "\<not> space y \<subseteq> space x" "\<not> space x \<subseteq> space y"
then show "y \<le> sigma (space x \<union> space y) {}"
by (intro less_eq_measure.intros) auto
next
assume "sets x = sets y" then show "y \<le> sup_measure' x y"
by (simp add: le_measure le_emeasure_sup_measure'2)
qed (auto intro: less_eq_measure.intros)
show "x \<le> y \<Longrightarrow> z \<le> y \<Longrightarrow> sup x z \<le> y"
unfolding sup_measure_def
proof (intro sup_lexord[where P="\<lambda>x. x \<le> y"])
assume "x \<le> y" "z \<le> y" and [simp]: "space x = space z" "sets x = sets z"
from \<open>x \<le> y\<close> show "sup_measure' x z \<le> y"
proof cases
case 1 then show ?thesis
by (intro less_eq_measure.intros(1)) simp
next
case 2 then show ?thesis
by (intro less_eq_measure.intros(2)) simp_all
next
case 3 with \<open>z \<le> y\<close> \<open>x \<le> y\<close> show ?thesis
by (auto simp add: le_measure intro!: emeasure_sup_measure'_le2)
qed
next
assume **: "x \<le> y" "z \<le> y" "space x = space z" "\<not> sets z \<subseteq> sets x" "\<not> sets x \<subseteq> sets z"
then have *: "sets x \<union> sets z \<subseteq> Pow (space x)"
using sets.space_closed by auto
show "sigma (space x) (sets x \<union> sets z) \<le> y"
unfolding sigma_le_iff[OF *] using ** by (auto simp: le_measure_iff split: if_split_asm)
next
assume "x \<le> y" "z \<le> y" "\<not> space z \<subseteq> space x" "\<not> space x \<subseteq> space z"
then have "space x \<subseteq> space y" "space z \<subseteq> space y"
by (auto simp: le_measure_iff split: if_split_asm)
then show "sigma (space x \<union> space z) {} \<le> y"
by (simp add: sigma_le_iff)
qed
qed
end
lemma space_empty_eq_bot: "space a = {} \<longleftrightarrow> a = bot"
using space_empty[of a] by (auto intro!: measure_eqI)
lemma sets_eq_iff_bounded: "A \<le> B \<Longrightarrow> B \<le> C \<Longrightarrow> sets A = sets C \<Longrightarrow> sets B = sets A"
by (auto dest: sets_eq_imp_space_eq simp add: le_measure_iff split: if_split_asm)
lemma sets_sup: "sets A = sets M \<Longrightarrow> sets B = sets M \<Longrightarrow> sets (sup A B) = sets M"
by (auto simp add: sup_measure_def sup_lexord_def dest: sets_eq_imp_space_eq)
lemma le_measureD1: "A \<le> B \<Longrightarrow> space A \<le> space B"
by (auto simp: le_measure_iff split: if_split_asm)
lemma le_measureD2: "A \<le> B \<Longrightarrow> space A = space B \<Longrightarrow> sets A \<le> sets B"
by (auto simp: le_measure_iff split: if_split_asm)
lemma le_measureD3: "A \<le> B \<Longrightarrow> sets A = sets B \<Longrightarrow> emeasure A X \<le> emeasure B X"
by (auto simp: le_measure_iff le_fun_def dest: sets_eq_imp_space_eq split: if_split_asm)
lemma UN_space_closed: "\<Union>(sets ` S) \<subseteq> Pow (\<Union>(space ` S))"
using sets.space_closed by auto
definition\<^marker>\<open>tag important\<close>
Sup_lexord :: "('a \<Rightarrow> 'b::complete_lattice) \<Rightarrow> ('a set \<Rightarrow> 'a) \<Rightarrow> ('a set \<Rightarrow> 'a) \<Rightarrow> 'a set \<Rightarrow> 'a"
where
"Sup_lexord k c s A =
(let U = (SUP a\<in>A. k a)
in if \<exists>a\<in>A. k a = U then c {a\<in>A. k a = U} else s A)"
lemma Sup_lexord:
"(\<And>a S. a \<in> A \<Longrightarrow> k a = (SUP a\<in>A. k a) \<Longrightarrow> S = {a'\<in>A. k a' = k a} \<Longrightarrow> P (c S)) \<Longrightarrow> ((\<And>a. a \<in> A \<Longrightarrow> k a \<noteq> (SUP a\<in>A. k a)) \<Longrightarrow> P (s A)) \<Longrightarrow>
P (Sup_lexord k c s A)"
by (auto simp: Sup_lexord_def Let_def)
lemma Sup_lexord1:
assumes A: "A \<noteq> {}" "(\<And>a. a \<in> A \<Longrightarrow> k a = (\<Union>a\<in>A. k a))" "P (c A)"
shows "P (Sup_lexord k c s A)"
unfolding Sup_lexord_def Let_def
proof (clarsimp, safe)
show "\<forall>a\<in>A. k a \<noteq> (\<Union>x\<in>A. k x) \<Longrightarrow> P (s A)"
by (metis assms(1,2) ex_in_conv)
next
fix a assume "a \<in> A" "k a = (\<Union>x\<in>A. k x)"
then have "{a \<in> A. k a = (\<Union>x\<in>A. k x)} = {a \<in> A. k a = k a}"
by (metis A(2)[symmetric])
then show "P (c {a \<in> A. k a = (\<Union>x\<in>A. k x)})"
by (simp add: A(3))
qed
instantiation measure :: (type) complete_lattice
begin
interpretation sup_measure: comm_monoid_set sup "bot :: 'a measure"
by standard (auto intro!: antisym)
lemma sup_measure_F_mono':
"finite J \<Longrightarrow> finite I \<Longrightarrow> sup_measure.F id I \<le> sup_measure.F id (I \<union> J)"
proof (induction J rule: finite_induct)
case empty then show ?case
by simp
next
case (insert i J)
show ?case
proof cases
assume "i \<in> I" with insert show ?thesis
by (auto simp: insert_absorb)
next
assume "i \<notin> I"
have "sup_measure.F id I \<le> sup_measure.F id (I \<union> J)"
by (intro insert)
also have "\<dots> \<le> sup_measure.F id (insert i (I \<union> J))"
using insert \<open>i \<notin> I\<close> by (subst sup_measure.insert) auto
finally show ?thesis
by auto
qed
qed
lemma sup_measure_F_mono: "finite I \<Longrightarrow> J \<subseteq> I \<Longrightarrow> sup_measure.F id J \<le> sup_measure.F id I"
using sup_measure_F_mono'[of I J] by (auto simp: finite_subset Un_absorb1)
lemma sets_sup_measure_F:
"finite I \<Longrightarrow> I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> sets i = sets M) \<Longrightarrow> sets (sup_measure.F id I) = sets M"
by (induction I rule: finite_ne_induct) (simp_all add: sets_sup)
definition\<^marker>\<open>tag important\<close> Sup_measure' :: "'a measure set \<Rightarrow> 'a measure" where
"Sup_measure' M =
measure_of (\<Union>a\<in>M. space a) (\<Union>a\<in>M. sets a)
(\<lambda>X. (SUP P\<in>{P. finite P \<and> P \<subseteq> M }. sup_measure.F id P X))"
lemma space_Sup_measure'2: "space (Sup_measure' M) = (\<Union>m\<in>M. space m)"
unfolding Sup_measure'_def by (intro space_measure_of[OF UN_space_closed])
lemma sets_Sup_measure'2: "sets (Sup_measure' M) = sigma_sets (\<Union>m\<in>M. space m) (\<Union>m\<in>M. sets m)"
unfolding Sup_measure'_def by (intro sets_measure_of[OF UN_space_closed])
lemma sets_Sup_measure':
assumes sets_eq[simp]: "\<And>m. m \<in> M \<Longrightarrow> sets m = sets A" and "M \<noteq> {}"
shows "sets (Sup_measure' M) = sets A"
using sets_eq[THEN sets_eq_imp_space_eq, simp] \<open>M \<noteq> {}\<close> by (simp add: Sup_measure'_def)
lemma space_Sup_measure':
assumes sets_eq[simp]: "\<And>m. m \<in> M \<Longrightarrow> sets m = sets A" and "M \<noteq> {}"
shows "space (Sup_measure' M) = space A"
using sets_eq[THEN sets_eq_imp_space_eq, simp] \<open>M \<noteq> {}\<close>
by (simp add: Sup_measure'_def )
lemma emeasure_Sup_measure':
assumes sets_eq[simp]: "\<And>m. m \<in> M \<Longrightarrow> sets m = sets A" and "X \<in> sets A" "M \<noteq> {}"
shows "emeasure (Sup_measure' M) X = (SUP P\<in>{P. finite P \<and> P \<subseteq> M}. sup_measure.F id P X)"
(is "_ = ?S X")
using Sup_measure'_def
proof (rule emeasure_measure_of)
note sets_eq[THEN sets_eq_imp_space_eq, simp]
have *: "sets (Sup_measure' M) = sets A" "space (Sup_measure' M) = space A"
using \<open>M \<noteq> {}\<close> by (simp_all add: Sup_measure'_def)
let ?\<mu> = "sup_measure.F id"
show "countably_additive (sets (Sup_measure' M)) ?S"
proof (rule countably_additiveI, goal_cases)
case (1 F)
then have **: "range F \<subseteq> sets A"
by (auto simp: *)
show "(\<Sum>i. ?S (F i)) = ?S (\<Union>i. F i)"
proof (subst ennreal_suminf_SUP_eq_directed)
fix i j and N :: "nat set" assume ij: "i \<in> {P. finite P \<and> P \<subseteq> M}" "j \<in> {P. finite P \<and> P \<subseteq> M}"
have "(i \<noteq> {} \<longrightarrow> sets (?\<mu> i) = sets A) \<and> (j \<noteq> {} \<longrightarrow> sets (?\<mu> j) = sets A) \<and>
(i \<noteq> {} \<or> j \<noteq> {} \<longrightarrow> sets (?\<mu> (i \<union> j)) = sets A)"
using ij by (intro impI sets_sup_measure_F conjI) auto
then have "?\<mu> j (F n) \<le> ?\<mu> (i \<union> j) (F n) \<and> ?\<mu> i (F n) \<le> ?\<mu> (i \<union> j) (F n)" for n
using ij
by (cases "i = {}"; cases "j = {}")
(auto intro!: le_measureD3 sup_measure_F_mono simp: sets_sup_measure_F
simp del: id_apply)
with ij show "\<exists>k\<in>{P. finite P \<and> P \<subseteq> M}. \<forall>n\<in>N. ?\<mu> i (F n) \<le> ?\<mu> k (F n) \<and> ?\<mu> j (F n) \<le> ?\<mu> k (F n)"
by (safe intro!: bexI[of _ "i \<union> j"]) auto
next
show "(SUP P \<in> {P. finite P \<and> P \<subseteq> M}. \<Sum>n. ?\<mu> P (F n)) = (SUP P \<in> {P. finite P \<and> P \<subseteq> M}. ?\<mu> P (\<Union>(F ` UNIV)))"
proof (intro arg_cong [of _ _ Sup] image_cong refl)
fix i assume i: "i \<in> {P. finite P \<and> P \<subseteq> M}"
show "(\<Sum>n. ?\<mu> i (F n)) = ?\<mu> i (\<Union>(F ` UNIV))"
proof cases
assume "i \<noteq> {}" with i ** show ?thesis
apply (intro suminf_emeasure \<open>disjoint_family F\<close>)
apply (subst sets_sup_measure_F[OF _ _ sets_eq])
apply auto
done
qed simp
qed
qed
qed
show "positive (sets (Sup_measure' M)) ?S"
by (auto simp: positive_def bot_ennreal[symmetric])
show "X \<in> sets (Sup_measure' M)"
using assms * by auto
qed (rule UN_space_closed)
definition\<^marker>\<open>tag important\<close> Sup_measure :: "'a measure set \<Rightarrow> 'a measure" where
"Sup_measure =
Sup_lexord space
(Sup_lexord sets Sup_measure'
(\<lambda>U. sigma (\<Union>u\<in>U. space u) (\<Union>u\<in>U. sets u)))
(\<lambda>U. sigma (\<Union>u\<in>U. space u) {})"
definition\<^marker>\<open>tag important\<close> Inf_measure :: "'a measure set \<Rightarrow> 'a measure" where
"Inf_measure A = Sup {x. \<forall>a\<in>A. x \<le> a}"
definition\<^marker>\<open>tag important\<close> inf_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure" where
"inf_measure a b = Inf {a, b}"
definition\<^marker>\<open>tag important\<close> top_measure :: "'a measure" where
"top_measure = Inf {}"
instance
proof
note UN_space_closed [simp]
show upper: "x \<le> Sup A" if x: "x \<in> A" for x :: "'a measure" and A
unfolding Sup_measure_def
proof (intro Sup_lexord[where P="\<lambda>y. x \<le> y"])
assume "\<And>a. a \<in> A \<Longrightarrow> space a \<noteq> (\<Union>a\<in>A. space a)"
from this[OF \<open>x \<in> A\<close>] \<open>x \<in> A\<close> show "x \<le> sigma (\<Union>a\<in>A. space a) {}"
by (intro less_eq_measure.intros) auto
next
fix a S assume "a \<in> A" and a: "space a = (\<Union>a\<in>A. space a)" and S: "S = {a' \<in> A. space a' = space a}"
and neq: "\<And>aa. aa \<in> S \<Longrightarrow> sets aa \<noteq> (\<Union>a\<in>S. sets a)"
have sp_a: "space a = (\<Union>(space ` S))"
using \<open>a\<in>A\<close> by (auto simp: S)
show "x \<le> sigma (\<Union>(space ` S)) (\<Union>(sets ` S))"
proof cases
assume [simp]: "space x = space a"
have "sets x \<subset> (\<Union>a\<in>S. sets a)"
using \<open>x\<in>A\<close> neq[of x] by (auto simp: S)
also have "\<dots> \<subseteq> sigma_sets (\<Union>x\<in>S. space x) (\<Union>x\<in>S. sets x)"
by (rule sigma_sets_superset_generator)
finally show ?thesis
by (intro less_eq_measure.intros(2)) (simp_all add: sp_a)
next
assume "space x \<noteq> space a"
moreover have "space x \<le> space a"
unfolding a using \<open>x\<in>A\<close> by auto
ultimately show ?thesis
by (intro less_eq_measure.intros) (simp add: less_le sp_a)
qed
next
fix a b S S' assume "a \<in> A" and a: "space a = (\<Union>a\<in>A. space a)" and S: "S = {a' \<in> A. space a' = space a}"
and "b \<in> S" and b: "sets b = (\<Union>a\<in>S. sets a)" and S': "S' = {a' \<in> S. sets a' = sets b}"
then have "S' \<noteq> {}" "space b = space a"
by auto
have sets_eq: "\<And>x. x \<in> S' \<Longrightarrow> sets x = sets b"
by (auto simp: S')
note sets_eq[THEN sets_eq_imp_space_eq, simp]
have *: "sets (Sup_measure' S') = sets b" "space (Sup_measure' S') = space b"
using \<open>S' \<noteq> {}\<close> by (simp_all add: Sup_measure'_def sets_eq)
show "x \<le> Sup_measure' S'"
proof cases
assume "x \<in> S"
with \<open>b \<in> S\<close> have "space x = space b"
by (simp add: S)
show ?thesis
proof cases
assume "x \<in> S'"
show "x \<le> Sup_measure' S'"
proof (intro le_measure[THEN iffD2] ballI)
show "sets x = sets (Sup_measure' S')"
using \<open>x\<in>S'\<close> * by (simp add: S')
fix X assume "X \<in> sets x"
show "emeasure x X \<le> emeasure (Sup_measure' S') X"
proof (subst emeasure_Sup_measure'[OF _ \<open>X \<in> sets x\<close>])
show "emeasure x X \<le> (SUP P \<in> {P. finite P \<and> P \<subseteq> S'}. emeasure (sup_measure.F id P) X)"
using \<open>x\<in>S'\<close> by (intro SUP_upper2[where i="{x}"]) auto
qed (insert \<open>x\<in>S'\<close> S', auto)
qed
next
assume "x \<notin> S'"
then have "sets x \<noteq> sets b"
using \<open>x\<in>S\<close> by (auto simp: S')
moreover have "sets x \<le> sets b"
using \<open>x\<in>S\<close> unfolding b by auto
ultimately show ?thesis
using * \<open>x \<in> S\<close>
by (intro less_eq_measure.intros(2))
(simp_all add: * \<open>space x = space b\<close> less_le)
qed
next
assume "x \<notin> S"
with \<open>x\<in>A\<close> \<open>x \<notin> S\<close> \<open>space b = space a\<close> show ?thesis
by (intro less_eq_measure.intros)
(simp_all add: * less_le a SUP_upper S)
qed
qed
show least: "Sup A \<le> x" if x: "\<And>z. z \<in> A \<Longrightarrow> z \<le> x" for x :: "'a measure" and A
unfolding Sup_measure_def
proof (intro Sup_lexord[where P="\<lambda>y. y \<le> x"])
assume "\<And>a. a \<in> A \<Longrightarrow> space a \<noteq> (\<Union>a\<in>A. space a)"
show "sigma (\<Union>(space ` A)) {} \<le> x"
using x[THEN le_measureD1] by (subst sigma_le_iff) auto
next
fix a S assume "a \<in> A" "space a = (\<Union>a\<in>A. space a)" and S: "S = {a' \<in> A. space a' = space a}"
"\<And>a. a \<in> S \<Longrightarrow> sets a \<noteq> (\<Union>a\<in>S. sets a)"
have "\<Union>(space ` S) \<subseteq> space x"
using S le_measureD1[OF x] by auto
moreover
have "\<Union>(space ` S) = space a"
using \<open>a\<in>A\<close> S by auto
then have "space x = \<Union>(space ` S) \<Longrightarrow> \<Union>(sets ` S) \<subseteq> sets x"
using \<open>a \<in> A\<close> le_measureD2[OF x] by (auto simp: S)
ultimately show "sigma (\<Union>(space ` S)) (\<Union>(sets ` S)) \<le> x"
by (subst sigma_le_iff) simp_all
next
fix a b S S' assume "a \<in> A" and a: "space a = (\<Union>a\<in>A. space a)" and S: "S = {a' \<in> A. space a' = space a}"
and "b \<in> S" and b: "sets b = (\<Union>a\<in>S. sets a)" and S': "S' = {a' \<in> S. sets a' = sets b}"
then have "S' \<noteq> {}" "space b = space a"
by auto
have sets_eq: "\<And>x. x \<in> S' \<Longrightarrow> sets x = sets b"
by (auto simp: S')
note sets_eq[THEN sets_eq_imp_space_eq, simp]
have *: "sets (Sup_measure' S') = sets b" "space (Sup_measure' S') = space b"
using \<open>S' \<noteq> {}\<close> by (simp_all add: Sup_measure'_def sets_eq)
show "Sup_measure' S' \<le> x"
proof cases
assume "space x = space a"
show ?thesis
proof cases
assume **: "sets x = sets b"
show ?thesis
proof (intro le_measure[THEN iffD2] ballI)
show ***: "sets (Sup_measure' S') = sets x"
by (simp add: * **)
fix X assume "X \<in> sets (Sup_measure' S')"
show "emeasure (Sup_measure' S') X \<le> emeasure x X"
unfolding ***
proof (subst emeasure_Sup_measure'[OF _ \<open>X \<in> sets (Sup_measure' S')\<close>])
show "(SUP P \<in> {P. finite P \<and> P \<subseteq> S'}. emeasure (sup_measure.F id P) X) \<le> emeasure x X"
proof (safe intro!: SUP_least)
fix P assume P: "finite P" "P \<subseteq> S'"
show "emeasure (sup_measure.F id P) X \<le> emeasure x X"
proof cases
assume "P = {}" then show ?thesis
by auto
next
assume "P \<noteq> {}"
from P have "finite P" "P \<subseteq> A"
unfolding S' S by (simp_all add: subset_eq)
then have "sup_measure.F id P \<le> x"
by (induction P) (auto simp: x)
moreover have "sets (sup_measure.F id P) = sets x"
using \<open>finite P\<close> \<open>P \<noteq> {}\<close> \<open>P \<subseteq> S'\<close> \<open>sets x = sets b\<close>
by (intro sets_sup_measure_F) (auto simp: S')
ultimately show "emeasure (sup_measure.F id P) X \<le> emeasure x X"
by (rule le_measureD3)
qed
qed
show "m \<in> S' \<Longrightarrow> sets m = sets (Sup_measure' S')" for m
unfolding * by (simp add: S')
qed fact
qed
next
assume "sets x \<noteq> sets b"
moreover have "sets b \<le> sets x"
unfolding b S using x[THEN le_measureD2] \<open>space x = space a\<close> by auto
ultimately show "Sup_measure' S' \<le> x"
using \<open>space x = space a\<close> \<open>b \<in> S\<close>
by (intro less_eq_measure.intros(2)) (simp_all add: * S)
qed
next
assume "space x \<noteq> space a"
then have "space a < space x"
using le_measureD1[OF x[OF \<open>a\<in>A\<close>]] by auto
then show "Sup_measure' S' \<le> x"
by (intro less_eq_measure.intros) (simp add: * \<open>space b = space a\<close>)
qed
qed
show "Sup {} = (bot::'a measure)" "Inf {} = (top::'a measure)"
by (auto intro!: antisym least simp: top_measure_def)
show lower: "x \<in> A \<Longrightarrow> Inf A \<le> x" for x :: "'a measure" and A
unfolding Inf_measure_def by (intro least) auto
show greatest: "(\<And>z. z \<in> A \<Longrightarrow> x \<le> z) \<Longrightarrow> x \<le> Inf A" for x :: "'a measure" and A
unfolding Inf_measure_def by (intro upper) auto
show "inf x y \<le> x" "inf x y \<le> y" "x \<le> y \<Longrightarrow> x \<le> z \<Longrightarrow> x \<le> inf y z" for x y z :: "'a measure"
by (auto simp: inf_measure_def intro!: lower greatest)
qed
end
lemma sets_SUP:
assumes "\<And>x. x \<in> I \<Longrightarrow> sets (M x) = sets N"
shows "I \<noteq> {} \<Longrightarrow> sets (SUP i\<in>I. M i) = sets N"
unfolding Sup_measure_def
using assms assms[THEN sets_eq_imp_space_eq]
sets_Sup_measure'[where A=N and M="M`I"]
by (intro Sup_lexord1[where P="\<lambda>x. sets x = sets N"]) auto
lemma emeasure_SUP:
assumes sets: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sets N" "X \<in> sets N" "I \<noteq> {}"
shows "emeasure (SUP i\<in>I. M i) X = (SUP J\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. emeasure (SUP i\<in>J. M i) X)"
proof -
interpret sup_measure: comm_monoid_set sup "bot :: 'b measure"
by standard (auto intro!: antisym)
have eq: "finite J \<Longrightarrow> sup_measure.F id J = (SUP i\<in>J. i)" for J :: "'b measure set"
by (induction J rule: finite_induct) auto
have 1: "J \<noteq> {} \<Longrightarrow> J \<subseteq> I \<Longrightarrow> sets (SUP x\<in>J. M x) = sets N" for J
by (intro sets_SUP sets) (auto )
from \<open>I \<noteq> {}\<close> obtain i where "i\<in>I" by auto
have "Sup_measure' (M`I) X = (SUP P\<in>{P. finite P \<and> P \<subseteq> M`I}. sup_measure.F id P X)"
using sets by (intro emeasure_Sup_measure') auto
also have "Sup_measure' (M`I) = (SUP i\<in>I. M i)"
unfolding Sup_measure_def using \<open>I \<noteq> {}\<close> sets sets(1)[THEN sets_eq_imp_space_eq]
by (intro Sup_lexord1[where P="\<lambda>x. _ = x"]) auto
also have "(SUP P\<in>{P. finite P \<and> P \<subseteq> M`I}. sup_measure.F id P X) =
(SUP J\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. (SUP i\<in>J. M i) X)"
proof (intro SUP_eq)
fix J assume "J \<in> {P. finite P \<and> P \<subseteq> M`I}"
then obtain J' where J': "J' \<subseteq> I" "finite J'" and J: "J = M`J'" and "finite J"
using finite_subset_image[of J M I] by auto
show "\<exists>j\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. sup_measure.F id J X \<le> (SUP i\<in>j. M i) X"
proof cases
assume "J' = {}" with \<open>i \<in> I\<close> show ?thesis
by (auto simp add: J)
next
assume "J' \<noteq> {}" with J J' show ?thesis
by (intro bexI[of _ "J'"]) (auto simp add: eq simp del: id_apply)
qed
next
fix J assume J: "J \<in> {P. P \<noteq> {} \<and> finite P \<and> P \<subseteq> I}"
show "\<exists>J'\<in>{J. finite J \<and> J \<subseteq> M`I}. (SUP i\<in>J. M i) X \<le> sup_measure.F id J' X"
using J by (intro bexI[of _ "M`J"]) (auto simp add: eq simp del: id_apply)
qed
finally show ?thesis .
qed
lemma emeasure_SUP_chain:
assumes sets: "\<And>i. i \<in> A \<Longrightarrow> sets (M i) = sets N" "X \<in> sets N"
assumes ch: "Complete_Partial_Order.chain (\<le>) (M ` A)" and "A \<noteq> {}"
shows "emeasure (SUP i\<in>A. M i) X = (SUP i\<in>A. emeasure (M i) X)"
proof (subst emeasure_SUP[OF sets \<open>A \<noteq> {}\<close>])
show "(SUP J\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> A}. emeasure (Sup (M ` J)) X) = (SUP i\<in>A. emeasure (M i) X)"
proof (rule SUP_eq)
fix J assume "J \<in> {J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> A}"
then have J: "Complete_Partial_Order.chain (\<le>) (M ` J)" "finite J" "J \<noteq> {}" and "J \<subseteq> A"
using ch[THEN chain_subset, of "M`J"] by auto
with in_chain_finite[OF J(1)] obtain j where "j \<in> J" "(SUP j\<in>J. M j) = M j"
by auto
with \<open>J \<subseteq> A\<close> show "\<exists>j\<in>A. emeasure (Sup (M ` J)) X \<le> emeasure (M j) X"
by auto
next
fix j assume "j\<in>A" then show "\<exists>i\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> A}. emeasure (M j) X \<le> emeasure (Sup (M ` i)) X"
by (intro bexI[of _ "{j}"]) auto
qed
qed
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Supremum of a set of \<open>\<sigma>\<close>-algebras\<close>
lemma space_Sup_eq_UN: "space (Sup M) = (\<Union>x\<in>M. space x)"
unfolding Sup_measure_def
apply (intro Sup_lexord[where P="\<lambda>x. space x = _"])
apply (subst space_Sup_measure'2)
apply auto []
apply (subst space_measure_of[OF UN_space_closed])
apply auto
done
lemma sets_Sup_eq:
assumes *: "\<And>m. m \<in> M \<Longrightarrow> space m = X" and "M \<noteq> {}"
shows "sets (Sup M) = sigma_sets X (\<Union>x\<in>M. sets x)"
unfolding Sup_measure_def
apply (rule Sup_lexord1)
apply fact
apply (simp add: assms)
apply (rule Sup_lexord)
subgoal premises that for a S
unfolding that(3) that(2)[symmetric]
using that(1)
apply (subst sets_Sup_measure'2)
apply (intro arg_cong2[where f=sigma_sets])
apply (auto simp: *)
done
apply (subst sets_measure_of[OF UN_space_closed])
apply (simp add: assms)
done
lemma in_sets_Sup: "(\<And>m. m \<in> M \<Longrightarrow> space m = X) \<Longrightarrow> m \<in> M \<Longrightarrow> A \<in> sets m \<Longrightarrow> A \<in> sets (Sup M)"
by (subst sets_Sup_eq[where X=X]) auto
lemma Sup_lexord_rel:
assumes "\<And>i. i \<in> I \<Longrightarrow> k (A i) = k (B i)"
"R (c (A ` {a \<in> I. k (B a) = (SUP x\<in>I. k (B x))})) (c (B ` {a \<in> I. k (B a) = (SUP x\<in>I. k (B x))}))"
"R (s (A`I)) (s (B`I))"
shows "R (Sup_lexord k c s (A`I)) (Sup_lexord k c s (B`I))"
proof -
have "A ` {a \<in> I. k (B a) = (SUP x\<in>I. k (B x))} = {a \<in> A ` I. k a = (SUP x\<in>I. k (B x))}"
using assms(1) by auto
moreover have "B ` {a \<in> I. k (B a) = (SUP x\<in>I. k (B x))} = {a \<in> B ` I. k a = (SUP x\<in>I. k (B x))}"
by auto
ultimately show ?thesis
using assms by (auto simp: Sup_lexord_def Let_def image_comp)
qed
lemma sets_SUP_cong:
assumes eq: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sets (N i)" shows "sets (SUP i\<in>I. M i) = sets (SUP i\<in>I. N i)"
unfolding Sup_measure_def
using eq eq[THEN sets_eq_imp_space_eq]
apply (intro Sup_lexord_rel[where R="\<lambda>x y. sets x = sets y"])
apply simp
apply simp
apply (simp add: sets_Sup_measure'2)
apply (intro arg_cong2[where f="\<lambda>x y. sets (sigma x y)"])
apply auto
done
lemma sets_Sup_in_sets:
assumes "M \<noteq> {}"
assumes "\<And>m. m \<in> M \<Longrightarrow> space m = space N"
assumes "\<And>m. m \<in> M \<Longrightarrow> sets m \<subseteq> sets N"
shows "sets (Sup M) \<subseteq> sets N"
proof -
have *: "\<Union>(space ` M) = space N"
using assms by auto
show ?thesis
unfolding * using assms by (subst sets_Sup_eq[of M "space N"]) (auto intro!: sets.sigma_sets_subset)
qed
lemma measurable_Sup1:
assumes m: "m \<in> M" and f: "f \<in> measurable m N"
and const_space: "\<And>m n. m \<in> M \<Longrightarrow> n \<in> M \<Longrightarrow> space m = space n"
shows "f \<in> measurable (Sup M) N"
proof -
have "space (Sup M) = space m"
using m by (auto simp add: space_Sup_eq_UN dest: const_space)
then show ?thesis
using m f unfolding measurable_def by (auto intro: in_sets_Sup[OF const_space])
qed
lemma measurable_Sup2:
assumes M: "M \<noteq> {}"
assumes f: "\<And>m. m \<in> M \<Longrightarrow> f \<in> measurable N m"
and const_space: "\<And>m n. m \<in> M \<Longrightarrow> n \<in> M \<Longrightarrow> space m = space n"
shows "f \<in> measurable N (Sup M)"
proof -
from M obtain m where "m \<in> M" by auto
have space_eq: "\<And>n. n \<in> M \<Longrightarrow> space n = space m"
by (intro const_space \<open>m \<in> M\<close>)
have "f \<in> measurable N (sigma (\<Union>m\<in>M. space m) (\<Union>m\<in>M. sets m))"
proof (rule measurable_measure_of)
show "f \<in> space N \<rightarrow> \<Union>(space ` M)"
using measurable_space[OF f] M by auto
qed (auto intro: measurable_sets f dest: sets.sets_into_space)
also have "measurable N (sigma (\<Union>m\<in>M. space m) (\<Union>m\<in>M. sets m)) = measurable N (Sup M)"
apply (intro measurable_cong_sets refl)
apply (subst sets_Sup_eq[OF space_eq M])
apply simp
apply (subst sets_measure_of[OF UN_space_closed])
apply (simp add: space_eq M)
done
finally show ?thesis .
qed
lemma measurable_SUP2:
"I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> f \<in> measurable N (M i)) \<Longrightarrow>
(\<And>i j. i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> space (M i) = space (M j)) \<Longrightarrow> f \<in> measurable N (SUP i\<in>I. M i)"
by (auto intro!: measurable_Sup2)
lemma sets_Sup_sigma:
assumes [simp]: "M \<noteq> {}" and M: "\<And>m. m \<in> M \<Longrightarrow> m \<subseteq> Pow \<Omega>"
shows "sets (SUP m\<in>M. sigma \<Omega> m) = sets (sigma \<Omega> (\<Union>M))"
proof -
{ fix a m assume "a \<in> sigma_sets \<Omega> m" "m \<in> M"
then have "a \<in> sigma_sets \<Omega> (\<Union>M)"
by induction (auto intro: sigma_sets.intros(2-)) }
then show "sets (SUP m\<in>M. sigma \<Omega> m) = sets (sigma \<Omega> (\<Union>M))"
apply (subst sets_Sup_eq[where X="\<Omega>"])
apply (auto simp add: M) []
apply auto []
apply (simp add: space_measure_of_conv M Union_least)
apply (rule sigma_sets_eqI)
apply auto
done
qed
lemma Sup_sigma:
assumes [simp]: "M \<noteq> {}" and M: "\<And>m. m \<in> M \<Longrightarrow> m \<subseteq> Pow \<Omega>"
shows "(SUP m\<in>M. sigma \<Omega> m) = (sigma \<Omega> (\<Union>M))"
proof (intro antisym SUP_least)
have *: "\<Union>M \<subseteq> Pow \<Omega>"
using M by auto
show "sigma \<Omega> (\<Union>M) \<le> (SUP m\<in>M. sigma \<Omega> m)"
proof (intro less_eq_measure.intros(3))
show "space (sigma \<Omega> (\<Union>M)) = space (SUP m\<in>M. sigma \<Omega> m)"
"sets (sigma \<Omega> (\<Union>M)) = sets (SUP m\<in>M. sigma \<Omega> m)"
using sets_Sup_sigma[OF assms] sets_Sup_sigma[OF assms, THEN sets_eq_imp_space_eq]
by auto
qed (simp add: emeasure_sigma le_fun_def)
fix m assume "m \<in> M" then show "sigma \<Omega> m \<le> sigma \<Omega> (\<Union>M)"
by (subst sigma_le_iff) (auto simp add: M *)
qed
lemma SUP_sigma_sigma:
"M \<noteq> {} \<Longrightarrow> (\<And>m. m \<in> M \<Longrightarrow> f m \<subseteq> Pow \<Omega>) \<Longrightarrow> (SUP m\<in>M. sigma \<Omega> (f m)) = sigma \<Omega> (\<Union>m\<in>M. f m)"
using Sup_sigma[of "f`M" \<Omega>] by (auto simp: image_comp)
lemma sets_vimage_Sup_eq:
assumes *: "M \<noteq> {}" "f \<in> X \<rightarrow> Y" "\<And>m. m \<in> M \<Longrightarrow> space m = Y"
shows "sets (vimage_algebra X f (Sup M)) = sets (SUP m \<in> M. vimage_algebra X f m)"
(is "?IS = ?SI")
proof
show "?IS \<subseteq> ?SI"
apply (intro sets_image_in_sets measurable_Sup2)
apply (simp add: space_Sup_eq_UN *)
apply (simp add: *)
apply (intro measurable_Sup1)
apply (rule imageI)
apply assumption
apply (rule measurable_vimage_algebra1)
apply (auto simp: *)
done
show "?SI \<subseteq> ?IS"
apply (intro sets_Sup_in_sets)
apply (auto simp: *) []
apply (auto simp: *) []
apply (elim imageE)
apply simp
apply (rule sets_image_in_sets)
apply simp
apply (simp add: measurable_def)
apply (simp add: * space_Sup_eq_UN sets_vimage_algebra2)
apply (auto intro: in_sets_Sup[OF *(3)])
done
qed
lemma restrict_space_eq_vimage_algebra':
"sets (restrict_space M \<Omega>) = sets (vimage_algebra (\<Omega> \<inter> space M) (\<lambda>x. x) M)"
proof -
have *: "{A \<inter> (\<Omega> \<inter> space M) |A. A \<in> sets M} = {A \<inter> \<Omega> |A. A \<in> sets M}"
using sets.sets_into_space[of _ M] by blast
show ?thesis
unfolding restrict_space_def
by (subst sets_measure_of)
(auto simp add: image_subset_iff sets_vimage_algebra * dest: sets.sets_into_space intro!: arg_cong2[where f=sigma_sets])
qed
lemma sigma_le_sets:
assumes [simp]: "A \<subseteq> Pow X" shows "sets (sigma X A) \<subseteq> sets N \<longleftrightarrow> X \<in> sets N \<and> A \<subseteq> sets N"
proof
have "X \<in> sigma_sets X A" "A \<subseteq> sigma_sets X A"
by (auto intro: sigma_sets_top)
moreover assume "sets (sigma X A) \<subseteq> sets N"
ultimately show "X \<in> sets N \<and> A \<subseteq> sets N"
by auto
next
assume *: "X \<in> sets N \<and> A \<subseteq> sets N"
{ fix Y assume "Y \<in> sigma_sets X A" from this * have "Y \<in> sets N"
by induction auto }
then show "sets (sigma X A) \<subseteq> sets N"
by auto
qed
lemma measurable_iff_sets:
"f \<in> measurable M N \<longleftrightarrow> (f \<in> space M \<rightarrow> space N \<and> sets (vimage_algebra (space M) f N) \<subseteq> sets M)"
proof -
have *: "{f -` A \<inter> space M |A. A \<in> sets N} \<subseteq> Pow (space M)"
by auto
show ?thesis
unfolding measurable_def
by (auto simp add: vimage_algebra_def sigma_le_sets[OF *])
qed
lemma sets_vimage_algebra_space: "X \<in> sets (vimage_algebra X f M)"
using sets.top[of "vimage_algebra X f M"] by simp
lemma measurable_mono:
assumes N: "sets N' \<le> sets N" "space N = space N'"
assumes M: "sets M \<le> sets M'" "space M = space M'"
shows "measurable M N \<subseteq> measurable M' N'"
unfolding measurable_def
proof safe
fix f A assume "f \<in> space M \<rightarrow> space N" "A \<in> sets N'"
moreover assume "\<forall>y\<in>sets N. f -` y \<inter> space M \<in> sets M" note this[THEN bspec, of A]
ultimately show "f -` A \<inter> space M' \<in> sets M'"
using assms by auto
qed (insert N M, auto)
lemma measurable_Sup_measurable:
assumes f: "f \<in> space N \<rightarrow> A"
shows "f \<in> measurable N (Sup {M. space M = A \<and> f \<in> measurable N M})"
proof (rule measurable_Sup2)
show "{M. space M = A \<and> f \<in> measurable N M} \<noteq> {}"
using f unfolding ex_in_conv[symmetric]
by (intro exI[of _ "sigma A {}"]) (auto intro!: measurable_measure_of)
qed auto
lemma (in sigma_algebra) sigma_sets_subset':
assumes a: "a \<subseteq> M" "\<Omega>' \<in> M"
shows "sigma_sets \<Omega>' a \<subseteq> M"
proof
show "x \<in> M" if x: "x \<in> sigma_sets \<Omega>' a" for x
using x by (induct rule: sigma_sets.induct) (insert a, auto)
qed
lemma in_sets_SUP: "i \<in> I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> space (M i) = Y) \<Longrightarrow> X \<in> sets (M i) \<Longrightarrow> X \<in> sets (SUP i\<in>I. M i)"
by (intro in_sets_Sup[where X=Y]) auto
lemma measurable_SUP1:
"i \<in> I \<Longrightarrow> f \<in> measurable (M i) N \<Longrightarrow> (\<And>m n. m \<in> I \<Longrightarrow> n \<in> I \<Longrightarrow> space (M m) = space (M n)) \<Longrightarrow>
f \<in> measurable (SUP i\<in>I. M i) N"
by (auto intro: measurable_Sup1)
lemma sets_image_in_sets':
assumes X: "X \<in> sets N"
assumes f: "\<And>A. A \<in> sets M \<Longrightarrow> f -` A \<inter> X \<in> sets N"
shows "sets (vimage_algebra X f M) \<subseteq> sets N"
unfolding sets_vimage_algebra
by (rule sets.sigma_sets_subset') (auto intro!: measurable_sets X f)
lemma mono_vimage_algebra:
"sets M \<le> sets N \<Longrightarrow> sets (vimage_algebra X f M) \<subseteq> sets (vimage_algebra X f N)"
using sets.top[of "sigma X {f -` A \<inter> X |A. A \<in> sets N}"]
unfolding vimage_algebra_def
apply (subst (asm) space_measure_of)
apply auto []
apply (subst sigma_le_sets)
apply auto
done
lemma mono_restrict_space: "sets M \<le> sets N \<Longrightarrow> sets (restrict_space M X) \<subseteq> sets (restrict_space N X)"
unfolding sets_restrict_space by (rule image_mono)
lemma sets_eq_bot: "sets M = {{}} \<longleftrightarrow> M = bot"
apply safe
apply (intro measure_eqI)
apply auto
done
lemma sets_eq_bot2: "{{}} = sets M \<longleftrightarrow> M = bot"
using sets_eq_bot[of M] by blast
lemma (in finite_measure) countable_support:
"countable {x. measure M {x} \<noteq> 0}"
proof cases
assume "measure M (space M) = 0"
with bounded_measure measure_le_0_iff have "{x. measure M {x} \<noteq> 0} = {}"
by auto
then show ?thesis
by simp
next
let ?M = "measure M (space M)" and ?m = "\<lambda>x. measure M {x}"
assume "?M \<noteq> 0"
then have *: "{x. ?m x \<noteq> 0} = (\<Union>n. {x. ?M / Suc n < ?m x})"
using reals_Archimedean[of "?m x / ?M" for x]
by (auto simp: field_simps not_le[symmetric] divide_le_0_iff measure_le_0_iff)
have **: "\<And>n. finite {x. ?M / Suc n < ?m x}"
proof (rule ccontr)
fix n assume "infinite {x. ?M / Suc n < ?m x}" (is "infinite ?X")
then obtain X where "finite X" "card X = Suc (Suc n)" "X \<subseteq> ?X"
by (metis infinite_arbitrarily_large)
from this(3) have *: "\<And>x. x \<in> X \<Longrightarrow> ?M / Suc n \<le> ?m x"
by auto
{ fix x assume "x \<in> X"
from \<open>?M \<noteq> 0\<close> *[OF this] have "?m x \<noteq> 0" by (auto simp: field_simps measure_le_0_iff)
then have "{x} \<in> sets M" by (auto dest: measure_notin_sets) }
note singleton_sets = this
have "?M < (\<Sum>x\<in>X. ?M / Suc n)"
using \<open>?M \<noteq> 0\<close>
by (simp add: \<open>card X = Suc (Suc n)\<close> field_simps less_le)
also have "\<dots> \<le> (\<Sum>x\<in>X. ?m x)"
by (rule sum_mono) fact
also have "\<dots> = measure M (\<Union>x\<in>X. {x})"
using singleton_sets \<open>finite X\<close>
by (intro finite_measure_finite_Union[symmetric]) (auto simp: disjoint_family_on_def)
finally have "?M < measure M (\<Union>x\<in>X. {x})" .
moreover have "measure M (\<Union>x\<in>X. {x}) \<le> ?M"
using singleton_sets[THEN sets.sets_into_space] by (intro finite_measure_mono) auto
ultimately show False by simp
qed
show ?thesis
unfolding * by (intro countable_UN countableI_type countable_finite[OF **])
qed
end
diff --git a/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy b/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy
--- a/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy
+++ b/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy
@@ -1,2507 +1,2506 @@
(* Title: HOL/Analysis/Nonnegative_Lebesgue_Integration.thy
Author: Johannes Hölzl, TU München
Author: Armin Heller, TU München
*)
section \<open>Lebesgue Integration for Nonnegative Functions\<close>
theory Nonnegative_Lebesgue_Integration
imports Measure_Space Borel_Space
begin
subsection\<^marker>\<open>tag unimportant\<close> \<open>Approximating functions\<close>
lemma AE_upper_bound_inf_ennreal:
fixes F G::"'a \<Rightarrow> ennreal"
assumes "\<And>e. (e::real) > 0 \<Longrightarrow> AE x in M. F x \<le> G x + e"
shows "AE x in M. F x \<le> G x"
proof -
have "AE x in M. \<forall>n::nat. F x \<le> G x + ennreal (1 / Suc n)"
using assms by (auto simp: AE_all_countable)
then show ?thesis
proof (eventually_elim)
fix x assume x: "\<forall>n::nat. F x \<le> G x + ennreal (1 / Suc n)"
show "F x \<le> G x"
proof (rule ennreal_le_epsilon)
fix e :: real assume "0 < e"
then obtain n where n: "1 / Suc n < e"
by (blast elim: nat_approx_posE)
have "F x \<le> G x + 1 / Suc n"
using x by simp
also have "\<dots> \<le> G x + e"
using n by (intro add_mono ennreal_leI) auto
finally show "F x \<le> G x + ennreal e" .
qed
qed
qed
lemma AE_upper_bound_inf:
fixes F G::"'a \<Rightarrow> real"
assumes "\<And>e. e > 0 \<Longrightarrow> AE x in M. F x \<le> G x + e"
shows "AE x in M. F x \<le> G x"
proof -
have "AE x in M. F x \<le> G x + 1/real (n+1)" for n::nat
by (rule assms, auto)
then have "AE x in M. \<forall>n::nat \<in> UNIV. F x \<le> G x + 1/real (n+1)"
by (rule AE_ball_countable', auto)
moreover
{
fix x assume i: "\<forall>n::nat \<in> UNIV. F x \<le> G x + 1/real (n+1)"
have "(\<lambda>n. G x + 1/real (n+1)) \<longlonglongrightarrow> G x + 0"
by (rule tendsto_add, simp, rule LIMSEQ_ignore_initial_segment[OF lim_1_over_n, of 1])
then have "F x \<le> G x" using i LIMSEQ_le_const by fastforce
}
ultimately show ?thesis by auto
qed
lemma not_AE_zero_ennreal_E:
fixes f::"'a \<Rightarrow> ennreal"
assumes "\<not> (AE x in M. f x = 0)" and [measurable]: "f \<in> borel_measurable M"
shows "\<exists>A\<in>sets M. \<exists>e::real>0. emeasure M A > 0 \<and> (\<forall>x \<in> A. f x \<ge> e)"
proof -
{ assume "\<not> (\<exists>e::real>0. {x \<in> space M. f x \<ge> e} \<notin> null_sets M)"
then have "0 < e \<Longrightarrow> AE x in M. f x \<le> e" for e :: real
by (auto simp: not_le less_imp_le dest!: AE_not_in)
then have "AE x in M. f x \<le> 0"
by (intro AE_upper_bound_inf_ennreal[where G="\<lambda>_. 0"]) simp
then have False
using assms by auto }
then obtain e::real where e: "e > 0" "{x \<in> space M. f x \<ge> e} \<notin> null_sets M" by auto
define A where "A = {x \<in> space M. f x \<ge> e}"
have 1 [measurable]: "A \<in> sets M" unfolding A_def by auto
have 2: "emeasure M A > 0"
using e(2) A_def \<open>A \<in> sets M\<close> by auto
have 3: "\<And>x. x \<in> A \<Longrightarrow> f x \<ge> e" unfolding A_def by auto
show ?thesis using e(1) 1 2 3 by blast
qed
lemma not_AE_zero_E:
fixes f::"'a \<Rightarrow> real"
assumes "AE x in M. f x \<ge> 0"
"\<not>(AE x in M. f x = 0)"
and [measurable]: "f \<in> borel_measurable M"
shows "\<exists>A e. A \<in> sets M \<and> e>0 \<and> emeasure M A > 0 \<and> (\<forall>x \<in> A. f x \<ge> e)"
proof -
have "\<exists>e. e > 0 \<and> {x \<in> space M. f x \<ge> e} \<notin> null_sets M"
proof (rule ccontr)
assume *: "\<not>(\<exists>e. e > 0 \<and> {x \<in> space M. f x \<ge> e} \<notin> null_sets M)"
{
fix e::real assume "e > 0"
then have "{x \<in> space M. f x \<ge> e} \<in> null_sets M" using * by blast
then have "AE x in M. x \<notin> {x \<in> space M. f x \<ge> e}" using AE_not_in by blast
then have "AE x in M. f x \<le> e" by auto
}
then have "AE x in M. f x \<le> 0" by (rule AE_upper_bound_inf, auto)
then have "AE x in M. f x = 0" using assms(1) by auto
then show False using assms(2) by auto
qed
then obtain e where e: "e>0" "{x \<in> space M. f x \<ge> e} \<notin> null_sets M" by auto
define A where "A = {x \<in> space M. f x \<ge> e}"
have 1 [measurable]: "A \<in> sets M" unfolding A_def by auto
have 2: "emeasure M A > 0"
using e(2) A_def \<open>A \<in> sets M\<close> by auto
have 3: "\<And>x. x \<in> A \<Longrightarrow> f x \<ge> e" unfolding A_def by auto
show ?thesis
using e(1) 1 2 3 by blast
qed
subsection "Simple function"
text \<open>
Our simple functions are not restricted to nonnegative real numbers. Instead
they are just functions with a finite range and are measurable when singleton
sets are measurable.
\<close>
definition\<^marker>\<open>tag important\<close> "simple_function M g \<longleftrightarrow>
finite (g ` space M) \<and>
(\<forall>x \<in> g ` space M. g -` {x} \<inter> space M \<in> sets M)"
lemma simple_functionD:
assumes "simple_function M g"
shows "finite (g ` space M)" and "g -` X \<inter> space M \<in> sets M"
proof -
show "finite (g ` space M)"
using assms unfolding simple_function_def by auto
have "g -` X \<inter> space M = g -` (X \<inter> g`space M) \<inter> space M" by auto
also have "\<dots> = (\<Union>x\<in>X \<inter> g`space M. g-`{x} \<inter> space M)" by auto
finally show "g -` X \<inter> space M \<in> sets M" using assms
by (auto simp del: UN_simps simp: simple_function_def)
qed
lemma measurable_simple_function[measurable_dest]:
"simple_function M f \<Longrightarrow> f \<in> measurable M (count_space UNIV)"
unfolding simple_function_def measurable_def
proof safe
fix A assume "finite (f ` space M)" "\<forall>x\<in>f ` space M. f -` {x} \<inter> space M \<in> sets M"
then have "(\<Union>x\<in>f ` space M. if x \<in> A then f -` {x} \<inter> space M else {}) \<in> sets M"
by (intro sets.finite_UN) auto
also have "(\<Union>x\<in>f ` space M. if x \<in> A then f -` {x} \<inter> space M else {}) = f -` A \<inter> space M"
by (auto split: if_split_asm)
finally show "f -` A \<inter> space M \<in> sets M" .
qed simp
lemma borel_measurable_simple_function:
"simple_function M f \<Longrightarrow> f \<in> borel_measurable M"
by (auto dest!: measurable_simple_function simp: measurable_def)
lemma simple_function_measurable2[intro]:
assumes "simple_function M f" "simple_function M g"
shows "f -` A \<inter> g -` B \<inter> space M \<in> sets M"
proof -
have "f -` A \<inter> g -` B \<inter> space M = (f -` A \<inter> space M) \<inter> (g -` B \<inter> space M)"
by auto
then show ?thesis using assms[THEN simple_functionD(2)] by auto
qed
lemma simple_function_indicator_representation:
fixes f ::"'a \<Rightarrow> ennreal"
assumes f: "simple_function M f" and x: "x \<in> space M"
shows "f x = (\<Sum>y \<in> f ` space M. y * indicator (f -` {y} \<inter> space M) x)"
(is "?l = ?r")
proof -
have "?r = (\<Sum>y \<in> f ` space M.
(if y = f x then y * indicator (f -` {y} \<inter> space M) x else 0))"
by (auto intro!: sum.cong)
also have "... = f x * indicator (f -` {f x} \<inter> space M) x"
- using assms by (auto dest: simple_functionD simp: sum.delta)
+ using assms by (auto dest: simple_functionD)
also have "... = f x" using x by (auto simp: indicator_def)
finally show ?thesis by auto
qed
lemma simple_function_notspace:
"simple_function M (\<lambda>x. h x * indicator (- space M) x::ennreal)" (is "simple_function M ?h")
proof -
have "?h ` space M \<subseteq> {0}" unfolding indicator_def by auto
hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset)
have "?h -` {0} \<inter> space M = space M" by auto
thus ?thesis unfolding simple_function_def by (auto simp add: image_constant_conv)
qed
lemma simple_function_cong:
assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"
shows "simple_function M f \<longleftrightarrow> simple_function M g"
proof -
have "\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
using assms by auto
with assms show ?thesis
by (simp add: simple_function_def cong: image_cong)
qed
lemma simple_function_cong_algebra:
assumes "sets N = sets M" "space N = space M"
shows "simple_function M f \<longleftrightarrow> simple_function N f"
unfolding simple_function_def assms ..
lemma simple_function_borel_measurable:
fixes f :: "'a \<Rightarrow> 'x::{t2_space}"
assumes "f \<in> borel_measurable M" and "finite (f ` space M)"
shows "simple_function M f"
using assms unfolding simple_function_def
by (auto intro: borel_measurable_vimage)
lemma simple_function_iff_borel_measurable:
fixes f :: "'a \<Rightarrow> 'x::{t2_space}"
shows "simple_function M f \<longleftrightarrow> finite (f ` space M) \<and> f \<in> borel_measurable M"
by (metis borel_measurable_simple_function simple_functionD(1) simple_function_borel_measurable)
lemma simple_function_eq_measurable:
"simple_function M f \<longleftrightarrow> finite (f`space M) \<and> f \<in> measurable M (count_space UNIV)"
using measurable_simple_function[of M f] by (fastforce simp: simple_function_def)
lemma simple_function_const[intro, simp]:
"simple_function M (\<lambda>x. c)"
by (auto intro: finite_subset simp: simple_function_def)
lemma simple_function_compose[intro, simp]:
assumes "simple_function M f"
shows "simple_function M (g \<circ> f)"
unfolding simple_function_def
proof safe
show "finite ((g \<circ> f) ` space M)"
using assms unfolding simple_function_def image_comp [symmetric] by auto
next
fix x assume "x \<in> space M"
let ?G = "g -` {g (f x)} \<inter> (f`space M)"
have *: "(g \<circ> f) -` {(g \<circ> f) x} \<inter> space M =
(\<Union>x\<in>?G. f -` {x} \<inter> space M)" by auto
show "(g \<circ> f) -` {(g \<circ> f) x} \<inter> space M \<in> sets M"
using assms unfolding simple_function_def *
by (rule_tac sets.finite_UN) auto
qed
lemma simple_function_indicator[intro, simp]:
assumes "A \<in> sets M"
shows "simple_function M (indicator A)"
proof -
have "indicator A ` space M \<subseteq> {0, 1}" (is "?S \<subseteq> _")
by (auto simp: indicator_def)
hence "finite ?S" by (rule finite_subset) simp
moreover have "- A \<inter> space M = space M - A" by auto
ultimately show ?thesis unfolding simple_function_def
using assms by (auto simp: indicator_def [abs_def])
qed
lemma simple_function_Pair[intro, simp]:
assumes "simple_function M f"
assumes "simple_function M g"
shows "simple_function M (\<lambda>x. (f x, g x))" (is "simple_function M ?p")
unfolding simple_function_def
proof safe
show "finite (?p ` space M)"
using assms unfolding simple_function_def
by (rule_tac finite_subset[of _ "f`space M \<times> g`space M"]) auto
next
fix x assume "x \<in> space M"
have "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M =
(f -` {f x} \<inter> space M) \<inter> (g -` {g x} \<inter> space M)"
by auto
with \<open>x \<in> space M\<close> show "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M \<in> sets M"
using assms unfolding simple_function_def by auto
qed
lemma simple_function_compose1:
assumes "simple_function M f"
shows "simple_function M (\<lambda>x. g (f x))"
using simple_function_compose[OF assms, of g]
by (simp add: comp_def)
lemma simple_function_compose2:
assumes "simple_function M f" and "simple_function M g"
shows "simple_function M (\<lambda>x. h (f x) (g x))"
proof -
have "simple_function M ((\<lambda>(x, y). h x y) \<circ> (\<lambda>x. (f x, g x)))"
using assms by auto
thus ?thesis by (simp_all add: comp_def)
qed
lemmas simple_function_add[intro, simp] = simple_function_compose2[where h="(+)"]
and simple_function_diff[intro, simp] = simple_function_compose2[where h="(-)"]
and simple_function_uminus[intro, simp] = simple_function_compose[where g="uminus"]
and simple_function_mult[intro, simp] = simple_function_compose2[where h="(*)"]
and simple_function_div[intro, simp] = simple_function_compose2[where h="(/)"]
and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"]
and simple_function_max[intro, simp] = simple_function_compose2[where h=max]
lemma simple_function_sum[intro, simp]:
assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function M (f i)"
shows "simple_function M (\<lambda>x. \<Sum>i\<in>P. f i x)"
proof cases
assume "finite P" from this assms show ?thesis by induct auto
qed auto
lemma simple_function_ennreal[intro, simp]:
fixes f g :: "'a \<Rightarrow> real" assumes sf: "simple_function M f"
shows "simple_function M (\<lambda>x. ennreal (f x))"
by (rule simple_function_compose1[OF sf])
lemma simple_function_real_of_nat[intro, simp]:
fixes f g :: "'a \<Rightarrow> nat" assumes sf: "simple_function M f"
shows "simple_function M (\<lambda>x. real (f x))"
by (rule simple_function_compose1[OF sf])
lemma\<^marker>\<open>tag important\<close> borel_measurable_implies_simple_function_sequence:
fixes u :: "'a \<Rightarrow> ennreal"
assumes u[measurable]: "u \<in> borel_measurable M"
shows "\<exists>f. incseq f \<and> (\<forall>i. (\<forall>x. f i x < top) \<and> simple_function M (f i)) \<and> u = (SUP i. f i)"
proof -
define f where [abs_def]:
"f i x = real_of_int (floor (enn2real (min i (u x)) * 2^i)) / 2^i" for i x
have [simp]: "0 \<le> f i x" for i x
by (auto simp: f_def intro!: divide_nonneg_nonneg mult_nonneg_nonneg enn2real_nonneg)
have *: "2^n * real_of_int x = real_of_int (2^n * x)" for n x
by simp
have "real_of_int \<lfloor>real i * 2 ^ i\<rfloor> = real_of_int \<lfloor>i * 2 ^ i\<rfloor>" for i
by (intro arg_cong[where f=real_of_int]) simp
then have [simp]: "real_of_int \<lfloor>real i * 2 ^ i\<rfloor> = i * 2 ^ i" for i
unfolding floor_of_nat by simp
have "incseq f"
proof (intro monoI le_funI)
fix m n :: nat and x assume "m \<le> n"
moreover
{ fix d :: nat
have "\<lfloor>2^d::real\<rfloor> * \<lfloor>2^m * enn2real (min (of_nat m) (u x))\<rfloor> \<le>
\<lfloor>2^d * (2^m * enn2real (min (of_nat m) (u x)))\<rfloor>"
- by (rule le_mult_floor) (auto simp: enn2real_nonneg)
+ by (rule le_mult_floor) (auto)
also have "\<dots> \<le> \<lfloor>2^d * (2^m * enn2real (min (of_nat d + of_nat m) (u x)))\<rfloor>"
by (intro floor_mono mult_mono enn2real_mono min.mono)
- (auto simp: enn2real_nonneg min_less_iff_disj of_nat_less_top)
+ (auto simp: min_less_iff_disj of_nat_less_top)
finally have "f m x \<le> f (m + d) x"
unfolding f_def
by (auto simp: field_simps power_add * simp del: of_int_mult) }
ultimately show "f m x \<le> f n x"
by (auto simp add: le_iff_add)
qed
then have inc_f: "incseq (\<lambda>i. ennreal (f i x))" for x
by (auto simp: incseq_def le_fun_def)
then have "incseq (\<lambda>i x. ennreal (f i x))"
by (auto simp: incseq_def le_fun_def)
moreover
have "simple_function M (f i)" for i
proof (rule simple_function_borel_measurable)
have "\<lfloor>enn2real (min (of_nat i) (u x)) * 2 ^ i\<rfloor> \<le> \<lfloor>int i * 2 ^ i\<rfloor>" for x
by (cases "u x" rule: ennreal_cases)
(auto split: split_min intro!: floor_mono)
then have "f i ` space M \<subseteq> (\<lambda>n. real_of_int n / 2^i) ` {0 .. of_nat i * 2^i}"
- unfolding floor_of_int by (auto simp: f_def enn2real_nonneg intro!: imageI)
+ unfolding floor_of_int by (auto simp: f_def intro!: imageI)
then show "finite (f i ` space M)"
by (rule finite_subset) auto
show "f i \<in> borel_measurable M"
unfolding f_def enn2real_def by measurable
qed
moreover
{ fix x
have "(SUP i. ennreal (f i x)) = u x"
proof (cases "u x" rule: ennreal_cases)
case top then show ?thesis
by (simp add: f_def inf_min[symmetric] ennreal_of_nat_eq_real_of_nat[symmetric]
ennreal_SUP_of_nat_eq_top)
next
case (real r)
obtain n where "r \<le> of_nat n" using real_arch_simple by auto
then have min_eq_r: "\<forall>\<^sub>F x in sequentially. min (real x) r = r"
by (auto simp: eventually_sequentially intro!: exI[of _ n] split: split_min)
have "(\<lambda>i. real_of_int \<lfloor>min (real i) r * 2^i\<rfloor> / 2^i) \<longlonglongrightarrow> r"
proof (rule tendsto_sandwich)
show "(\<lambda>n. r - (1/2)^n) \<longlonglongrightarrow> r"
by (auto intro!: tendsto_eq_intros LIMSEQ_power_zero)
show "\<forall>\<^sub>F n in sequentially. real_of_int \<lfloor>min (real n) r * 2 ^ n\<rfloor> / 2 ^ n \<le> r"
using min_eq_r by eventually_elim (auto simp: field_simps)
have *: "r * (2 ^ n * 2 ^ n) \<le> 2^n + 2^n * real_of_int \<lfloor>r * 2 ^ n\<rfloor>" for n
using real_of_int_floor_ge_diff_one[of "r * 2^n", THEN mult_left_mono, of "2^n"]
by (auto simp: field_simps)
show "\<forall>\<^sub>F n in sequentially. r - (1/2)^n \<le> real_of_int \<lfloor>min (real n) r * 2 ^ n\<rfloor> / 2 ^ n"
using min_eq_r by eventually_elim (insert *, auto simp: field_simps)
qed auto
then have "(\<lambda>i. ennreal (f i x)) \<longlonglongrightarrow> ennreal r"
by (simp add: real f_def ennreal_of_nat_eq_real_of_nat min_ennreal)
from LIMSEQ_unique[OF LIMSEQ_SUP[OF inc_f] this]
show ?thesis
by (simp add: real)
qed }
ultimately show ?thesis
by (intro exI [of _ "\<lambda>i x. ennreal (f i x)"]) (auto simp add: image_comp)
qed
lemma borel_measurable_implies_simple_function_sequence':
fixes u :: "'a \<Rightarrow> ennreal"
assumes u: "u \<in> borel_measurable M"
obtains f where
"\<And>i. simple_function M (f i)" "incseq f" "\<And>i x. f i x < top" "\<And>x. (SUP i. f i x) = u x"
using borel_measurable_implies_simple_function_sequence [OF u]
by (metis SUP_apply)
lemma\<^marker>\<open>tag important\<close> simple_function_induct
[consumes 1, case_names cong set mult add, induct set: simple_function]:
fixes u :: "'a \<Rightarrow> ennreal"
assumes u: "simple_function M u"
assumes cong: "\<And>f g. simple_function M f \<Longrightarrow> simple_function M g \<Longrightarrow> (AE x in M. f x = g x) \<Longrightarrow> P f \<Longrightarrow> P g"
assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)"
assumes mult: "\<And>u c. P u \<Longrightarrow> P (\<lambda>x. c * u x)"
assumes add: "\<And>u v. P u \<Longrightarrow> P v \<Longrightarrow> P (\<lambda>x. v x + u x)"
shows "P u"
proof (rule cong)
from AE_space show "AE x in M. (\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x) = u x"
proof eventually_elim
fix x assume x: "x \<in> space M"
from simple_function_indicator_representation[OF u x]
show "(\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x) = u x" ..
qed
next
from u have "finite (u ` space M)"
unfolding simple_function_def by auto
then show "P (\<lambda>x. \<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x)"
proof induct
case empty show ?case
using set[of "{}"] by (simp add: indicator_def[abs_def])
qed (auto intro!: add mult set simple_functionD u)
next
show "simple_function M (\<lambda>x. (\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x))"
apply (subst simple_function_cong)
apply (rule simple_function_indicator_representation[symmetric])
apply (auto intro: u)
done
qed fact
lemma simple_function_induct_nn[consumes 1, case_names cong set mult add]:
fixes u :: "'a \<Rightarrow> ennreal"
assumes u: "simple_function M u"
assumes cong: "\<And>f g. simple_function M f \<Longrightarrow> simple_function M g \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> P f \<Longrightarrow> P g"
assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)"
assumes mult: "\<And>u c. simple_function M u \<Longrightarrow> P u \<Longrightarrow> P (\<lambda>x. c * u x)"
assumes add: "\<And>u v. simple_function M u \<Longrightarrow> P u \<Longrightarrow> simple_function M v \<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)"
shows "P u"
proof -
show ?thesis
proof (rule cong)
fix x assume x: "x \<in> space M"
from simple_function_indicator_representation[OF u x]
show "(\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x) = u x" ..
next
show "simple_function M (\<lambda>x. (\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x))"
apply (subst simple_function_cong)
apply (rule simple_function_indicator_representation[symmetric])
apply (auto intro: u)
done
next
from u have "finite (u ` space M)"
unfolding simple_function_def by auto
then show "P (\<lambda>x. \<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x)"
proof induct
case empty show ?case
using set[of "{}"] by (simp add: indicator_def[abs_def])
next
case (insert x S)
{ fix z have "(\<Sum>y\<in>S. y * indicator (u -` {y} \<inter> space M) z) = 0 \<or>
x * indicator (u -` {x} \<inter> space M) z = 0"
using insert by (subst sum_eq_0_iff) (auto simp: indicator_def) }
note disj = this
from insert show ?case
by (auto intro!: add mult set simple_functionD u simple_function_sum disj)
qed
qed fact
qed
lemma\<^marker>\<open>tag important\<close> borel_measurable_induct
[consumes 1, case_names cong set mult add seq, induct set: borel_measurable]:
fixes u :: "'a \<Rightarrow> ennreal"
assumes u: "u \<in> borel_measurable M"
assumes cong: "\<And>f g. f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> P g \<Longrightarrow> P f"
assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)"
assumes mult': "\<And>u c. c < top \<Longrightarrow> u \<in> borel_measurable M \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x < top) \<Longrightarrow> P u \<Longrightarrow> P (\<lambda>x. c * u x)"
assumes add: "\<And>u v. u \<in> borel_measurable M\<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x < top) \<Longrightarrow> P u \<Longrightarrow> v \<in> borel_measurable M \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> v x < top) \<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. x \<in> space M \<Longrightarrow> U i x < top) \<Longrightarrow> (\<And>i. P (U i)) \<Longrightarrow> incseq U \<Longrightarrow> u = (SUP i. U i) \<Longrightarrow> P (SUP i. U i)"
shows "P u"
using u
proof (induct rule: borel_measurable_implies_simple_function_sequence')
fix U assume 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) = u x"
have u_eq: "u = (SUP i. U i)"
using u by (auto simp add: image_comp sup)
have not_inf: "\<And>x i. x \<in> space M \<Longrightarrow> U i x < top"
using U by (auto simp: image_iff eq_commute)
from U have "\<And>i. U i \<in> borel_measurable M"
by (simp add: borel_measurable_simple_function)
show "P u"
unfolding u_eq
proof (rule seq)
fix i show "P (U i)"
using \<open>simple_function M (U i)\<close> not_inf[of _ i]
proof (induct rule: simple_function_induct_nn)
case (mult u c)
show ?case
proof cases
assume "c = 0 \<or> space M = {} \<or> (\<forall>x\<in>space M. u x = 0)"
with mult(1) show ?thesis
by (intro cong[of "\<lambda>x. c * u x" "indicator {}"] set)
(auto dest!: borel_measurable_simple_function)
next
assume "\<not> (c = 0 \<or> space M = {} \<or> (\<forall>x\<in>space M. u x = 0))"
then obtain x where "space M \<noteq> {}" and x: "x \<in> space M" "u x \<noteq> 0" "c \<noteq> 0"
by auto
with mult(3)[of x] have "c < top"
by (auto simp: ennreal_mult_less_top)
then have u_fin: "x' \<in> space M \<Longrightarrow> u x' < top" for x'
using mult(3)[of x'] \<open>c \<noteq> 0\<close> by (auto simp: ennreal_mult_less_top)
then have "P u"
by (rule mult)
with u_fin \<open>c < top\<close> mult(1) show ?thesis
by (intro mult') (auto dest!: borel_measurable_simple_function)
qed
qed (auto intro: cong intro!: set add dest!: borel_measurable_simple_function)
qed fact+
qed
lemma simple_function_If_set:
assumes sf: "simple_function M f" "simple_function M g" and A: "A \<inter> space M \<in> sets M"
shows "simple_function M (\<lambda>x. if x \<in> A then f x else g x)" (is "simple_function M ?IF")
proof -
define F where "F x = f -` {x} \<inter> space M" for x
define G where "G x = g -` {x} \<inter> space M" for x
show ?thesis unfolding simple_function_def
proof safe
have "?IF ` space M \<subseteq> f ` space M \<union> g ` space M" by auto
from finite_subset[OF this] assms
show "finite (?IF ` space M)" unfolding simple_function_def by auto
next
fix x assume "x \<in> space M"
then have *: "?IF -` {?IF x} \<inter> space M = (if x \<in> A
then ((F (f x) \<inter> (A \<inter> space M)) \<union> (G (f x) - (G (f x) \<inter> (A \<inter> space M))))
else ((F (g x) \<inter> (A \<inter> space M)) \<union> (G (g x) - (G (g x) \<inter> (A \<inter> space M)))))"
using sets.sets_into_space[OF A] by (auto split: if_split_asm simp: G_def F_def)
have [intro]: "\<And>x. F x \<in> sets M" "\<And>x. G x \<in> sets M"
unfolding F_def G_def using sf[THEN simple_functionD(2)] by auto
show "?IF -` {?IF x} \<inter> space M \<in> sets M" unfolding * using A by auto
qed
qed
lemma simple_function_If:
assumes sf: "simple_function M f" "simple_function M g" and P: "{x\<in>space M. P x} \<in> sets M"
shows "simple_function M (\<lambda>x. if P x then f x else g x)"
proof -
have "{x\<in>space M. P x} = {x. P x} \<inter> space M" by auto
with simple_function_If_set[OF sf, of "{x. P x}"] P show ?thesis by simp
qed
lemma simple_function_subalgebra:
assumes "simple_function N f"
and N_subalgebra: "sets N \<subseteq> sets M" "space N = space M"
shows "simple_function M f"
using assms unfolding simple_function_def by auto
lemma simple_function_comp:
assumes T: "T \<in> measurable M M'"
and f: "simple_function M' f"
shows "simple_function M (\<lambda>x. f (T x))"
proof (intro simple_function_def[THEN iffD2] conjI ballI)
have "(\<lambda>x. f (T x)) ` space M \<subseteq> f ` space M'"
using T unfolding measurable_def by auto
then show "finite ((\<lambda>x. f (T x)) ` space M)"
using f unfolding simple_function_def by (auto intro: finite_subset)
fix i assume i: "i \<in> (\<lambda>x. f (T x)) ` space M"
then have "i \<in> f ` space M'"
using T unfolding measurable_def by auto
then have "f -` {i} \<inter> space M' \<in> sets M'"
using f unfolding simple_function_def by auto
then have "T -` (f -` {i} \<inter> space M') \<inter> space M \<in> sets M"
using T unfolding measurable_def by auto
also have "T -` (f -` {i} \<inter> space M') \<inter> space M = (\<lambda>x. f (T x)) -` {i} \<inter> space M"
using T unfolding measurable_def by auto
finally show "(\<lambda>x. f (T x)) -` {i} \<inter> space M \<in> sets M" .
qed
subsection "Simple integral"
definition\<^marker>\<open>tag important\<close> simple_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> ennreal" ("integral\<^sup>S") where
"integral\<^sup>S M f = (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M))"
syntax
"_simple_integral" :: "pttrn \<Rightarrow> ennreal \<Rightarrow> 'a measure \<Rightarrow> ennreal" ("\<integral>\<^sup>S _. _ \<partial>_" [60,61] 110)
translations
"\<integral>\<^sup>S x. f \<partial>M" == "CONST simple_integral M (%x. f)"
lemma simple_integral_cong:
assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"
shows "integral\<^sup>S M f = integral\<^sup>S M g"
proof -
have "f ` space M = g ` space M"
"\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
using assms by (auto intro!: image_eqI)
thus ?thesis unfolding simple_integral_def by simp
qed
lemma simple_integral_const[simp]:
"(\<integral>\<^sup>Sx. c \<partial>M) = c * (emeasure M) (space M)"
proof (cases "space M = {}")
case True thus ?thesis unfolding simple_integral_def by simp
next
case False hence "(\<lambda>x. c) ` space M = {c}" by auto
thus ?thesis unfolding simple_integral_def by simp
qed
lemma simple_function_partition:
assumes f: "simple_function 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 "integral\<^sup>S M f = (\<Sum>y\<in>g ` space M. v y * emeasure M {x\<in>space M. g x = y})"
(is "_ = ?r")
proof -
from f g have [simp]: "finite (f`space M)" "finite (g`space M)"
by (auto simp: simple_function_def)
from f g have [measurable]: "f \<in> measurable M (count_space UNIV)" "g \<in> measurable M (count_space UNIV)"
by (auto intro: measurable_simple_function)
{ 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 "integral\<^sup>S M f =
(\<Sum>y\<in>f`space M. y * (\<Sum>z\<in>g`space M.
if \<exists>x\<in>space M. y = f x \<and> z = g x then emeasure M {x\<in>space M. g x = z} else 0))"
unfolding simple_integral_def
proof (safe intro!: sum.cong ennreal_mult_left_cong)
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:"(\<Union>i\<in>{z. \<exists>x\<in>space M. f y = f x \<and> z = g x}. {x \<in> space M. g x = i}) =
f -` {f y} \<inter> space M"
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
then show "emeasure M (f -` {f y} \<inter> space M) =
(\<Sum>z\<in>g ` space M. if \<exists>x\<in>space M. f y = f x \<and> z = g x then emeasure M {x \<in> space M. g x = z} else 0)"
apply (simp add: sum.If_cases)
apply (subst sum_emeasure)
apply (auto simp: disjoint_family_on_def eq)
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 y * emeasure M {x\<in>space M. g x = z} else 0))"
by (auto intro!: sum.cong simp: sum_distrib_left)
also have "\<dots> = ?r"
by (subst sum.swap)
(auto intro!: sum.cong simp: sum.If_cases scaleR_sum_right[symmetric] eq)
finally show "integral\<^sup>S M f = ?r" .
qed
lemma simple_integral_add[simp]:
assumes f: "simple_function M f" and "\<And>x. 0 \<le> f x" and g: "simple_function M g" and "\<And>x. 0 \<le> g x"
shows "(\<integral>\<^sup>Sx. f x + g x \<partial>M) = integral\<^sup>S M f + integral\<^sup>S M g"
proof -
have "(\<integral>\<^sup>Sx. f x + g x \<partial>M) =
(\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. (fst y + snd y) * emeasure M {x\<in>space M. (f x, g x) = y})"
by (intro simple_function_partition) (auto intro: f g)
also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. fst y * emeasure M {x\<in>space M. (f x, g x) = y}) +
(\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. snd y * emeasure M {x\<in>space M. (f x, g x) = y})"
using assms(2,4) by (auto intro!: sum.cong distrib_right simp: sum.distrib[symmetric])
also have "(\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. fst y * emeasure M {x\<in>space M. (f x, g x) = y}) = (\<integral>\<^sup>Sx. f x \<partial>M)"
by (intro simple_function_partition[symmetric]) (auto intro: f g)
also have "(\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. snd y * emeasure M {x\<in>space M. (f x, g x) = y}) = (\<integral>\<^sup>Sx. g x \<partial>M)"
by (intro simple_function_partition[symmetric]) (auto intro: f g)
finally show ?thesis .
qed
lemma simple_integral_sum[simp]:
assumes "\<And>i x. i \<in> P \<Longrightarrow> 0 \<le> f i x"
assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function M (f i)"
shows "(\<integral>\<^sup>Sx. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^sup>S M (f i))"
proof cases
assume "finite P"
from this assms show ?thesis
- by induct (auto simp: simple_function_sum simple_integral_add sum_nonneg)
+ by induct (auto)
qed auto
lemma simple_integral_mult[simp]:
assumes f: "simple_function M f"
shows "(\<integral>\<^sup>Sx. c * f x \<partial>M) = c * integral\<^sup>S M f"
proof -
have "(\<integral>\<^sup>Sx. c * f x \<partial>M) = (\<Sum>y\<in>f ` space M. (c * y) * emeasure M {x\<in>space M. f x = y})"
using f by (intro simple_function_partition) auto
also have "\<dots> = c * integral\<^sup>S M f"
using f unfolding simple_integral_def
by (subst sum_distrib_left) (auto simp: mult.assoc Int_def conj_commute)
finally show ?thesis .
qed
lemma simple_integral_mono_AE:
assumes f[measurable]: "simple_function M f" and g[measurable]: "simple_function M g"
and mono: "AE x in M. f x \<le> g x"
shows "integral\<^sup>S M f \<le> integral\<^sup>S M g"
proof -
let ?\<mu> = "\<lambda>P. emeasure M {x\<in>space M. P x}"
have "integral\<^sup>S M f = (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. fst y * ?\<mu> (\<lambda>x. (f x, g x) = y))"
using f g by (intro simple_function_partition) auto
also have "\<dots> \<le> (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. snd y * ?\<mu> (\<lambda>x. (f x, g x) = y))"
proof (clarsimp intro!: sum_mono)
fix x assume "x \<in> space M"
let ?M = "?\<mu> (\<lambda>y. f y = f x \<and> g y = g x)"
show "f x * ?M \<le> g x * ?M"
proof cases
assume "?M \<noteq> 0"
then have "0 < ?M"
by (simp add: less_le)
also have "\<dots> \<le> ?\<mu> (\<lambda>y. f x \<le> g x)"
using mono by (intro emeasure_mono_AE) auto
finally have "\<not> \<not> f x \<le> g x"
by (intro notI) auto
then show ?thesis
by (intro mult_right_mono) auto
qed simp
qed
also have "\<dots> = integral\<^sup>S M g"
using f g by (intro simple_function_partition[symmetric]) auto
finally show ?thesis .
qed
lemma simple_integral_mono:
assumes "simple_function M f" and "simple_function M g"
and mono: "\<And> x. x \<in> space M \<Longrightarrow> f x \<le> g x"
shows "integral\<^sup>S M f \<le> integral\<^sup>S M g"
using assms by (intro simple_integral_mono_AE) auto
lemma simple_integral_cong_AE:
assumes "simple_function M f" and "simple_function M g"
and "AE x in M. f x = g x"
shows "integral\<^sup>S M f = integral\<^sup>S M g"
using assms by (auto simp: eq_iff intro!: simple_integral_mono_AE)
lemma simple_integral_cong':
assumes sf: "simple_function M f" "simple_function M g"
and mea: "(emeasure M) {x\<in>space M. f x \<noteq> g x} = 0"
shows "integral\<^sup>S M f = integral\<^sup>S M g"
proof (intro simple_integral_cong_AE sf AE_I)
show "(emeasure M) {x\<in>space M. f x \<noteq> g x} = 0" by fact
show "{x \<in> space M. f x \<noteq> g x} \<in> sets M"
using sf[THEN borel_measurable_simple_function] by auto
qed simp
lemma simple_integral_indicator:
assumes A: "A \<in> sets M"
assumes f: "simple_function M f"
shows "(\<integral>\<^sup>Sx. f x * indicator A x \<partial>M) =
(\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M \<inter> A))"
proof -
have eq: "(\<lambda>x. (f x, indicator A x)) ` space M \<inter> {x. snd x = 1} = (\<lambda>x. (f x, 1::ennreal))`A"
using A[THEN sets.sets_into_space] by (auto simp: indicator_def image_iff split: if_split_asm)
have eq2: "\<And>x. f x \<notin> f ` A \<Longrightarrow> f -` {f x} \<inter> space M \<inter> A = {}"
by (auto simp: image_iff)
have "(\<integral>\<^sup>Sx. f x * indicator A x \<partial>M) =
(\<Sum>y\<in>(\<lambda>x. (f x, indicator A x))`space M. (fst y * snd y) * emeasure M {x\<in>space M. (f x, indicator A x) = y})"
using assms by (intro simple_function_partition) auto
also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, indicator A x::ennreal))`space M.
if snd y = 1 then fst y * emeasure M (f -` {fst y} \<inter> space M \<inter> A) else 0)"
by (auto simp: indicator_def split: if_split_asm intro!: arg_cong2[where f="(*)"] arg_cong2[where f=emeasure] sum.cong)
also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, 1::ennreal))`A. fst y * emeasure M (f -` {fst y} \<inter> space M \<inter> A))"
using assms by (subst sum.If_cases) (auto intro!: simple_functionD(1) simp: eq)
also have "\<dots> = (\<Sum>y\<in>fst`(\<lambda>x. (f x, 1::ennreal))`A. y * emeasure M (f -` {y} \<inter> space M \<inter> A))"
by (subst sum.reindex [of fst]) (auto simp: inj_on_def)
also have "\<dots> = (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M \<inter> A))"
using A[THEN sets.sets_into_space]
by (intro sum.mono_neutral_cong_left simple_functionD f) (auto simp: image_comp comp_def eq2)
finally show ?thesis .
qed
lemma simple_integral_indicator_only[simp]:
assumes "A \<in> sets M"
shows "integral\<^sup>S M (indicator A) = emeasure M A"
using simple_integral_indicator[OF assms, of "\<lambda>x. 1"] sets.sets_into_space[OF assms]
by (simp_all add: image_constant_conv Int_absorb1 split: if_split_asm)
lemma simple_integral_null_set:
assumes "simple_function M u" "\<And>x. 0 \<le> u x" and "N \<in> null_sets M"
shows "(\<integral>\<^sup>Sx. u x * indicator N x \<partial>M) = 0"
proof -
have "AE x in M. indicator N x = (0 :: ennreal)"
using \<open>N \<in> null_sets M\<close> by (auto simp: indicator_def intro!: AE_I[of _ _ N])
then have "(\<integral>\<^sup>Sx. u x * indicator N x \<partial>M) = (\<integral>\<^sup>Sx. 0 \<partial>M)"
using assms apply (intro simple_integral_cong_AE) by auto
then show ?thesis by simp
qed
lemma simple_integral_cong_AE_mult_indicator:
assumes sf: "simple_function M f" and eq: "AE x in M. x \<in> S" and "S \<in> sets M"
shows "integral\<^sup>S M f = (\<integral>\<^sup>Sx. f x * indicator S x \<partial>M)"
using assms by (intro simple_integral_cong_AE) auto
lemma simple_integral_cmult_indicator:
assumes A: "A \<in> sets M"
shows "(\<integral>\<^sup>Sx. c * indicator A x \<partial>M) = c * emeasure M A"
using simple_integral_mult[OF simple_function_indicator[OF A]]
unfolding simple_integral_indicator_only[OF A] by simp
lemma simple_integral_nonneg:
assumes f: "simple_function M f" and ae: "AE x in M. 0 \<le> f x"
shows "0 \<le> integral\<^sup>S M f"
proof -
have "integral\<^sup>S M (\<lambda>x. 0) \<le> integral\<^sup>S M f"
using simple_integral_mono_AE[OF _ f ae] by auto
then show ?thesis by simp
qed
subsection \<open>Integral on nonnegative functions\<close>
definition\<^marker>\<open>tag important\<close> nn_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> ennreal" ("integral\<^sup>N") where
"integral\<^sup>N M f = (SUP g \<in> {g. simple_function M g \<and> g \<le> f}. integral\<^sup>S M g)"
syntax
"_nn_integral" :: "pttrn \<Rightarrow> ennreal \<Rightarrow> 'a measure \<Rightarrow> ennreal" ("\<integral>\<^sup>+((2 _./ _)/ \<partial>_)" [60,61] 110)
translations
"\<integral>\<^sup>+x. f \<partial>M" == "CONST nn_integral M (\<lambda>x. f)"
lemma nn_integral_def_finite:
"integral\<^sup>N M f = (SUP g \<in> {g. simple_function M g \<and> g \<le> f \<and> (\<forall>x. g x < top)}. integral\<^sup>S M g)"
(is "_ = Sup (?A ` ?f)")
unfolding nn_integral_def
proof (safe intro!: antisym SUP_least)
fix g assume g[measurable]: "simple_function M g" "g \<le> f"
show "integral\<^sup>S M g \<le> Sup (?A ` ?f)"
proof cases
assume ae: "AE x in M. g x \<noteq> top"
let ?G = "{x \<in> space M. g x \<noteq> top}"
have "integral\<^sup>S M g = integral\<^sup>S M (\<lambda>x. g x * indicator ?G x)"
proof (rule simple_integral_cong_AE)
show "AE x in M. g x = g x * indicator ?G x"
using ae AE_space by eventually_elim auto
qed (insert g, auto)
also have "\<dots> \<le> Sup (?A ` ?f)"
using g by (intro SUP_upper) (auto simp: le_fun_def less_top split: split_indicator)
finally show ?thesis .
next
assume nAE: "\<not> (AE x in M. g x \<noteq> top)"
then have "emeasure M {x\<in>space M. g x = top} \<noteq> 0" (is "emeasure M ?G \<noteq> 0")
by (subst (asm) AE_iff_measurable[OF _ refl]) auto
then have "top = (SUP n. (\<integral>\<^sup>Sx. of_nat n * indicator ?G x \<partial>M))"
by (simp add: ennreal_SUP_of_nat_eq_top ennreal_top_eq_mult_iff SUP_mult_right_ennreal[symmetric])
also have "\<dots> \<le> Sup (?A ` ?f)"
using g
by (safe intro!: SUP_least SUP_upper)
(auto simp: le_fun_def of_nat_less_top top_unique[symmetric] split: split_indicator
intro: order_trans[of _ "g x" "f x" for x, OF order_trans[of _ top]])
finally show ?thesis
by (simp add: top_unique del: SUP_eq_top_iff Sup_eq_top_iff)
qed
qed (auto intro: SUP_upper)
lemma nn_integral_mono_AE:
assumes ae: "AE x in M. u x \<le> v x" shows "integral\<^sup>N M u \<le> integral\<^sup>N M v"
unfolding nn_integral_def
proof (safe intro!: SUP_mono)
fix n assume n: "simple_function M n" "n \<le> u"
from ae[THEN AE_E] guess N . note N = this
then have ae_N: "AE x in M. x \<notin> N" by (auto intro: AE_not_in)
let ?n = "\<lambda>x. n x * indicator (space M - N) x"
have "AE x in M. n x \<le> ?n x" "simple_function M ?n"
using n N ae_N by auto
moreover
{ fix x have "?n x \<le> v x"
proof cases
assume x: "x \<in> space M - N"
with N have "u x \<le> v x" by auto
with n(2)[THEN le_funD, of x] x show ?thesis
by (auto simp: max_def split: if_split_asm)
qed simp }
then have "?n \<le> v" by (auto simp: le_funI)
moreover have "integral\<^sup>S M n \<le> integral\<^sup>S M ?n"
using ae_N N n by (auto intro!: simple_integral_mono_AE)
ultimately show "\<exists>m\<in>{g. simple_function M g \<and> g \<le> v}. integral\<^sup>S M n \<le> integral\<^sup>S M m"
by force
qed
lemma nn_integral_mono:
"(\<And>x. x \<in> space M \<Longrightarrow> u x \<le> v x) \<Longrightarrow> integral\<^sup>N M u \<le> integral\<^sup>N M v"
by (auto intro: nn_integral_mono_AE)
lemma mono_nn_integral: "mono F \<Longrightarrow> mono (\<lambda>x. integral\<^sup>N M (F x))"
by (auto simp add: mono_def le_fun_def intro!: nn_integral_mono)
lemma nn_integral_cong_AE:
"AE x in M. u x = v x \<Longrightarrow> integral\<^sup>N M u = integral\<^sup>N M v"
by (auto simp: eq_iff intro!: nn_integral_mono_AE)
lemma nn_integral_cong:
"(\<And>x. x \<in> space M \<Longrightarrow> u x = v x) \<Longrightarrow> integral\<^sup>N M u = integral\<^sup>N M v"
by (auto intro: nn_integral_cong_AE)
lemma nn_integral_cong_simp:
"(\<And>x. x \<in> space M =simp=> u x = v x) \<Longrightarrow> integral\<^sup>N M u = integral\<^sup>N M v"
by (auto intro: nn_integral_cong simp: simp_implies_def)
lemma incseq_nn_integral:
assumes "incseq f" shows "incseq (\<lambda>i. integral\<^sup>N M (f i))"
proof -
have "\<And>i x. f i x \<le> f (Suc i) x"
using assms by (auto dest!: incseq_SucD simp: le_fun_def)
then show ?thesis
by (auto intro!: incseq_SucI nn_integral_mono)
qed
lemma nn_integral_eq_simple_integral:
assumes f: "simple_function M f" shows "integral\<^sup>N M f = integral\<^sup>S M f"
proof -
let ?f = "\<lambda>x. f x * indicator (space M) x"
have f': "simple_function M ?f" using f by auto
have "integral\<^sup>N M ?f \<le> integral\<^sup>S M ?f" using f'
by (force intro!: SUP_least simple_integral_mono simp: le_fun_def nn_integral_def)
moreover have "integral\<^sup>S M ?f \<le> integral\<^sup>N M ?f"
unfolding nn_integral_def
using f' by (auto intro!: SUP_upper)
ultimately show ?thesis
by (simp cong: nn_integral_cong simple_integral_cong)
qed
text \<open>Beppo-Levi monotone convergence theorem\<close>
lemma nn_integral_monotone_convergence_SUP:
assumes f: "incseq f" and [measurable]: "\<And>i. f i \<in> borel_measurable M"
shows "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^sup>N M (f i))"
proof (rule antisym)
show "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) \<le> (SUP i. (\<integral>\<^sup>+ x. f i x \<partial>M))"
unfolding nn_integral_def_finite[of _ "\<lambda>x. SUP i. f i x"]
proof (safe intro!: SUP_least)
fix u assume sf_u[simp]: "simple_function M u" and
u: "u \<le> (\<lambda>x. SUP i. f i x)" and u_range: "\<forall>x. u x < top"
note sf_u[THEN borel_measurable_simple_function, measurable]
show "integral\<^sup>S M u \<le> (SUP j. \<integral>\<^sup>+x. f j x \<partial>M)"
proof (rule ennreal_approx_unit)
fix a :: ennreal assume "a < 1"
let ?au = "\<lambda>x. a * u x"
let ?B = "\<lambda>c i. {x\<in>space M. ?au x = c \<and> c \<le> f i x}"
have "integral\<^sup>S M ?au = (\<Sum>c\<in>?au`space M. c * (SUP i. emeasure M (?B c i)))"
unfolding simple_integral_def
proof (intro sum.cong ennreal_mult_left_cong refl)
fix c assume "c \<in> ?au ` space M" "c \<noteq> 0"
{ fix x' assume x': "x' \<in> space M" "?au x' = c"
with \<open>c \<noteq> 0\<close> u_range have "?au x' < 1 * u x'"
by (intro ennreal_mult_strict_right_mono \<open>a < 1\<close>) (auto simp: less_le)
also have "\<dots> \<le> (SUP i. f i x')"
using u by (auto simp: le_fun_def)
finally have "\<exists>i. ?au x' \<le> f i x'"
by (auto simp: less_SUP_iff intro: less_imp_le) }
then have *: "?au -` {c} \<inter> space M = (\<Union>i. ?B c i)"
by auto
show "emeasure M (?au -` {c} \<inter> space M) = (SUP i. emeasure M (?B c i))"
unfolding * using f
by (intro SUP_emeasure_incseq[symmetric])
(auto simp: incseq_def le_fun_def intro: order_trans)
qed
also have "\<dots> = (SUP i. \<Sum>c\<in>?au`space M. c * emeasure M (?B c i))"
unfolding SUP_mult_left_ennreal using f
by (intro ennreal_SUP_sum[symmetric])
(auto intro!: mult_mono emeasure_mono simp: incseq_def le_fun_def intro: order_trans)
also have "\<dots> \<le> (SUP i. integral\<^sup>N M (f i))"
proof (intro SUP_subset_mono order_refl)
fix i
have "(\<Sum>c\<in>?au`space M. c * emeasure M (?B c i)) =
(\<integral>\<^sup>Sx. (a * u x) * indicator {x\<in>space M. a * u x \<le> f i x} x \<partial>M)"
by (subst simple_integral_indicator)
(auto intro!: sum.cong ennreal_mult_left_cong arg_cong2[where f=emeasure])
also have "\<dots> = (\<integral>\<^sup>+x. (a * u x) * indicator {x\<in>space M. a * u x \<le> f i x} x \<partial>M)"
by (rule nn_integral_eq_simple_integral[symmetric]) simp
also have "\<dots> \<le> (\<integral>\<^sup>+x. f i x \<partial>M)"
by (intro nn_integral_mono) (auto split: split_indicator)
finally show "(\<Sum>c\<in>?au`space M. c * emeasure M (?B c i)) \<le> (\<integral>\<^sup>+x. f i x \<partial>M)" .
qed
finally show "a * integral\<^sup>S M u \<le> (SUP i. integral\<^sup>N M (f i))"
by simp
qed
qed
qed (auto intro!: SUP_least SUP_upper nn_integral_mono)
lemma sup_continuous_nn_integral[order_continuous_intros]:
assumes f: "\<And>y. sup_continuous (f y)"
assumes [measurable]: "\<And>x. (\<lambda>y. f y x) \<in> borel_measurable M"
shows "sup_continuous (\<lambda>x. (\<integral>\<^sup>+y. f y x \<partial>M))"
unfolding sup_continuous_def
proof safe
fix C :: "nat \<Rightarrow> 'b" assume C: "incseq C"
with sup_continuous_mono[OF f] show "(\<integral>\<^sup>+ y. f y (Sup (C ` UNIV)) \<partial>M) = (SUP i. \<integral>\<^sup>+ y. f y (C i) \<partial>M)"
unfolding sup_continuousD[OF f C]
by (subst nn_integral_monotone_convergence_SUP) (auto simp: mono_def le_fun_def)
qed
theorem nn_integral_monotone_convergence_SUP_AE:
assumes f: "\<And>i. AE x in M. f i x \<le> f (Suc i) x" "\<And>i. f i \<in> borel_measurable M"
shows "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^sup>N M (f i))"
proof -
from f have "AE x in M. \<forall>i. f i x \<le> f (Suc i) x"
by (simp add: AE_all_countable)
from this[THEN AE_E] guess N . note N = this
let ?f = "\<lambda>i x. if x \<in> space M - N then f i x else 0"
have f_eq: "AE x in M. \<forall>i. ?f i x = f i x" using N by (auto intro!: AE_I[of _ _ N])
then have "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (\<integral>\<^sup>+ x. (SUP i. ?f i x) \<partial>M)"
by (auto intro!: nn_integral_cong_AE)
also have "\<dots> = (SUP i. (\<integral>\<^sup>+ x. ?f i x \<partial>M))"
proof (rule nn_integral_monotone_convergence_SUP)
show "incseq ?f" using N(1) by (force intro!: incseq_SucI le_funI)
{ fix i show "(\<lambda>x. if x \<in> space M - N then f i x else 0) \<in> borel_measurable M"
using f N(3) by (intro measurable_If_set) auto }
qed
also have "\<dots> = (SUP i. (\<integral>\<^sup>+ x. f i x \<partial>M))"
using f_eq by (force intro!: arg_cong[where f = "\<lambda>f. Sup (range f)"] nn_integral_cong_AE ext)
finally show ?thesis .
qed
lemma nn_integral_monotone_convergence_simple:
"incseq f \<Longrightarrow> (\<And>i. simple_function M (f i)) \<Longrightarrow> (SUP i. \<integral>\<^sup>S x. f i x \<partial>M) = (\<integral>\<^sup>+x. (SUP i. f i x) \<partial>M)"
using nn_integral_monotone_convergence_SUP[of f M]
by (simp add: nn_integral_eq_simple_integral[symmetric] borel_measurable_simple_function)
lemma SUP_simple_integral_sequences:
assumes f: "incseq f" "\<And>i. simple_function M (f i)"
and g: "incseq g" "\<And>i. simple_function M (g i)"
and eq: "AE x in M. (SUP i. f i x) = (SUP i. g i x)"
shows "(SUP i. integral\<^sup>S M (f i)) = (SUP i. integral\<^sup>S M (g i))"
(is "Sup (?F ` _) = Sup (?G ` _)")
proof -
have "(SUP i. integral\<^sup>S M (f i)) = (\<integral>\<^sup>+x. (SUP i. f i x) \<partial>M)"
using f by (rule nn_integral_monotone_convergence_simple)
also have "\<dots> = (\<integral>\<^sup>+x. (SUP i. g i x) \<partial>M)"
unfolding eq[THEN nn_integral_cong_AE] ..
also have "\<dots> = (SUP i. ?G i)"
using g by (rule nn_integral_monotone_convergence_simple[symmetric])
finally show ?thesis by simp
qed
lemma nn_integral_const[simp]: "(\<integral>\<^sup>+ x. c \<partial>M) = c * emeasure M (space M)"
by (subst nn_integral_eq_simple_integral) auto
lemma nn_integral_linear:
assumes f: "f \<in> borel_measurable M" and g: "g \<in> borel_measurable M"
shows "(\<integral>\<^sup>+ x. a * f x + g x \<partial>M) = a * integral\<^sup>N M f + integral\<^sup>N M g"
(is "integral\<^sup>N M ?L = _")
proof -
from borel_measurable_implies_simple_function_sequence'[OF f(1)] guess u .
note u = nn_integral_monotone_convergence_simple[OF this(2,1)] this
from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess v .
note v = nn_integral_monotone_convergence_simple[OF this(2,1)] this
let ?L' = "\<lambda>i x. a * u i x + v i x"
have "?L \<in> borel_measurable M" using assms by auto
from borel_measurable_implies_simple_function_sequence'[OF this] guess l .
note l = nn_integral_monotone_convergence_simple[OF this(2,1)] this
have inc: "incseq (\<lambda>i. a * integral\<^sup>S M (u i))" "incseq (\<lambda>i. integral\<^sup>S M (v i))"
using u v by (auto simp: incseq_Suc_iff le_fun_def intro!: add_mono mult_left_mono simple_integral_mono)
have l': "(SUP i. integral\<^sup>S M (l i)) = (SUP i. integral\<^sup>S M (?L' i))"
proof (rule SUP_simple_integral_sequences[OF l(3,2)])
show "incseq ?L'" "\<And>i. simple_function M (?L' i)"
using u v unfolding incseq_Suc_iff le_fun_def
by (auto intro!: add_mono mult_left_mono)
{ fix x
have "(SUP i. a * u i x + v i x) = a * (SUP i. u i x) + (SUP i. v i x)"
using u(3) v(3) u(4)[of _ x] v(4)[of _ x] unfolding SUP_mult_left_ennreal
by (auto intro!: ennreal_SUP_add simp: incseq_Suc_iff le_fun_def add_mono mult_left_mono) }
then show "AE x in M. (SUP i. l i x) = (SUP i. ?L' i x)"
unfolding l(5) using u(5) v(5) by (intro AE_I2) auto
qed
also have "\<dots> = (SUP i. a * integral\<^sup>S M (u i) + integral\<^sup>S M (v i))"
using u(2) v(2) by auto
finally show ?thesis
unfolding l(5)[symmetric] l(1)[symmetric]
by (simp add: ennreal_SUP_add[OF inc] v u SUP_mult_left_ennreal[symmetric])
qed
lemma nn_integral_cmult: "f \<in> borel_measurable M \<Longrightarrow> (\<integral>\<^sup>+ x. c * f x \<partial>M) = c * integral\<^sup>N M f"
using nn_integral_linear[of f M "\<lambda>x. 0" c] by simp
lemma nn_integral_multc: "f \<in> borel_measurable M \<Longrightarrow> (\<integral>\<^sup>+ x. f x * c \<partial>M) = integral\<^sup>N M f * c"
unfolding mult.commute[of _ c] nn_integral_cmult by simp
lemma nn_integral_divide: "f \<in> borel_measurable M \<Longrightarrow> (\<integral>\<^sup>+ x. f x / c \<partial>M) = (\<integral>\<^sup>+ x. f x \<partial>M) / c"
unfolding divide_ennreal_def by (rule nn_integral_multc)
lemma nn_integral_indicator[simp]: "A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+ x. indicator A x\<partial>M) = (emeasure M) A"
by (subst nn_integral_eq_simple_integral) (auto simp: simple_integral_indicator)
lemma nn_integral_cmult_indicator: "A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+ x. c * indicator A x \<partial>M) = c * emeasure M A"
- by (subst nn_integral_eq_simple_integral)
- (auto simp: simple_function_indicator simple_integral_indicator)
+ by (subst nn_integral_eq_simple_integral) (auto)
lemma nn_integral_indicator':
assumes [measurable]: "A \<inter> space M \<in> sets M"
shows "(\<integral>\<^sup>+ x. indicator A x \<partial>M) = emeasure M (A \<inter> space M)"
proof -
have "(\<integral>\<^sup>+ x. indicator A x \<partial>M) = (\<integral>\<^sup>+ x. indicator (A \<inter> space M) x \<partial>M)"
by (intro nn_integral_cong) (simp split: split_indicator)
also have "\<dots> = emeasure M (A \<inter> space M)"
by simp
finally show ?thesis .
qed
lemma nn_integral_indicator_singleton[simp]:
assumes [measurable]: "{y} \<in> sets M" shows "(\<integral>\<^sup>+x. f x * indicator {y} x \<partial>M) = f y * emeasure M {y}"
proof -
have "(\<integral>\<^sup>+x. f x * indicator {y} x \<partial>M) = (\<integral>\<^sup>+x. f y * indicator {y} x \<partial>M)"
by (auto intro!: nn_integral_cong split: split_indicator)
then show ?thesis
by (simp add: nn_integral_cmult)
qed
lemma nn_integral_set_ennreal:
"(\<integral>\<^sup>+x. ennreal (f x) * indicator A x \<partial>M) = (\<integral>\<^sup>+x. ennreal (f x * indicator A x) \<partial>M)"
by (rule nn_integral_cong) (simp split: split_indicator)
lemma nn_integral_indicator_singleton'[simp]:
assumes [measurable]: "{y} \<in> sets M"
shows "(\<integral>\<^sup>+x. ennreal (f x * indicator {y} x) \<partial>M) = f y * emeasure M {y}"
- by (subst nn_integral_set_ennreal[symmetric]) (simp add: nn_integral_indicator_singleton)
+ by (subst nn_integral_set_ennreal[symmetric]) (simp)
lemma nn_integral_add:
"f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<integral>\<^sup>+ x. f x + g x \<partial>M) = integral\<^sup>N M f + integral\<^sup>N M g"
using nn_integral_linear[of f M g 1] by simp
lemma nn_integral_sum:
"(\<And>i. i \<in> P \<Longrightarrow> f i \<in> borel_measurable M) \<Longrightarrow> (\<integral>\<^sup>+ x. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^sup>N M (f i))"
by (induction P rule: infinite_finite_induct) (auto simp: nn_integral_add)
theorem nn_integral_suminf:
assumes f: "\<And>i. f i \<in> borel_measurable M"
shows "(\<integral>\<^sup>+ x. (\<Sum>i. f i x) \<partial>M) = (\<Sum>i. integral\<^sup>N M (f i))"
proof -
have all_pos: "AE x in M. \<forall>i. 0 \<le> f i x"
using assms by (auto simp: AE_all_countable)
have "(\<Sum>i. integral\<^sup>N M (f i)) = (SUP n. \<Sum>i<n. integral\<^sup>N M (f i))"
by (rule suminf_eq_SUP)
also have "\<dots> = (SUP n. \<integral>\<^sup>+x. (\<Sum>i<n. f i x) \<partial>M)"
unfolding nn_integral_sum[OF f] ..
also have "\<dots> = \<integral>\<^sup>+x. (SUP n. \<Sum>i<n. f i x) \<partial>M" using f all_pos
by (intro nn_integral_monotone_convergence_SUP_AE[symmetric])
(elim AE_mp, auto simp: sum_nonneg simp del: sum.lessThan_Suc intro!: AE_I2 sum_mono2)
also have "\<dots> = \<integral>\<^sup>+x. (\<Sum>i. f i x) \<partial>M" using all_pos
by (intro nn_integral_cong_AE) (auto simp: suminf_eq_SUP)
finally show ?thesis by simp
qed
lemma nn_integral_bound_simple_function:
assumes bnd: "\<And>x. x \<in> space M \<Longrightarrow> f x < \<infinity>"
assumes f[measurable]: "simple_function M f"
assumes supp: "emeasure M {x\<in>space M. f x \<noteq> 0} < \<infinity>"
shows "nn_integral M f < \<infinity>"
proof cases
assume "space M = {}"
then have "nn_integral M f = (\<integral>\<^sup>+x. 0 \<partial>M)"
by (intro nn_integral_cong) auto
then show ?thesis by simp
next
assume "space M \<noteq> {}"
with simple_functionD(1)[OF f] bnd have bnd: "0 \<le> Max (f`space M) \<and> Max (f`space M) < \<infinity>"
by (subst Max_less_iff) (auto simp: Max_ge_iff)
have "nn_integral M f \<le> (\<integral>\<^sup>+x. Max (f`space M) * indicator {x\<in>space M. f x \<noteq> 0} x \<partial>M)"
proof (rule nn_integral_mono)
fix x assume "x \<in> space M"
with f show "f x \<le> Max (f ` space M) * indicator {x \<in> space M. f x \<noteq> 0} x"
by (auto split: split_indicator intro!: Max_ge simple_functionD)
qed
also have "\<dots> < \<infinity>"
using bnd supp by (subst nn_integral_cmult) (auto simp: ennreal_mult_less_top)
finally show ?thesis .
qed
theorem nn_integral_Markov_inequality:
assumes u: "u \<in> borel_measurable M" and "A \<in> sets M"
shows "(emeasure M) ({x\<in>space M. 1 \<le> c * u x} \<inter> A) \<le> c * (\<integral>\<^sup>+ x. u x * indicator A x \<partial>M)"
(is "(emeasure M) ?A \<le> _ * ?PI")
proof -
have "?A \<in> sets M"
using \<open>A \<in> sets M\<close> u by auto
hence "(emeasure M) ?A = (\<integral>\<^sup>+ x. indicator ?A x \<partial>M)"
using nn_integral_indicator by simp
also have "\<dots> \<le> (\<integral>\<^sup>+ x. c * (u x * indicator A x) \<partial>M)"
using u by (auto intro!: nn_integral_mono_AE simp: indicator_def)
also have "\<dots> = c * (\<integral>\<^sup>+ x. u x * indicator A x \<partial>M)"
using assms by (auto intro!: nn_integral_cmult)
finally show ?thesis .
qed
lemma nn_integral_noteq_infinite:
assumes g: "g \<in> borel_measurable M" and "integral\<^sup>N M g \<noteq> \<infinity>"
shows "AE x in M. g x \<noteq> \<infinity>"
proof (rule ccontr)
assume c: "\<not> (AE x in M. g x \<noteq> \<infinity>)"
have "(emeasure M) {x\<in>space M. g x = \<infinity>} \<noteq> 0"
using c g by (auto simp add: AE_iff_null)
then have "0 < (emeasure M) {x\<in>space M. g x = \<infinity>}"
by (auto simp: zero_less_iff_neq_zero)
then have "\<infinity> = \<infinity> * (emeasure M) {x\<in>space M. g x = \<infinity>}"
by (auto simp: ennreal_top_eq_mult_iff)
also have "\<dots> \<le> (\<integral>\<^sup>+x. \<infinity> * indicator {x\<in>space M. g x = \<infinity>} x \<partial>M)"
using g by (subst nn_integral_cmult_indicator) auto
also have "\<dots> \<le> integral\<^sup>N M g"
using assms by (auto intro!: nn_integral_mono_AE simp: indicator_def)
finally show False
using \<open>integral\<^sup>N M g \<noteq> \<infinity>\<close> by (auto simp: top_unique)
qed
lemma nn_integral_PInf:
assumes f: "f \<in> borel_measurable M" and not_Inf: "integral\<^sup>N M f \<noteq> \<infinity>"
shows "emeasure M (f -` {\<infinity>} \<inter> space M) = 0"
proof -
have "\<infinity> * emeasure M (f -` {\<infinity>} \<inter> space M) = (\<integral>\<^sup>+ x. \<infinity> * indicator (f -` {\<infinity>} \<inter> space M) x \<partial>M)"
using f by (subst nn_integral_cmult_indicator) (auto simp: measurable_sets)
also have "\<dots> \<le> integral\<^sup>N M f"
by (auto intro!: nn_integral_mono simp: indicator_def)
finally have "\<infinity> * (emeasure M) (f -` {\<infinity>} \<inter> space M) \<le> integral\<^sup>N M f"
by simp
then show ?thesis
using assms by (auto simp: ennreal_top_mult top_unique split: if_split_asm)
qed
lemma simple_integral_PInf:
"simple_function M f \<Longrightarrow> integral\<^sup>S M f \<noteq> \<infinity> \<Longrightarrow> emeasure M (f -` {\<infinity>} \<inter> space M) = 0"
by (rule nn_integral_PInf) (auto simp: nn_integral_eq_simple_integral borel_measurable_simple_function)
lemma nn_integral_PInf_AE:
assumes "f \<in> borel_measurable M" "integral\<^sup>N M f \<noteq> \<infinity>" shows "AE x in M. f x \<noteq> \<infinity>"
proof (rule AE_I)
show "(emeasure M) (f -` {\<infinity>} \<inter> space M) = 0"
by (rule nn_integral_PInf[OF assms])
show "f -` {\<infinity>} \<inter> space M \<in> sets M"
using assms by (auto intro: borel_measurable_vimage)
qed auto
lemma nn_integral_diff:
assumes f: "f \<in> borel_measurable M"
and g: "g \<in> borel_measurable M"
and fin: "integral\<^sup>N M g \<noteq> \<infinity>"
and mono: "AE x in M. g x \<le> f x"
shows "(\<integral>\<^sup>+ x. f x - g x \<partial>M) = integral\<^sup>N M f - integral\<^sup>N M g"
proof -
have diff: "(\<lambda>x. f x - g x) \<in> borel_measurable M"
using assms by auto
have "AE x in M. f x = f x - g x + g x"
using diff_add_cancel_ennreal mono nn_integral_noteq_infinite[OF g fin] assms by auto
then have **: "integral\<^sup>N M f = (\<integral>\<^sup>+x. f x - g x \<partial>M) + integral\<^sup>N M g"
unfolding nn_integral_add[OF diff g, symmetric]
by (rule nn_integral_cong_AE)
show ?thesis unfolding **
using fin
by (cases rule: ennreal2_cases[of "\<integral>\<^sup>+ x. f x - g x \<partial>M" "integral\<^sup>N M g"]) auto
qed
lemma nn_integral_mult_bounded_inf:
assumes f: "f \<in> borel_measurable M" "(\<integral>\<^sup>+x. f x \<partial>M) < \<infinity>" and c: "c \<noteq> \<infinity>" and ae: "AE x in M. g x \<le> c * f x"
shows "(\<integral>\<^sup>+x. g x \<partial>M) < \<infinity>"
proof -
have "(\<integral>\<^sup>+x. g x \<partial>M) \<le> (\<integral>\<^sup>+x. c * f x \<partial>M)"
by (intro nn_integral_mono_AE ae)
also have "(\<integral>\<^sup>+x. c * f x \<partial>M) < \<infinity>"
using c f by (subst nn_integral_cmult) (auto simp: ennreal_mult_less_top top_unique not_less)
finally show ?thesis .
qed
text \<open>Fatou's lemma: convergence theorem on limes inferior\<close>
lemma nn_integral_monotone_convergence_INF_AE':
assumes f: "\<And>i. AE x in M. f (Suc i) x \<le> f i x" and [measurable]: "\<And>i. f i \<in> borel_measurable M"
and *: "(\<integral>\<^sup>+ x. f 0 x \<partial>M) < \<infinity>"
shows "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (INF i. integral\<^sup>N M (f i))"
proof (rule ennreal_minus_cancel)
have "integral\<^sup>N M (f 0) - (\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (\<integral>\<^sup>+x. f 0 x - (INF i. f i x) \<partial>M)"
proof (rule nn_integral_diff[symmetric])
have "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) \<le> (\<integral>\<^sup>+ x. f 0 x \<partial>M)"
by (intro nn_integral_mono INF_lower) simp
with * show "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) \<noteq> \<infinity>"
by simp
qed (auto intro: INF_lower)
also have "\<dots> = (\<integral>\<^sup>+x. (SUP i. f 0 x - f i x) \<partial>M)"
by (simp add: ennreal_INF_const_minus)
also have "\<dots> = (SUP i. (\<integral>\<^sup>+x. f 0 x - f i x \<partial>M))"
proof (intro nn_integral_monotone_convergence_SUP_AE)
show "AE x in M. f 0 x - f i x \<le> f 0 x - f (Suc i) x" for i
using f[of i] by eventually_elim (auto simp: ennreal_mono_minus)
qed simp
also have "\<dots> = (SUP i. nn_integral M (f 0) - (\<integral>\<^sup>+x. f i x \<partial>M))"
proof (subst nn_integral_diff[symmetric])
fix i
have dec: "AE x in M. \<forall>i. f (Suc i) x \<le> f i x"
unfolding AE_all_countable using f by auto
then show "AE x in M. f i x \<le> f 0 x"
using dec by eventually_elim (auto intro: lift_Suc_antimono_le[of "\<lambda>i. f i x" 0 i for x])
then have "(\<integral>\<^sup>+ x. f i x \<partial>M) \<le> (\<integral>\<^sup>+ x. f 0 x \<partial>M)"
by (rule nn_integral_mono_AE)
with * show "(\<integral>\<^sup>+ x. f i x \<partial>M) \<noteq> \<infinity>"
by simp
qed (insert f, auto simp: decseq_def le_fun_def)
finally show "integral\<^sup>N M (f 0) - (\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) =
integral\<^sup>N M (f 0) - (INF i. \<integral>\<^sup>+ x. f i x \<partial>M)"
by (simp add: ennreal_INF_const_minus)
qed (insert *, auto intro!: nn_integral_mono intro: INF_lower)
theorem nn_integral_monotone_convergence_INF_AE:
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ennreal"
assumes f: "\<And>i. AE x in M. f (Suc i) x \<le> f i x"
and [measurable]: "\<And>i. f i \<in> borel_measurable M"
and fin: "(\<integral>\<^sup>+ x. f i x \<partial>M) < \<infinity>"
shows "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (INF i. integral\<^sup>N M (f i))"
proof -
{ fix f :: "nat \<Rightarrow> ennreal" and j assume "decseq f"
then have "(INF i. f i) = (INF i. f (i + j))"
apply (intro INF_eq)
apply (rule_tac x="i" in bexI)
apply (auto simp: decseq_def le_fun_def)
done }
note INF_shift = this
have mono: "AE x in M. \<forall>i. f (Suc i) x \<le> f i x"
using f by (auto simp: AE_all_countable)
then have "AE x in M. (INF i. f i x) = (INF n. f (n + i) x)"
by eventually_elim (auto intro!: decseq_SucI INF_shift)
then have "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (\<integral>\<^sup>+ x. (INF n. f (n + i) x) \<partial>M)"
by (rule nn_integral_cong_AE)
also have "\<dots> = (INF n. (\<integral>\<^sup>+ x. f (n + i) x \<partial>M))"
by (rule nn_integral_monotone_convergence_INF_AE') (insert assms, auto)
also have "\<dots> = (INF n. (\<integral>\<^sup>+ x. f n x \<partial>M))"
by (intro INF_shift[symmetric] decseq_SucI nn_integral_mono_AE f)
finally show ?thesis .
qed
lemma nn_integral_monotone_convergence_INF_decseq:
assumes f: "decseq f" and *: "\<And>i. f i \<in> borel_measurable M" "(\<integral>\<^sup>+ x. f i x \<partial>M) < \<infinity>"
shows "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (INF i. integral\<^sup>N M (f i))"
using nn_integral_monotone_convergence_INF_AE[of f M i, OF _ *] f by (auto simp: decseq_Suc_iff le_fun_def)
theorem nn_integral_liminf:
fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ennreal"
assumes u: "\<And>i. u i \<in> borel_measurable M"
shows "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) \<le> liminf (\<lambda>n. integral\<^sup>N M (u n))"
proof -
have "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) = (SUP n. \<integral>\<^sup>+ x. (INF i\<in>{n..}. u i x) \<partial>M)"
unfolding liminf_SUP_INF using u
by (intro nn_integral_monotone_convergence_SUP_AE)
(auto intro!: AE_I2 intro: INF_greatest INF_superset_mono)
also have "\<dots> \<le> liminf (\<lambda>n. integral\<^sup>N M (u n))"
by (auto simp: liminf_SUP_INF intro!: SUP_mono INF_greatest nn_integral_mono INF_lower)
finally show ?thesis .
qed
theorem nn_integral_limsup:
fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ennreal"
assumes [measurable]: "\<And>i. u i \<in> borel_measurable M" "w \<in> borel_measurable M"
assumes bounds: "\<And>i. AE x in M. u i x \<le> w x" and w: "(\<integral>\<^sup>+x. w x \<partial>M) < \<infinity>"
shows "limsup (\<lambda>n. integral\<^sup>N M (u n)) \<le> (\<integral>\<^sup>+ x. limsup (\<lambda>n. u n x) \<partial>M)"
proof -
have bnd: "AE x in M. \<forall>i. u i x \<le> w x"
using bounds by (auto simp: AE_all_countable)
then have "(\<integral>\<^sup>+ x. (SUP n. u n x) \<partial>M) \<le> (\<integral>\<^sup>+ x. w x \<partial>M)"
by (auto intro!: nn_integral_mono_AE elim: eventually_mono intro: SUP_least)
then have "(\<integral>\<^sup>+ x. limsup (\<lambda>n. u n x) \<partial>M) = (INF n. \<integral>\<^sup>+ x. (SUP i\<in>{n..}. u i x) \<partial>M)"
unfolding limsup_INF_SUP using bnd w
by (intro nn_integral_monotone_convergence_INF_AE')
(auto intro!: AE_I2 intro: SUP_least SUP_subset_mono)
also have "\<dots> \<ge> limsup (\<lambda>n. integral\<^sup>N M (u n))"
by (auto simp: limsup_INF_SUP intro!: INF_mono SUP_least exI nn_integral_mono SUP_upper)
finally (xtrans) show ?thesis .
qed
lemma nn_integral_LIMSEQ:
assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M"
and u: "\<And>x. (\<lambda>i. f i x) \<longlonglongrightarrow> u x"
shows "(\<lambda>n. integral\<^sup>N M (f n)) \<longlonglongrightarrow> integral\<^sup>N M u"
proof -
have "(\<lambda>n. integral\<^sup>N M (f n)) \<longlonglongrightarrow> (SUP n. integral\<^sup>N M (f n))"
using f by (intro LIMSEQ_SUP[of "\<lambda>n. integral\<^sup>N M (f n)"] incseq_nn_integral)
also have "(SUP n. integral\<^sup>N M (f n)) = integral\<^sup>N M (\<lambda>x. SUP n. f n x)"
using f by (intro nn_integral_monotone_convergence_SUP[symmetric])
also have "integral\<^sup>N M (\<lambda>x. SUP n. f n x) = integral\<^sup>N M (\<lambda>x. u x)"
using f by (subst LIMSEQ_SUP[THEN LIMSEQ_unique, OF _ u]) (auto simp: incseq_def le_fun_def)
finally show ?thesis .
qed
theorem nn_integral_dominated_convergence:
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. 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. u i x \<partial>M)) \<longlonglongrightarrow> (\<integral>\<^sup>+x. u' x \<partial>M)"
proof -
have "limsup (\<lambda>n. integral\<^sup>N M (u n)) \<le> (\<integral>\<^sup>+ x. limsup (\<lambda>n. u n x) \<partial>M)"
by (intro nn_integral_limsup[OF _ _ bound w]) auto
moreover have "(\<integral>\<^sup>+ x. limsup (\<lambda>n. u n x) \<partial>M) = (\<integral>\<^sup>+ x. u' x \<partial>M)"
using u' by (intro nn_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot)
moreover have "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) = (\<integral>\<^sup>+ x. u' x \<partial>M)"
using u' by (intro nn_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot)
moreover have "(\<integral>\<^sup>+x. liminf (\<lambda>n. u n x) \<partial>M) \<le> liminf (\<lambda>n. integral\<^sup>N M (u n))"
by (intro nn_integral_liminf) auto
moreover have "liminf (\<lambda>n. integral\<^sup>N M (u n)) \<le> limsup (\<lambda>n. integral\<^sup>N M (u n))"
by (intro Liminf_le_Limsup sequentially_bot)
ultimately show ?thesis
by (intro Liminf_eq_Limsup) auto
qed
lemma inf_continuous_nn_integral[order_continuous_intros]:
assumes f: "\<And>y. inf_continuous (f y)"
assumes [measurable]: "\<And>x. (\<lambda>y. f y x) \<in> borel_measurable M"
assumes bnd: "\<And>x. (\<integral>\<^sup>+ y. f y x \<partial>M) \<noteq> \<infinity>"
shows "inf_continuous (\<lambda>x. (\<integral>\<^sup>+y. f y x \<partial>M))"
unfolding inf_continuous_def
proof safe
fix C :: "nat \<Rightarrow> 'b" assume C: "decseq C"
then show "(\<integral>\<^sup>+ y. f y (Inf (C ` UNIV)) \<partial>M) = (INF i. \<integral>\<^sup>+ y. f y (C i) \<partial>M)"
using inf_continuous_mono[OF f] bnd
by (auto simp add: inf_continuousD[OF f C] fun_eq_iff antimono_def mono_def le_fun_def less_top
intro!: nn_integral_monotone_convergence_INF_decseq)
qed
lemma nn_integral_null_set:
assumes "N \<in> null_sets M" shows "(\<integral>\<^sup>+ x. u x * indicator N x \<partial>M) = 0"
proof -
have "(\<integral>\<^sup>+ x. u x * indicator N x \<partial>M) = (\<integral>\<^sup>+ x. 0 \<partial>M)"
proof (intro nn_integral_cong_AE AE_I)
show "{x \<in> space M. u x * indicator N x \<noteq> 0} \<subseteq> N"
by (auto simp: indicator_def)
show "(emeasure M) N = 0" "N \<in> sets M"
using assms by auto
qed
then show ?thesis by simp
qed
lemma nn_integral_0_iff:
assumes u: "u \<in> borel_measurable M"
shows "integral\<^sup>N M u = 0 \<longleftrightarrow> emeasure M {x\<in>space M. u x \<noteq> 0} = 0"
(is "_ \<longleftrightarrow> (emeasure M) ?A = 0")
proof -
have u_eq: "(\<integral>\<^sup>+ x. u x * indicator ?A x \<partial>M) = integral\<^sup>N M u"
by (auto intro!: nn_integral_cong simp: indicator_def)
show ?thesis
proof
assume "(emeasure M) ?A = 0"
with nn_integral_null_set[of ?A M u] u
show "integral\<^sup>N M u = 0" by (simp add: u_eq null_sets_def)
next
assume *: "integral\<^sup>N M u = 0"
let ?M = "\<lambda>n. {x \<in> space M. 1 \<le> real (n::nat) * u x}"
have "0 = (SUP n. (emeasure M) (?M n \<inter> ?A))"
proof -
{ fix n :: nat
from nn_integral_Markov_inequality[OF u, of ?A "of_nat n"] u
have "(emeasure M) (?M n \<inter> ?A) \<le> 0"
by (simp add: ennreal_of_nat_eq_real_of_nat u_eq *)
moreover have "0 \<le> (emeasure M) (?M n \<inter> ?A)" using u by auto
ultimately have "(emeasure M) (?M n \<inter> ?A) = 0" by auto }
thus ?thesis by simp
qed
also have "\<dots> = (emeasure M) (\<Union>n. ?M n \<inter> ?A)"
proof (safe intro!: SUP_emeasure_incseq)
fix n show "?M n \<inter> ?A \<in> sets M"
using u by (auto intro!: sets.Int)
next
show "incseq (\<lambda>n. {x \<in> space M. 1 \<le> real n * u x} \<inter> {x \<in> space M. u x \<noteq> 0})"
proof (safe intro!: incseq_SucI)
fix n :: nat and x
assume *: "1 \<le> real n * u x"
also have "real n * u x \<le> real (Suc n) * u x"
by (auto intro!: mult_right_mono)
finally show "1 \<le> real (Suc n) * u x" by auto
qed
qed
also have "\<dots> = (emeasure M) {x\<in>space M. 0 < u x}"
proof (safe intro!: arg_cong[where f="(emeasure M)"])
fix x assume "0 < u x" and [simp, intro]: "x \<in> space M"
show "x \<in> (\<Union>n. ?M n \<inter> ?A)"
proof (cases "u x" rule: ennreal_cases)
case (real r) with \<open>0 < u x\<close> have "0 < r" by auto
obtain j :: nat where "1 / r \<le> real j" using real_arch_simple ..
hence "1 / r * r \<le> real j * r" unfolding mult_le_cancel_right using \<open>0 < r\<close> by auto
hence "1 \<le> real j * r" using real \<open>0 < r\<close> by auto
thus ?thesis using \<open>0 < r\<close> real
by (auto simp: ennreal_of_nat_eq_real_of_nat ennreal_1[symmetric] ennreal_mult[symmetric]
simp del: ennreal_1)
qed (insert \<open>0 < u x\<close>, auto simp: ennreal_mult_top)
qed (auto simp: zero_less_iff_neq_zero)
finally show "emeasure M ?A = 0"
by (simp add: zero_less_iff_neq_zero)
qed
qed
lemma nn_integral_0_iff_AE:
assumes u: "u \<in> borel_measurable M"
shows "integral\<^sup>N M u = 0 \<longleftrightarrow> (AE x in M. u x = 0)"
proof -
have sets: "{x\<in>space M. u x \<noteq> 0} \<in> sets M"
using u by auto
show "integral\<^sup>N M u = 0 \<longleftrightarrow> (AE x in M. u x = 0)"
using nn_integral_0_iff[of u] AE_iff_null[OF sets] u by auto
qed
lemma AE_iff_nn_integral:
"{x\<in>space M. P x} \<in> sets M \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> integral\<^sup>N M (indicator {x. \<not> P x}) = 0"
by (subst nn_integral_0_iff_AE) (auto simp: indicator_def[abs_def])
lemma nn_integral_less:
assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
assumes f: "(\<integral>\<^sup>+x. f x \<partial>M) \<noteq> \<infinity>"
assumes ord: "AE x in M. f x \<le> g x" "\<not> (AE x in M. g x \<le> f x)"
shows "(\<integral>\<^sup>+x. f x \<partial>M) < (\<integral>\<^sup>+x. g x \<partial>M)"
proof -
have "0 < (\<integral>\<^sup>+x. g x - f x \<partial>M)"
proof (intro order_le_neq_trans notI)
assume "0 = (\<integral>\<^sup>+x. g x - f x \<partial>M)"
then have "AE x in M. g x - f x = 0"
using nn_integral_0_iff_AE[of "\<lambda>x. g x - f x" M] by simp
with ord(1) have "AE x in M. g x \<le> f x"
by eventually_elim (auto simp: ennreal_minus_eq_0)
with ord show False
by simp
qed simp
also have "\<dots> = (\<integral>\<^sup>+x. g x \<partial>M) - (\<integral>\<^sup>+x. f x \<partial>M)"
using f by (subst nn_integral_diff) (auto simp: ord)
finally show ?thesis
using f by (auto dest!: ennreal_minus_pos_iff[rotated] simp: less_top)
qed
lemma nn_integral_subalgebra:
assumes f: "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>N N f = integral\<^sup>N M f"
proof -
have [simp]: "\<And>f :: 'a \<Rightarrow> ennreal. f \<in> borel_measurable N \<Longrightarrow> f \<in> borel_measurable M"
using N by (auto simp: measurable_def)
have [simp]: "\<And>P. (AE x in N. P x) \<Longrightarrow> (AE x in M. P x)"
using N by (auto simp add: eventually_ae_filter null_sets_def subset_eq)
have [simp]: "\<And>A. A \<in> sets N \<Longrightarrow> A \<in> sets M"
using N by auto
from f show ?thesis
apply induct
apply (simp_all add: nn_integral_add nn_integral_cmult nn_integral_monotone_convergence_SUP N image_comp)
apply (auto intro!: nn_integral_cong cong: nn_integral_cong simp: N(2)[symmetric])
done
qed
lemma nn_integral_nat_function:
fixes f :: "'a \<Rightarrow> nat"
assumes "f \<in> measurable M (count_space UNIV)"
shows "(\<integral>\<^sup>+x. of_nat (f x) \<partial>M) = (\<Sum>t. emeasure M {x\<in>space M. t < f x})"
proof -
define F where "F i = {x\<in>space M. i < f x}" for i
with assms have [measurable]: "\<And>i. F i \<in> sets M"
by auto
{ fix x assume "x \<in> space M"
have "(\<lambda>i. if i < f x then 1 else 0) sums (of_nat (f x)::real)"
using sums_If_finite[of "\<lambda>i. i < f x" "\<lambda>_. 1::real"] by simp
then have "(\<lambda>i. ennreal (if i < f x then 1 else 0)) sums of_nat(f x)"
unfolding ennreal_of_nat_eq_real_of_nat
by (subst sums_ennreal) auto
moreover have "\<And>i. ennreal (if i < f x then 1 else 0) = indicator (F i) x"
using \<open>x \<in> space M\<close> by (simp add: one_ennreal_def F_def)
ultimately have "of_nat (f x) = (\<Sum>i. indicator (F i) x :: ennreal)"
by (simp add: sums_iff) }
then have "(\<integral>\<^sup>+x. of_nat (f x) \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>i. indicator (F i) x) \<partial>M)"
by (simp cong: nn_integral_cong)
also have "\<dots> = (\<Sum>i. emeasure M (F i))"
by (simp add: nn_integral_suminf)
finally show ?thesis
by (simp add: F_def)
qed
theorem nn_integral_lfp:
assumes sets[simp]: "\<And>s. sets (M s) = sets N"
assumes f: "sup_continuous f"
assumes g: "sup_continuous g"
assumes meas: "\<And>F. F \<in> borel_measurable N \<Longrightarrow> f F \<in> borel_measurable N"
assumes step: "\<And>F s. F \<in> borel_measurable N \<Longrightarrow> integral\<^sup>N (M s) (f F) = g (\<lambda>s. integral\<^sup>N (M s) F) s"
shows "(\<integral>\<^sup>+\<omega>. lfp f \<omega> \<partial>M s) = lfp g s"
proof (subst lfp_transfer_bounded[where \<alpha>="\<lambda>F s. \<integral>\<^sup>+x. F x \<partial>M s" and g=g and f=f and P="\<lambda>f. f \<in> borel_measurable N", symmetric])
fix C :: "nat \<Rightarrow> 'b \<Rightarrow> ennreal" assume "incseq C" "\<And>i. C i \<in> borel_measurable N"
then show "(\<lambda>s. \<integral>\<^sup>+x. (SUP i. C i) x \<partial>M s) = (SUP i. (\<lambda>s. \<integral>\<^sup>+x. C i x \<partial>M s))"
unfolding SUP_apply[abs_def]
by (subst nn_integral_monotone_convergence_SUP)
(auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure] cong: measurable_cong_sets)
qed (auto simp add: step le_fun_def SUP_apply[abs_def] bot_fun_def bot_ennreal intro!: meas f g)
theorem nn_integral_gfp:
assumes sets[simp]: "\<And>s. sets (M s) = sets N"
assumes f: "inf_continuous f" and g: "inf_continuous g"
assumes meas: "\<And>F. F \<in> borel_measurable N \<Longrightarrow> f F \<in> borel_measurable N"
assumes bound: "\<And>F s. F \<in> borel_measurable N \<Longrightarrow> (\<integral>\<^sup>+x. f F x \<partial>M s) < \<infinity>"
assumes non_zero: "\<And>s. emeasure (M s) (space (M s)) \<noteq> 0"
assumes step: "\<And>F s. F \<in> borel_measurable N \<Longrightarrow> integral\<^sup>N (M s) (f F) = g (\<lambda>s. integral\<^sup>N (M s) F) s"
shows "(\<integral>\<^sup>+\<omega>. gfp f \<omega> \<partial>M s) = gfp g s"
proof (subst gfp_transfer_bounded[where \<alpha>="\<lambda>F s. \<integral>\<^sup>+x. F x \<partial>M s" and g=g and f=f
and P="\<lambda>F. F \<in> borel_measurable N \<and> (\<forall>s. (\<integral>\<^sup>+x. F x \<partial>M s) < \<infinity>)", symmetric])
fix C :: "nat \<Rightarrow> 'b \<Rightarrow> ennreal" assume "decseq C" "\<And>i. C i \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) (C i) < \<infinity>)"
then show "(\<lambda>s. \<integral>\<^sup>+x. (INF i. C i) x \<partial>M s) = (INF i. (\<lambda>s. \<integral>\<^sup>+x. C i x \<partial>M s))"
unfolding INF_apply[abs_def]
by (subst nn_integral_monotone_convergence_INF_decseq)
(auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure] cong: measurable_cong_sets)
next
show "\<And>x. g x \<le> (\<lambda>s. integral\<^sup>N (M s) (f top))"
by (subst step)
(auto simp add: top_fun_def less_le non_zero le_fun_def ennreal_top_mult
cong del: if_weak_cong intro!: monoD[OF inf_continuous_mono[OF g], THEN le_funD])
next
fix C assume "\<And>i::nat. C i \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) (C i) < \<infinity>)" "decseq C"
with bound show "Inf (C ` UNIV) \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) (Inf (C ` UNIV)) < \<infinity>)"
unfolding INF_apply[abs_def]
by (subst nn_integral_monotone_convergence_INF_decseq)
(auto simp: INF_less_iff cong: measurable_cong_sets intro!: borel_measurable_INF)
next
show "\<And>x. x \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) x < \<infinity>) \<Longrightarrow>
(\<lambda>s. integral\<^sup>N (M s) (f x)) = g (\<lambda>s. integral\<^sup>N (M s) x)"
by (subst step) auto
qed (insert bound, auto simp add: le_fun_def INF_apply[abs_def] top_fun_def intro!: meas f g)
(* TODO: rename? *)
subsection \<open>Integral under concrete measures\<close>
lemma nn_integral_mono_measure:
assumes "sets M = sets N" "M \<le> N" shows "nn_integral M f \<le> nn_integral N f"
unfolding nn_integral_def
proof (intro SUP_subset_mono)
note \<open>sets M = sets N\<close>[simp] \<open>sets M = sets N\<close>[THEN sets_eq_imp_space_eq, simp]
show "{g. simple_function M g \<and> g \<le> f} \<subseteq> {g. simple_function N g \<and> g \<le> f}"
by (simp add: simple_function_def)
show "integral\<^sup>S M x \<le> integral\<^sup>S N x" for x
using le_measureD3[OF \<open>M \<le> N\<close>]
by (auto simp add: simple_integral_def intro!: sum_mono mult_mono)
qed
lemma nn_integral_empty:
assumes "space M = {}"
shows "nn_integral M f = 0"
proof -
have "(\<integral>\<^sup>+ x. f x \<partial>M) = (\<integral>\<^sup>+ x. 0 \<partial>M)"
by(rule nn_integral_cong)(simp add: assms)
thus ?thesis by simp
qed
lemma nn_integral_bot[simp]: "nn_integral bot f = 0"
by (simp add: nn_integral_empty)
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Distributions\<close>
lemma nn_integral_distr:
assumes T: "T \<in> measurable M M'" and f: "f \<in> borel_measurable (distr M M' T)"
shows "integral\<^sup>N (distr M M' T) f = (\<integral>\<^sup>+ x. f (T x) \<partial>M)"
using f
proof induct
case (cong f g)
with T show ?case
apply (subst nn_integral_cong[of _ f g])
apply simp
apply (subst nn_integral_cong[of _ "\<lambda>x. f (T x)" "\<lambda>x. g (T x)"])
apply (simp add: measurable_def Pi_iff)
apply simp
done
next
case (set A)
then have eq: "\<And>x. x \<in> space M \<Longrightarrow> indicator A (T x) = indicator (T -` A \<inter> space M) x"
by (auto simp: indicator_def)
from set T show ?case
by (subst nn_integral_cong[OF eq])
(auto simp add: emeasure_distr intro!: nn_integral_indicator[symmetric] measurable_sets)
qed (simp_all add: measurable_compose[OF T] T nn_integral_cmult nn_integral_add
nn_integral_monotone_convergence_SUP le_fun_def incseq_def image_comp)
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Counting space\<close>
lemma simple_function_count_space[simp]:
"simple_function (count_space A) f \<longleftrightarrow> finite (f ` A)"
unfolding simple_function_def by simp
lemma nn_integral_count_space:
assumes A: "finite {a\<in>A. 0 < f a}"
shows "integral\<^sup>N (count_space A) f = (\<Sum>a|a\<in>A \<and> 0 < f a. f a)"
proof -
have *: "(\<integral>\<^sup>+x. max 0 (f x) \<partial>count_space A) =
(\<integral>\<^sup>+ x. (\<Sum>a|a\<in>A \<and> 0 < f a. f a * indicator {a} x) \<partial>count_space A)"
by (auto intro!: nn_integral_cong
simp add: indicator_def if_distrib sum.If_cases[OF A] max_def le_less)
also have "\<dots> = (\<Sum>a|a\<in>A \<and> 0 < f a. \<integral>\<^sup>+ x. f a * indicator {a} x \<partial>count_space A)"
by (subst nn_integral_sum) (simp_all add: AE_count_space less_imp_le)
also have "\<dots> = (\<Sum>a|a\<in>A \<and> 0 < f a. f a)"
by (auto intro!: sum.cong simp: one_ennreal_def[symmetric] max_def)
finally show ?thesis by (simp add: max.absorb2)
qed
lemma nn_integral_count_space_finite:
"finite A \<Longrightarrow> (\<integral>\<^sup>+x. f x \<partial>count_space A) = (\<Sum>a\<in>A. f a)"
by (auto intro!: sum.mono_neutral_left simp: nn_integral_count_space less_le)
lemma nn_integral_count_space':
assumes "finite A" "\<And>x. x \<in> B \<Longrightarrow> x \<notin> A \<Longrightarrow> f x = 0" "A \<subseteq> B"
shows "(\<integral>\<^sup>+x. f x \<partial>count_space B) = (\<Sum>x\<in>A. f x)"
proof -
have "(\<integral>\<^sup>+x. f x \<partial>count_space B) = (\<Sum>a | a \<in> B \<and> 0 < f a. f a)"
using assms(2,3)
by (intro nn_integral_count_space finite_subset[OF _ \<open>finite A\<close>]) (auto simp: less_le)
also have "\<dots> = (\<Sum>a\<in>A. f a)"
using assms by (intro sum.mono_neutral_cong_left) (auto simp: less_le)
finally show ?thesis .
qed
lemma nn_integral_bij_count_space:
assumes g: "bij_betw g A B"
shows "(\<integral>\<^sup>+x. f (g x) \<partial>count_space A) = (\<integral>\<^sup>+x. f x \<partial>count_space B)"
using g[THEN bij_betw_imp_funcset]
by (subst distr_bij_count_space[OF g, symmetric])
(auto intro!: nn_integral_distr[symmetric])
lemma nn_integral_indicator_finite:
fixes f :: "'a \<Rightarrow> ennreal"
assumes f: "finite A" and [measurable]: "\<And>a. a \<in> A \<Longrightarrow> {a} \<in> sets M"
shows "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = (\<Sum>x\<in>A. f x * emeasure M {x})"
proof -
from f have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>a\<in>A. f a * indicator {a} x) \<partial>M)"
by (intro nn_integral_cong) (auto simp: indicator_def if_distrib[where f="\<lambda>a. x * a" for x] sum.If_cases)
also have "\<dots> = (\<Sum>a\<in>A. f a * emeasure M {a})"
by (subst nn_integral_sum) auto
finally show ?thesis .
qed
lemma nn_integral_count_space_nat:
fixes f :: "nat \<Rightarrow> ennreal"
shows "(\<integral>\<^sup>+i. f i \<partial>count_space UNIV) = (\<Sum>i. f i)"
proof -
have "(\<integral>\<^sup>+i. f i \<partial>count_space UNIV) =
(\<integral>\<^sup>+i. (\<Sum>j. f j * indicator {j} i) \<partial>count_space UNIV)"
proof (intro nn_integral_cong)
fix i
have "f i = (\<Sum>j\<in>{i}. f j * indicator {j} i)"
by simp
also have "\<dots> = (\<Sum>j. f j * indicator {j} i)"
by (rule suminf_finite[symmetric]) auto
finally show "f i = (\<Sum>j. f j * indicator {j} i)" .
qed
also have "\<dots> = (\<Sum>j. (\<integral>\<^sup>+i. f j * indicator {j} i \<partial>count_space UNIV))"
by (rule nn_integral_suminf) auto
finally show ?thesis
by simp
qed
lemma nn_integral_enat_function:
assumes f: "f \<in> measurable M (count_space UNIV)"
shows "(\<integral>\<^sup>+ x. ennreal_of_enat (f x) \<partial>M) = (\<Sum>t. emeasure M {x \<in> space M. t < f x})"
proof -
define F where "F i = {x\<in>space M. i < f x}" for i :: nat
with assms have [measurable]: "\<And>i. F i \<in> sets M"
by auto
{ fix x assume "x \<in> space M"
have "(\<lambda>i::nat. if i < f x then 1 else 0) sums ennreal_of_enat (f x)"
using sums_If_finite[of "\<lambda>r. r < f x" "\<lambda>_. 1 :: ennreal"]
by (cases "f x") (simp_all add: sums_def of_nat_tendsto_top_ennreal)
also have "(\<lambda>i. (if i < f x then 1 else 0)) = (\<lambda>i. indicator (F i) x)"
using \<open>x \<in> space M\<close> by (simp add: one_ennreal_def F_def fun_eq_iff)
finally have "ennreal_of_enat (f x) = (\<Sum>i. indicator (F i) x)"
by (simp add: sums_iff) }
then have "(\<integral>\<^sup>+x. ennreal_of_enat (f x) \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>i. indicator (F i) x) \<partial>M)"
by (simp cong: nn_integral_cong)
also have "\<dots> = (\<Sum>i. emeasure M (F i))"
by (simp add: nn_integral_suminf)
finally show ?thesis
by (simp add: F_def)
qed
lemma nn_integral_count_space_nn_integral:
fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> ennreal"
assumes "countable I" and [measurable]: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable M"
shows "(\<integral>\<^sup>+x. \<integral>\<^sup>+i. f i x \<partial>count_space I \<partial>M) = (\<integral>\<^sup>+i. \<integral>\<^sup>+x. f i x \<partial>M \<partial>count_space I)"
proof cases
assume "finite I" then show ?thesis
by (simp add: nn_integral_count_space_finite nn_integral_sum)
next
assume "infinite I"
then have [simp]: "I \<noteq> {}"
by auto
note * = bij_betw_from_nat_into[OF \<open>countable I\<close> \<open>infinite I\<close>]
have **: "\<And>f. (\<And>i. 0 \<le> f i) \<Longrightarrow> (\<integral>\<^sup>+i. f i \<partial>count_space I) = (\<Sum>n. f (from_nat_into I n))"
by (simp add: nn_integral_bij_count_space[symmetric, OF *] nn_integral_count_space_nat)
show ?thesis
by (simp add: ** nn_integral_suminf from_nat_into)
qed
lemma of_bool_Bex_eq_nn_integral:
assumes unique: "\<And>x y. x \<in> X \<Longrightarrow> y \<in> X \<Longrightarrow> P x \<Longrightarrow> P y \<Longrightarrow> x = y"
shows "of_bool (\<exists>y\<in>X. P y) = (\<integral>\<^sup>+y. of_bool (P y) \<partial>count_space X)"
proof cases
assume "\<exists>y\<in>X. P y"
then obtain y where "P y" "y \<in> X" by auto
then show ?thesis
by (subst nn_integral_count_space'[where A="{y}"]) (auto dest: unique)
qed (auto cong: nn_integral_cong_simp)
lemma emeasure_UN_countable:
assumes sets[measurable]: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> sets M" and I[simp]: "countable I"
assumes disj: "disjoint_family_on X I"
shows "emeasure M (\<Union>(X ` I)) = (\<integral>\<^sup>+i. emeasure M (X i) \<partial>count_space I)"
proof -
have eq: "\<And>x. indicator (\<Union>(X ` I)) x = \<integral>\<^sup>+ i. indicator (X i) x \<partial>count_space I"
proof cases
fix x assume x: "x \<in> \<Union>(X ` I)"
then obtain j where j: "x \<in> X j" "j \<in> I"
by auto
with disj have "\<And>i. i \<in> I \<Longrightarrow> indicator (X i) x = (indicator {j} i::ennreal)"
by (auto simp: disjoint_family_on_def split: split_indicator)
with x j show "?thesis x"
by (simp cong: nn_integral_cong_simp)
qed (auto simp: nn_integral_0_iff_AE)
note sets.countable_UN'[unfolded subset_eq, measurable]
have "emeasure M (\<Union>(X ` I)) = (\<integral>\<^sup>+x. indicator (\<Union>(X ` I)) x \<partial>M)"
by simp
also have "\<dots> = (\<integral>\<^sup>+i. \<integral>\<^sup>+x. indicator (X i) x \<partial>M \<partial>count_space I)"
by (simp add: eq nn_integral_count_space_nn_integral)
finally show ?thesis
by (simp cong: nn_integral_cong_simp)
qed
lemma emeasure_countable_singleton:
assumes sets: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M" and X: "countable X"
shows "emeasure M X = (\<integral>\<^sup>+x. emeasure M {x} \<partial>count_space X)"
proof -
have "emeasure M (\<Union>i\<in>X. {i}) = (\<integral>\<^sup>+x. emeasure M {x} \<partial>count_space X)"
using assms by (intro emeasure_UN_countable) (auto simp: disjoint_family_on_def)
also have "(\<Union>i\<in>X. {i}) = X" by auto
finally show ?thesis .
qed
lemma measure_eqI_countable:
assumes [simp]: "sets M = Pow A" "sets N = Pow A" and A: "countable A"
assumes eq: "\<And>a. a \<in> A \<Longrightarrow> emeasure M {a} = emeasure N {a}"
shows "M = N"
proof (rule measure_eqI)
fix X assume "X \<in> sets M"
then have X: "X \<subseteq> A" by auto
moreover from A X have "countable X" by (auto dest: countable_subset)
ultimately have
"emeasure M X = (\<integral>\<^sup>+a. emeasure M {a} \<partial>count_space X)"
"emeasure N X = (\<integral>\<^sup>+a. emeasure N {a} \<partial>count_space X)"
by (auto intro!: emeasure_countable_singleton)
moreover have "(\<integral>\<^sup>+a. emeasure M {a} \<partial>count_space X) = (\<integral>\<^sup>+a. emeasure N {a} \<partial>count_space X)"
using X by (intro nn_integral_cong eq) auto
ultimately show "emeasure M X = emeasure N X"
by simp
qed simp
lemma measure_eqI_countable_AE:
assumes [simp]: "sets M = UNIV" "sets N = UNIV"
assumes ae: "AE x in M. x \<in> \<Omega>" "AE x in N. x \<in> \<Omega>" and [simp]: "countable \<Omega>"
assumes eq: "\<And>x. x \<in> \<Omega> \<Longrightarrow> emeasure M {x} = emeasure N {x}"
shows "M = N"
proof (rule measure_eqI)
fix A
have "emeasure N A = emeasure N {x\<in>\<Omega>. x \<in> A}"
using ae by (intro emeasure_eq_AE) auto
also have "\<dots> = (\<integral>\<^sup>+x. emeasure N {x} \<partial>count_space {x\<in>\<Omega>. x \<in> A})"
by (intro emeasure_countable_singleton) auto
also have "\<dots> = (\<integral>\<^sup>+x. emeasure M {x} \<partial>count_space {x\<in>\<Omega>. x \<in> A})"
by (intro nn_integral_cong eq[symmetric]) auto
also have "\<dots> = emeasure M {x\<in>\<Omega>. x \<in> A}"
by (intro emeasure_countable_singleton[symmetric]) auto
also have "\<dots> = emeasure M A"
using ae by (intro emeasure_eq_AE) auto
finally show "emeasure M A = emeasure N A" ..
qed simp
lemma nn_integral_monotone_convergence_SUP_nat:
fixes f :: "'a \<Rightarrow> nat \<Rightarrow> ennreal"
assumes chain: "Complete_Partial_Order.chain (\<le>) (f ` Y)"
and nonempty: "Y \<noteq> {}"
shows "(\<integral>\<^sup>+ x. (SUP i\<in>Y. f i x) \<partial>count_space UNIV) = (SUP i\<in>Y. (\<integral>\<^sup>+ x. f i x \<partial>count_space UNIV))"
(is "?lhs = ?rhs" is "integral\<^sup>N ?M _ = _")
proof (rule order_class.order.antisym)
show "?rhs \<le> ?lhs"
by (auto intro!: SUP_least SUP_upper nn_integral_mono)
next
have "\<exists>g. incseq g \<and> range g \<subseteq> (\<lambda>i. f i x) ` Y \<and> (SUP i\<in>Y. f i x) = (SUP i. g i)" for x
by (rule ennreal_Sup_countable_SUP) (simp add: nonempty)
then obtain g where incseq: "\<And>x. incseq (g x)"
and range: "\<And>x. range (g x) \<subseteq> (\<lambda>i. f i x) ` Y"
and sup: "\<And>x. (SUP i\<in>Y. f i x) = (SUP i. g x i)" by moura
from incseq have incseq': "incseq (\<lambda>i x. g x i)"
by(blast intro: incseq_SucI le_funI dest: incseq_SucD)
have "?lhs = \<integral>\<^sup>+ x. (SUP i. g x i) \<partial>?M" by(simp add: sup)
also have "\<dots> = (SUP i. \<integral>\<^sup>+ x. g x i \<partial>?M)" using incseq'
by(rule nn_integral_monotone_convergence_SUP) simp
also have "\<dots> \<le> (SUP i\<in>Y. \<integral>\<^sup>+ x. f i x \<partial>?M)"
proof(rule SUP_least)
fix n
have "\<And>x. \<exists>i. g x n = f i x \<and> i \<in> Y" using range by blast
then obtain I where I: "\<And>x. g x n = f (I x) x" "\<And>x. I x \<in> Y" by moura
have "(\<integral>\<^sup>+ x. g x n \<partial>count_space UNIV) = (\<Sum>x. g x n)"
by(rule nn_integral_count_space_nat)
also have "\<dots> = (SUP m. \<Sum>x<m. g x n)"
by(rule suminf_eq_SUP)
also have "\<dots> \<le> (SUP i\<in>Y. \<integral>\<^sup>+ x. f i x \<partial>?M)"
proof(rule SUP_mono)
fix m
show "\<exists>m'\<in>Y. (\<Sum>x<m. g x n) \<le> (\<integral>\<^sup>+ x. f m' x \<partial>?M)"
proof(cases "m > 0")
case False
thus ?thesis using nonempty by auto
next
case True
let ?Y = "I ` {..<m}"
have "f ` ?Y \<subseteq> f ` Y" using I by auto
with chain have chain': "Complete_Partial_Order.chain (\<le>) (f ` ?Y)" by(rule chain_subset)
hence "Sup (f ` ?Y) \<in> f ` ?Y"
by(rule ccpo_class.in_chain_finite)(auto simp add: True lessThan_empty_iff)
then obtain m' where "m' < m" and m': "(SUP i\<in>?Y. f i) = f (I m')" by auto
have "I m' \<in> Y" using I by blast
have "(\<Sum>x<m. g x n) \<le> (\<Sum>x<m. f (I m') x)"
proof(rule sum_mono)
fix x
assume "x \<in> {..<m}"
hence "x < m" by simp
have "g x n = f (I x) x" by(simp add: I)
also have "\<dots> \<le> (SUP i\<in>?Y. f i) x" unfolding Sup_fun_def image_image
using \<open>x \<in> {..<m}\<close> by (rule Sup_upper [OF imageI])
also have "\<dots> = f (I m') x" unfolding m' by simp
finally show "g x n \<le> f (I m') x" .
qed
also have "\<dots> \<le> (SUP m. (\<Sum>x<m. f (I m') x))"
by(rule SUP_upper) simp
also have "\<dots> = (\<Sum>x. f (I m') x)"
by(rule suminf_eq_SUP[symmetric])
also have "\<dots> = (\<integral>\<^sup>+ x. f (I m') x \<partial>?M)"
by(rule nn_integral_count_space_nat[symmetric])
finally show ?thesis using \<open>I m' \<in> Y\<close> by blast
qed
qed
finally show "(\<integral>\<^sup>+ x. g x n \<partial>count_space UNIV) \<le> \<dots>" .
qed
finally show "?lhs \<le> ?rhs" .
qed
lemma power_series_tendsto_at_left:
assumes nonneg: "\<And>i. 0 \<le> f i" and summable: "\<And>z. 0 \<le> z \<Longrightarrow> z < 1 \<Longrightarrow> summable (\<lambda>n. f n * z^n)"
shows "((\<lambda>z. ennreal (\<Sum>n. f n * z^n)) \<longlongrightarrow> (\<Sum>n. ennreal (f n))) (at_left (1::real))"
proof (intro tendsto_at_left_sequentially)
show "0 < (1::real)" by simp
fix S :: "nat \<Rightarrow> real" assume S: "\<And>n. S n < 1" "\<And>n. 0 < S n" "S \<longlonglongrightarrow> 1" "incseq S"
then have S_nonneg: "\<And>i. 0 \<le> S i" by (auto intro: less_imp_le)
have "(\<lambda>i. (\<integral>\<^sup>+n. f n * S i^n \<partial>count_space UNIV)) \<longlonglongrightarrow> (\<integral>\<^sup>+n. ennreal (f n) \<partial>count_space UNIV)"
proof (rule nn_integral_LIMSEQ)
show "incseq (\<lambda>i n. ennreal (f n * S i^n))"
using S by (auto intro!: mult_mono power_mono nonneg ennreal_leI
simp: incseq_def le_fun_def less_imp_le)
fix n have "(\<lambda>i. ennreal (f n * S i^n)) \<longlonglongrightarrow> ennreal (f n * 1^n)"
by (intro tendsto_intros tendsto_ennrealI S)
then show "(\<lambda>i. ennreal (f n * S i^n)) \<longlonglongrightarrow> ennreal (f n)"
by simp
qed (auto simp: S_nonneg intro!: mult_nonneg_nonneg nonneg)
also have "(\<lambda>i. (\<integral>\<^sup>+n. f n * S i^n \<partial>count_space UNIV)) = (\<lambda>i. \<Sum>n. f n * S i^n)"
by (subst nn_integral_count_space_nat)
(intro ext suminf_ennreal2 mult_nonneg_nonneg nonneg S_nonneg
zero_le_power summable S)+
also have "(\<integral>\<^sup>+n. ennreal (f n) \<partial>count_space UNIV) = (\<Sum>n. ennreal (f n))"
by (simp add: nn_integral_count_space_nat nonneg)
finally show "(\<lambda>n. ennreal (\<Sum>na. f na * S n ^ na)) \<longlonglongrightarrow> (\<Sum>n. ennreal (f n))" .
qed
subsubsection \<open>Measures with Restricted Space\<close>
lemma simple_function_restrict_space_ennreal:
fixes f :: "'a \<Rightarrow> ennreal"
assumes "\<Omega> \<inter> space M \<in> sets M"
shows "simple_function (restrict_space M \<Omega>) f \<longleftrightarrow> simple_function M (\<lambda>x. f x * indicator \<Omega> x)"
proof -
{ assume "finite (f ` space (restrict_space M \<Omega>))"
then have "finite (f ` space (restrict_space M \<Omega>) \<union> {0})" by simp
then have "finite ((\<lambda>x. f x * indicator \<Omega> x) ` space M)"
by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) }
moreover
{ assume "finite ((\<lambda>x. f x * indicator \<Omega> x) ` space M)"
then have "finite (f ` space (restrict_space M \<Omega>))"
by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) }
ultimately show ?thesis
unfolding
simple_function_iff_borel_measurable borel_measurable_restrict_space_iff_ennreal[OF assms]
by auto
qed
lemma simple_function_restrict_space:
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
assumes "\<Omega> \<inter> space M \<in> sets M"
shows "simple_function (restrict_space M \<Omega>) f \<longleftrightarrow> simple_function M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)"
proof -
{ assume "finite (f ` space (restrict_space M \<Omega>))"
then have "finite (f ` space (restrict_space M \<Omega>) \<union> {0})" by simp
then have "finite ((\<lambda>x. indicator \<Omega> x *\<^sub>R f x) ` space M)"
by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) }
moreover
{ assume "finite ((\<lambda>x. indicator \<Omega> x *\<^sub>R f x) ` space M)"
then have "finite (f ` space (restrict_space M \<Omega>))"
by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) }
ultimately show ?thesis
unfolding simple_function_iff_borel_measurable
borel_measurable_restrict_space_iff[OF assms]
by auto
qed
lemma simple_integral_restrict_space:
assumes \<Omega>: "\<Omega> \<inter> space M \<in> sets M" "simple_function (restrict_space M \<Omega>) f"
shows "simple_integral (restrict_space M \<Omega>) f = simple_integral M (\<lambda>x. f x * indicator \<Omega> x)"
using simple_function_restrict_space_ennreal[THEN iffD1, OF \<Omega>, THEN simple_functionD(1)]
by (auto simp add: space_restrict_space emeasure_restrict_space[OF \<Omega>(1)] le_infI2 simple_integral_def
split: split_indicator split_indicator_asm
intro!: sum.mono_neutral_cong_left ennreal_mult_left_cong arg_cong2[where f=emeasure])
lemma nn_integral_restrict_space:
assumes \<Omega>[simp]: "\<Omega> \<inter> space M \<in> sets M"
shows "nn_integral (restrict_space M \<Omega>) f = nn_integral M (\<lambda>x. f x * indicator \<Omega> x)"
proof -
let ?R = "restrict_space M \<Omega>" and ?X = "\<lambda>M f. {s. simple_function M s \<and> s \<le> f \<and> (\<forall>x. s x < top)}"
have "integral\<^sup>S ?R ` ?X ?R f = integral\<^sup>S M ` ?X M (\<lambda>x. f x * indicator \<Omega> x)"
proof (safe intro!: image_eqI)
fix s assume s: "simple_function ?R s" "s \<le> f" "\<forall>x. s x < top"
from s show "integral\<^sup>S (restrict_space M \<Omega>) s = integral\<^sup>S M (\<lambda>x. s x * indicator \<Omega> x)"
by (intro simple_integral_restrict_space) auto
from s show "simple_function M (\<lambda>x. s x * indicator \<Omega> x)"
by (simp add: simple_function_restrict_space_ennreal)
from s show "(\<lambda>x. s x * indicator \<Omega> x) \<le> (\<lambda>x. f x * indicator \<Omega> x)"
"\<And>x. s x * indicator \<Omega> x < top"
by (auto split: split_indicator simp: le_fun_def image_subset_iff)
next
fix s assume s: "simple_function M s" "s \<le> (\<lambda>x. f x * indicator \<Omega> x)" "\<forall>x. s x < top"
then have "simple_function M (\<lambda>x. s x * indicator (\<Omega> \<inter> space M) x)" (is ?s')
by (intro simple_function_mult simple_function_indicator) auto
also have "?s' \<longleftrightarrow> simple_function M (\<lambda>x. s x * indicator \<Omega> x)"
by (rule simple_function_cong) (auto split: split_indicator)
finally show sf: "simple_function (restrict_space M \<Omega>) s"
by (simp add: simple_function_restrict_space_ennreal)
from s have s_eq: "s = (\<lambda>x. s x * indicator \<Omega> x)"
by (auto simp add: fun_eq_iff le_fun_def image_subset_iff
split: split_indicator split_indicator_asm
intro: antisym)
show "integral\<^sup>S M s = integral\<^sup>S (restrict_space M \<Omega>) s"
by (subst s_eq) (rule simple_integral_restrict_space[symmetric, OF \<Omega> sf])
show "\<And>x. s x < top"
using s by (auto simp: image_subset_iff)
from s show "s \<le> f"
by (subst s_eq) (auto simp: image_subset_iff le_fun_def split: split_indicator split_indicator_asm)
qed
then show ?thesis
unfolding nn_integral_def_finite by (simp cong del: SUP_cong_simp)
qed
lemma nn_integral_count_space_indicator:
assumes "NO_MATCH (UNIV::'a set) (X::'a set)"
shows "(\<integral>\<^sup>+x. f x \<partial>count_space X) = (\<integral>\<^sup>+x. f x * indicator X x \<partial>count_space UNIV)"
by (simp add: nn_integral_restrict_space[symmetric] restrict_count_space)
lemma nn_integral_count_space_eq:
"(\<And>x. x \<in> A - B \<Longrightarrow> f x = 0) \<Longrightarrow> (\<And>x. x \<in> B - A \<Longrightarrow> f x = 0) \<Longrightarrow>
(\<integral>\<^sup>+x. f x \<partial>count_space A) = (\<integral>\<^sup>+x. f x \<partial>count_space B)"
by (auto simp: nn_integral_count_space_indicator intro!: nn_integral_cong split: split_indicator)
lemma nn_integral_ge_point:
assumes "x \<in> A"
shows "p x \<le> \<integral>\<^sup>+ x. p x \<partial>count_space A"
proof -
from assms have "p x \<le> \<integral>\<^sup>+ x. p x \<partial>count_space {x}"
by(auto simp add: nn_integral_count_space_finite max_def)
also have "\<dots> = \<integral>\<^sup>+ x'. p x' * indicator {x} x' \<partial>count_space A"
using assms by(auto simp add: nn_integral_count_space_indicator indicator_def intro!: nn_integral_cong)
also have "\<dots> \<le> \<integral>\<^sup>+ x. p x \<partial>count_space A"
by(rule nn_integral_mono)(simp add: indicator_def)
finally show ?thesis .
qed
subsubsection \<open>Measure spaces with an associated density\<close>
definition\<^marker>\<open>tag important\<close> density :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> 'a measure" where
"density M f = measure_of (space M) (sets M) (\<lambda>A. \<integral>\<^sup>+ x. f x * indicator A x \<partial>M)"
lemma
shows sets_density[simp, measurable_cong]: "sets (density M f) = sets M"
and space_density[simp]: "space (density M f) = space M"
by (auto simp: density_def)
(* FIXME: add conversion to simplify space, sets and measurable *)
lemma space_density_imp[measurable_dest]:
"\<And>x M f. x \<in> space (density M f) \<Longrightarrow> x \<in> space M" by auto
lemma
shows measurable_density_eq1[simp]: "g \<in> measurable (density Mg f) Mg' \<longleftrightarrow> g \<in> measurable Mg Mg'"
and measurable_density_eq2[simp]: "h \<in> measurable Mh (density Mh' f) \<longleftrightarrow> h \<in> measurable Mh Mh'"
and simple_function_density_eq[simp]: "simple_function (density Mu f) u \<longleftrightarrow> simple_function Mu u"
unfolding measurable_def simple_function_def by simp_all
lemma density_cong: "f \<in> borel_measurable M \<Longrightarrow> f' \<in> borel_measurable M \<Longrightarrow>
(AE x in M. f x = f' x) \<Longrightarrow> density M f = density M f'"
unfolding density_def by (auto intro!: measure_of_eq nn_integral_cong_AE sets.space_closed)
lemma emeasure_density:
assumes f[measurable]: "f \<in> borel_measurable M" and A[measurable]: "A \<in> sets M"
shows "emeasure (density M f) A = (\<integral>\<^sup>+ x. f x * indicator A x \<partial>M)"
(is "_ = ?\<mu> A")
unfolding density_def
proof (rule emeasure_measure_of_sigma)
show "sigma_algebra (space M) (sets M)" ..
show "positive (sets M) ?\<mu>"
using f by (auto simp: positive_def)
show "countably_additive (sets M) ?\<mu>"
proof (intro countably_additiveI)
fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> sets M"
then have "\<And>i. A i \<in> sets M" by auto
then have *: "\<And>i. (\<lambda>x. f x * indicator (A i) x) \<in> borel_measurable M"
by auto
assume disj: "disjoint_family A"
then have "(\<Sum>n. ?\<mu> (A n)) = (\<integral>\<^sup>+ x. (\<Sum>n. f x * indicator (A n) x) \<partial>M)"
using f * by (subst nn_integral_suminf) auto
also have "(\<integral>\<^sup>+ x. (\<Sum>n. f x * indicator (A n) x) \<partial>M) = (\<integral>\<^sup>+ x. f x * (\<Sum>n. indicator (A n) x) \<partial>M)"
using f by (auto intro!: ennreal_suminf_cmult nn_integral_cong_AE)
also have "\<dots> = (\<integral>\<^sup>+ x. f x * indicator (\<Union>n. A n) x \<partial>M)"
unfolding suminf_indicator[OF disj] ..
finally show "(\<Sum>i. \<integral>\<^sup>+ x. f x * indicator (A i) x \<partial>M) = \<integral>\<^sup>+ x. f x * indicator (\<Union>i. A i) x \<partial>M" .
qed
qed fact
lemma null_sets_density_iff:
assumes f: "f \<in> borel_measurable M"
shows "A \<in> null_sets (density M f) \<longleftrightarrow> A \<in> sets M \<and> (AE x in M. x \<in> A \<longrightarrow> f x = 0)"
proof -
{ assume "A \<in> sets M"
have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = 0 \<longleftrightarrow> emeasure M {x \<in> space M. f x * indicator A x \<noteq> 0} = 0"
using f \<open>A \<in> sets M\<close> by (intro nn_integral_0_iff) auto
also have "\<dots> \<longleftrightarrow> (AE x in M. f x * indicator A x = 0)"
using f \<open>A \<in> sets M\<close> by (intro AE_iff_measurable[OF _ refl, symmetric]) auto
also have "(AE x in M. f x * indicator A x = 0) \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> f x \<le> 0)"
by (auto simp add: indicator_def max_def split: if_split_asm)
finally have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = 0 \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> f x \<le> 0)" . }
with f show ?thesis
by (simp add: null_sets_def emeasure_density cong: conj_cong)
qed
lemma AE_density:
assumes f: "f \<in> borel_measurable M"
shows "(AE x in density M f. P x) \<longleftrightarrow> (AE x in M. 0 < f x \<longrightarrow> P x)"
proof
assume "AE x in density M f. P x"
with f obtain N where "{x \<in> space M. \<not> P x} \<subseteq> N" "N \<in> sets M" and ae: "AE x in M. x \<in> N \<longrightarrow> f x = 0"
by (auto simp: eventually_ae_filter null_sets_density_iff)
then have "AE x in M. x \<notin> N \<longrightarrow> P x" by auto
with ae show "AE x in M. 0 < f x \<longrightarrow> P x"
by (rule eventually_elim2) auto
next
fix N assume ae: "AE x in M. 0 < f x \<longrightarrow> P x"
then obtain N where "{x \<in> space M. \<not> (0 < f x \<longrightarrow> P x)} \<subseteq> N" "N \<in> null_sets M"
by (auto simp: eventually_ae_filter)
then have *: "{x \<in> space (density M f). \<not> P x} \<subseteq> N \<union> {x\<in>space M. f x = 0}"
"N \<union> {x\<in>space M. f x = 0} \<in> sets M" and ae2: "AE x in M. x \<notin> N"
using f by (auto simp: subset_eq zero_less_iff_neq_zero intro!: AE_not_in)
show "AE x in density M f. P x"
using ae2
unfolding eventually_ae_filter[of _ "density M f"] Bex_def null_sets_density_iff[OF f]
by (intro exI[of _ "N \<union> {x\<in>space M. f x = 0}"] conjI *) (auto elim: eventually_elim2)
qed
lemma\<^marker>\<open>tag important\<close> nn_integral_density:
assumes f: "f \<in> borel_measurable M"
assumes g: "g \<in> borel_measurable M"
shows "integral\<^sup>N (density M f) g = (\<integral>\<^sup>+ x. f x * g x \<partial>M)"
using g proof induct
case (cong u v)
then show ?case
apply (subst nn_integral_cong[OF cong(3)])
apply (simp_all cong: nn_integral_cong)
done
next
case (set A) then show ?case
by (simp add: emeasure_density f)
next
case (mult u c)
moreover have "\<And>x. f x * (c * u x) = c * (f x * u x)" by (simp add: field_simps)
ultimately show ?case
using f by (simp add: nn_integral_cmult)
next
case (add u v)
then have "\<And>x. f x * (v x + u x) = f x * v x + f x * u x"
by (simp add: distrib_left)
with add f show ?case
by (auto simp add: nn_integral_add intro!: nn_integral_add[symmetric])
next
case (seq U)
have eq: "AE x in M. f x * (SUP i. U i x) = (SUP i. f x * U i x)"
by eventually_elim (simp add: SUP_mult_left_ennreal seq)
from seq f show ?case
apply (simp add: nn_integral_monotone_convergence_SUP image_comp)
apply (subst nn_integral_cong_AE[OF eq])
apply (subst nn_integral_monotone_convergence_SUP_AE)
apply (auto simp: incseq_def le_fun_def intro!: mult_left_mono)
done
qed
lemma density_distr:
assumes [measurable]: "f \<in> borel_measurable N" "X \<in> measurable M N"
shows "density (distr M N X) f = distr (density M (\<lambda>x. f (X x))) N X"
by (intro measure_eqI)
(auto simp add: emeasure_density nn_integral_distr emeasure_distr
split: split_indicator intro!: nn_integral_cong)
lemma emeasure_restricted:
assumes S: "S \<in> sets M" and X: "X \<in> sets M"
shows "emeasure (density M (indicator S)) X = emeasure M (S \<inter> X)"
proof -
have "emeasure (density M (indicator S)) X = (\<integral>\<^sup>+x. indicator S x * indicator X x \<partial>M)"
using S X by (simp add: emeasure_density)
also have "\<dots> = (\<integral>\<^sup>+x. indicator (S \<inter> X) x \<partial>M)"
by (auto intro!: nn_integral_cong simp: indicator_def)
also have "\<dots> = emeasure M (S \<inter> X)"
using S X by (simp add: sets.Int)
finally show ?thesis .
qed
lemma measure_restricted:
"S \<in> sets M \<Longrightarrow> X \<in> sets M \<Longrightarrow> measure (density M (indicator S)) X = measure M (S \<inter> X)"
by (simp add: emeasure_restricted measure_def)
lemma (in finite_measure) finite_measure_restricted:
"S \<in> sets M \<Longrightarrow> finite_measure (density M (indicator S))"
by standard (simp add: emeasure_restricted)
lemma emeasure_density_const:
"A \<in> sets M \<Longrightarrow> emeasure (density M (\<lambda>_. c)) A = c * emeasure M A"
by (auto simp: nn_integral_cmult_indicator emeasure_density)
lemma measure_density_const:
"A \<in> sets M \<Longrightarrow> c \<noteq> \<infinity> \<Longrightarrow> measure (density M (\<lambda>_. c)) A = enn2real c * measure M A"
by (auto simp: emeasure_density_const measure_def enn2real_mult)
lemma density_density_eq:
"f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow>
density (density M f) g = density M (\<lambda>x. f x * g x)"
by (auto intro!: measure_eqI simp: emeasure_density nn_integral_density ac_simps)
lemma distr_density_distr:
assumes T: "T \<in> measurable M M'" and T': "T' \<in> measurable M' M"
and inv: "\<forall>x\<in>space M. T' (T x) = x"
assumes f: "f \<in> borel_measurable M'"
shows "distr (density (distr M M' T) f) M T' = density M (f \<circ> T)" (is "?R = ?L")
proof (rule measure_eqI)
fix A assume A: "A \<in> sets ?R"
{ fix x assume "x \<in> space M"
with sets.sets_into_space[OF A]
have "indicator (T' -` A \<inter> space M') (T x) = (indicator A x :: ennreal)"
using T inv by (auto simp: indicator_def measurable_space) }
with A T T' f show "emeasure ?R A = emeasure ?L A"
by (simp add: measurable_comp emeasure_density emeasure_distr
nn_integral_distr measurable_sets cong: nn_integral_cong)
qed simp
lemma density_density_divide:
fixes f g :: "'a \<Rightarrow> real"
assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
assumes g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x"
assumes ac: "AE x in M. f x = 0 \<longrightarrow> g x = 0"
shows "density (density M f) (\<lambda>x. g x / f x) = density M g"
proof -
have "density M g = density M (\<lambda>x. ennreal (f x) * ennreal (g x / f x))"
using f g ac by (auto intro!: density_cong measurable_If simp: ennreal_mult[symmetric])
then show ?thesis
using f g by (subst density_density_eq) auto
qed
lemma density_1: "density M (\<lambda>_. 1) = M"
by (intro measure_eqI) (auto simp: emeasure_density)
lemma emeasure_density_add:
assumes X: "X \<in> sets M"
assumes Mf[measurable]: "f \<in> borel_measurable M"
assumes Mg[measurable]: "g \<in> borel_measurable M"
shows "emeasure (density M f) X + emeasure (density M g) X =
emeasure (density M (\<lambda>x. f x + g x)) X"
using assms
apply (subst (1 2 3) emeasure_density, simp_all) []
apply (subst nn_integral_add[symmetric], simp_all) []
apply (intro nn_integral_cong, simp split: split_indicator)
done
subsubsection \<open>Point measure\<close>
definition\<^marker>\<open>tag important\<close> point_measure :: "'a set \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> 'a measure" where
"point_measure A f = density (count_space A) f"
lemma
shows space_point_measure: "space (point_measure A f) = A"
and sets_point_measure: "sets (point_measure A f) = Pow A"
by (auto simp: point_measure_def)
lemma sets_point_measure_count_space[measurable_cong]: "sets (point_measure A f) = sets (count_space A)"
by (simp add: sets_point_measure)
lemma measurable_point_measure_eq1[simp]:
"g \<in> measurable (point_measure A f) M \<longleftrightarrow> g \<in> A \<rightarrow> space M"
unfolding point_measure_def by simp
lemma measurable_point_measure_eq2_finite[simp]:
"finite A \<Longrightarrow>
g \<in> measurable M (point_measure A f) \<longleftrightarrow>
(g \<in> space M \<rightarrow> A \<and> (\<forall>a\<in>A. g -` {a} \<inter> space M \<in> sets M))"
unfolding point_measure_def by (simp add: measurable_count_space_eq2)
lemma simple_function_point_measure[simp]:
"simple_function (point_measure A f) g \<longleftrightarrow> finite (g ` A)"
by (simp add: point_measure_def)
lemma emeasure_point_measure:
assumes A: "finite {a\<in>X. 0 < f a}" "X \<subseteq> A"
shows "emeasure (point_measure A f) X = (\<Sum>a|a\<in>X \<and> 0 < f a. f a)"
proof -
have "{a. (a \<in> X \<longrightarrow> a \<in> A \<and> 0 < f a) \<and> a \<in> X} = {a\<in>X. 0 < f a}"
using \<open>X \<subseteq> A\<close> by auto
with A show ?thesis
by (simp add: emeasure_density nn_integral_count_space point_measure_def indicator_def)
qed
lemma emeasure_point_measure_finite:
"finite A \<Longrightarrow> X \<subseteq> A \<Longrightarrow> emeasure (point_measure A f) X = (\<Sum>a\<in>X. f a)"
by (subst emeasure_point_measure) (auto dest: finite_subset intro!: sum.mono_neutral_left simp: less_le)
lemma emeasure_point_measure_finite2:
"X \<subseteq> A \<Longrightarrow> finite X \<Longrightarrow> emeasure (point_measure A f) X = (\<Sum>a\<in>X. f a)"
by (subst emeasure_point_measure)
(auto dest: finite_subset intro!: sum.mono_neutral_left simp: less_le)
lemma null_sets_point_measure_iff:
"X \<in> null_sets (point_measure A f) \<longleftrightarrow> X \<subseteq> A \<and> (\<forall>x\<in>X. f x = 0)"
by (auto simp: AE_count_space null_sets_density_iff point_measure_def)
lemma AE_point_measure:
"(AE x in point_measure A f. P x) \<longleftrightarrow> (\<forall>x\<in>A. 0 < f x \<longrightarrow> P x)"
unfolding point_measure_def
by (subst AE_density) (auto simp: AE_density AE_count_space point_measure_def)
lemma nn_integral_point_measure:
"finite {a\<in>A. 0 < f a \<and> 0 < g a} \<Longrightarrow>
integral\<^sup>N (point_measure A f) g = (\<Sum>a|a\<in>A \<and> 0 < f a \<and> 0 < g a. f a * g a)"
unfolding point_measure_def
by (subst nn_integral_density)
(simp_all add: nn_integral_density nn_integral_count_space ennreal_zero_less_mult_iff)
lemma nn_integral_point_measure_finite:
"finite A \<Longrightarrow> integral\<^sup>N (point_measure A f) g = (\<Sum>a\<in>A. f a * g a)"
by (subst nn_integral_point_measure) (auto intro!: sum.mono_neutral_left simp: less_le)
subsubsection \<open>Uniform measure\<close>
definition\<^marker>\<open>tag important\<close> "uniform_measure M A = density M (\<lambda>x. indicator A x / emeasure M A)"
lemma
shows sets_uniform_measure[simp, measurable_cong]: "sets (uniform_measure M A) = sets M"
and space_uniform_measure[simp]: "space (uniform_measure M A) = space M"
by (auto simp: uniform_measure_def)
lemma emeasure_uniform_measure[simp]:
assumes A: "A \<in> sets M" and B: "B \<in> sets M"
shows "emeasure (uniform_measure M A) B = emeasure M (A \<inter> B) / emeasure M A"
proof -
from A B have "emeasure (uniform_measure M A) B = (\<integral>\<^sup>+x. (1 / emeasure M A) * indicator (A \<inter> B) x \<partial>M)"
by (auto simp add: uniform_measure_def emeasure_density divide_ennreal_def split: split_indicator
intro!: nn_integral_cong)
also have "\<dots> = emeasure M (A \<inter> B) / emeasure M A"
using A B
by (subst nn_integral_cmult_indicator) (simp_all add: sets.Int divide_ennreal_def mult.commute)
finally show ?thesis .
qed
lemma measure_uniform_measure[simp]:
assumes A: "emeasure M A \<noteq> 0" "emeasure M A \<noteq> \<infinity>" and B: "B \<in> sets M"
shows "measure (uniform_measure M A) B = measure M (A \<inter> B) / measure M A"
using emeasure_uniform_measure[OF emeasure_neq_0_sets[OF A(1)] B] A
by (cases "emeasure M A" "emeasure M (A \<inter> B)" rule: ennreal2_cases)
(simp_all add: measure_def divide_ennreal top_ennreal.rep_eq top_ereal_def ennreal_top_divide)
lemma AE_uniform_measureI:
"A \<in> sets M \<Longrightarrow> (AE x in M. x \<in> A \<longrightarrow> P x) \<Longrightarrow> (AE x in uniform_measure M A. P x)"
unfolding uniform_measure_def by (auto simp: AE_density divide_ennreal_def)
lemma emeasure_uniform_measure_1:
"emeasure M S \<noteq> 0 \<Longrightarrow> emeasure M S \<noteq> \<infinity> \<Longrightarrow> emeasure (uniform_measure M S) S = 1"
by (subst emeasure_uniform_measure)
(simp_all add: emeasure_neq_0_sets emeasure_eq_ennreal_measure divide_ennreal
zero_less_iff_neq_zero[symmetric])
lemma nn_integral_uniform_measure:
assumes f[measurable]: "f \<in> borel_measurable M" and S[measurable]: "S \<in> sets M"
shows "(\<integral>\<^sup>+x. f x \<partial>uniform_measure M S) = (\<integral>\<^sup>+x. f x * indicator S x \<partial>M) / emeasure M S"
proof -
{ assume "emeasure M S = \<infinity>"
then have ?thesis
by (simp add: uniform_measure_def nn_integral_density f) }
moreover
{ assume [simp]: "emeasure M S = 0"
then have ae: "AE x in M. x \<notin> S"
using sets.sets_into_space[OF S]
by (subst AE_iff_measurable[OF _ refl]) (simp_all add: subset_eq cong: rev_conj_cong)
from ae have "(\<integral>\<^sup>+ x. indicator S x / 0 * f x \<partial>M) = 0"
by (subst nn_integral_0_iff_AE) auto
moreover from ae have "(\<integral>\<^sup>+ x. f x * indicator S x \<partial>M) = 0"
by (subst nn_integral_0_iff_AE) auto
ultimately have ?thesis
by (simp add: uniform_measure_def nn_integral_density f) }
moreover have "emeasure M S \<noteq> 0 \<Longrightarrow> emeasure M S \<noteq> \<infinity> \<Longrightarrow> ?thesis"
unfolding uniform_measure_def
by (subst nn_integral_density)
(auto simp: ennreal_times_divide f nn_integral_divide[symmetric] mult.commute)
ultimately show ?thesis by blast
qed
lemma AE_uniform_measure:
assumes "emeasure M A \<noteq> 0" "emeasure M A < \<infinity>"
shows "(AE x in uniform_measure M A. P x) \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> P x)"
proof -
have "A \<in> sets M"
using \<open>emeasure M A \<noteq> 0\<close> by (metis emeasure_notin_sets)
moreover have "\<And>x. 0 < indicator A x / emeasure M A \<longleftrightarrow> x \<in> A"
using assms
by (cases "emeasure M A") (auto split: split_indicator simp: ennreal_zero_less_divide)
ultimately show ?thesis
unfolding uniform_measure_def by (simp add: AE_density)
qed
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Null measure\<close>
lemma null_measure_eq_density: "null_measure M = density M (\<lambda>_. 0)"
by (intro measure_eqI) (simp_all add: emeasure_density)
lemma nn_integral_null_measure[simp]: "(\<integral>\<^sup>+x. f x \<partial>null_measure M) = 0"
by (auto simp add: nn_integral_def simple_integral_def SUP_constant bot_ennreal_def le_fun_def
intro!: exI[of _ "\<lambda>x. 0"])
lemma density_null_measure[simp]: "density (null_measure M) f = null_measure M"
proof (intro measure_eqI)
fix A show "emeasure (density (null_measure M) f) A = emeasure (null_measure M) A"
by (simp add: density_def) (simp only: null_measure_def[symmetric] emeasure_null_measure)
qed simp
subsubsection \<open>Uniform count measure\<close>
definition\<^marker>\<open>tag important\<close> "uniform_count_measure A = point_measure A (\<lambda>x. 1 / card A)"
lemma
shows space_uniform_count_measure: "space (uniform_count_measure A) = A"
and sets_uniform_count_measure: "sets (uniform_count_measure A) = Pow A"
unfolding uniform_count_measure_def by (auto simp: space_point_measure sets_point_measure)
lemma sets_uniform_count_measure_count_space[measurable_cong]:
"sets (uniform_count_measure A) = sets (count_space A)"
by (simp add: sets_uniform_count_measure)
lemma emeasure_uniform_count_measure:
"finite A \<Longrightarrow> X \<subseteq> A \<Longrightarrow> emeasure (uniform_count_measure A) X = card X / card A"
by (simp add: emeasure_point_measure_finite uniform_count_measure_def divide_inverse ennreal_mult
ennreal_of_nat_eq_real_of_nat)
lemma measure_uniform_count_measure:
"finite A \<Longrightarrow> X \<subseteq> A \<Longrightarrow> measure (uniform_count_measure A) X = card X / card A"
by (simp add: emeasure_point_measure_finite uniform_count_measure_def measure_def enn2real_mult)
lemma space_uniform_count_measure_empty_iff [simp]:
"space (uniform_count_measure X) = {} \<longleftrightarrow> X = {}"
by(simp add: space_uniform_count_measure)
lemma sets_uniform_count_measure_eq_UNIV [simp]:
"sets (uniform_count_measure UNIV) = UNIV \<longleftrightarrow> True"
"UNIV = sets (uniform_count_measure UNIV) \<longleftrightarrow> True"
by(simp_all add: sets_uniform_count_measure)
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Scaled measure\<close>
lemma nn_integral_scale_measure:
assumes f: "f \<in> borel_measurable M"
shows "nn_integral (scale_measure r M) f = r * nn_integral M f"
using f
proof induction
case (cong f g)
thus ?case
by(simp add: cong.hyps space_scale_measure cong: nn_integral_cong_simp)
next
case (mult f c)
thus ?case
by(simp add: nn_integral_cmult max_def mult.assoc mult.left_commute)
next
case (add f g)
thus ?case
by(simp add: nn_integral_add distrib_left)
next
case (seq U)
thus ?case
by(simp add: nn_integral_monotone_convergence_SUP SUP_mult_left_ennreal image_comp)
qed simp
end
diff --git a/src/HOL/Analysis/Ordered_Euclidean_Space.thy b/src/HOL/Analysis/Ordered_Euclidean_Space.thy
--- a/src/HOL/Analysis/Ordered_Euclidean_Space.thy
+++ b/src/HOL/Analysis/Ordered_Euclidean_Space.thy
@@ -1,305 +1,305 @@
section \<open>Ordered Euclidean Space\<close>
theory Ordered_Euclidean_Space
imports
Convex_Euclidean_Space
"HOL-Library.Product_Order"
begin
text \<open>An ordering on euclidean spaces that will allow us to talk about intervals\<close>
class ordered_euclidean_space = ord + inf + sup + abs + Inf + Sup + euclidean_space +
assumes eucl_le: "x \<le> y \<longleftrightarrow> (\<forall>i\<in>Basis. x \<bullet> i \<le> y \<bullet> i)"
assumes eucl_less_le_not_le: "x < y \<longleftrightarrow> x \<le> y \<and> \<not> y \<le> x"
assumes eucl_inf: "inf x y = (\<Sum>i\<in>Basis. inf (x \<bullet> i) (y \<bullet> i) *\<^sub>R i)"
assumes eucl_sup: "sup x y = (\<Sum>i\<in>Basis. sup (x \<bullet> i) (y \<bullet> i) *\<^sub>R i)"
assumes eucl_Inf: "Inf X = (\<Sum>i\<in>Basis. (INF x\<in>X. x \<bullet> i) *\<^sub>R i)"
assumes eucl_Sup: "Sup X = (\<Sum>i\<in>Basis. (SUP x\<in>X. x \<bullet> i) *\<^sub>R i)"
assumes eucl_abs: "\<bar>x\<bar> = (\<Sum>i\<in>Basis. \<bar>x \<bullet> i\<bar> *\<^sub>R i)"
begin
subclass order
by standard
(auto simp: eucl_le eucl_less_le_not_le intro!: euclidean_eqI antisym intro: order.trans)
subclass ordered_ab_group_add_abs
by standard (auto simp: eucl_le inner_add_left eucl_abs abs_leI)
subclass ordered_real_vector
by standard (auto simp: eucl_le intro!: mult_left_mono mult_right_mono)
subclass lattice
by standard (auto simp: eucl_inf eucl_sup eucl_le)
subclass distrib_lattice
by standard (auto simp: eucl_inf eucl_sup sup_inf_distrib1 intro!: euclidean_eqI)
subclass conditionally_complete_lattice
proof
fix z::'a and X::"'a set"
assume "X \<noteq> {}"
hence "\<And>i. (\<lambda>x. x \<bullet> i) ` X \<noteq> {}" by simp
thus "(\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> z \<le> Inf X" "(\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X \<le> z"
by (auto simp: eucl_Inf eucl_Sup eucl_le
intro!: cInf_greatest cSup_least)
qed (force intro!: cInf_lower cSup_upper
simp: bdd_below_def bdd_above_def preorder_class.bdd_below_def preorder_class.bdd_above_def
eucl_Inf eucl_Sup eucl_le)+
lemma inner_Basis_inf_left: "i \<in> Basis \<Longrightarrow> inf x y \<bullet> i = inf (x \<bullet> i) (y \<bullet> i)"
and inner_Basis_sup_left: "i \<in> Basis \<Longrightarrow> sup x y \<bullet> i = sup (x \<bullet> i) (y \<bullet> i)"
- by (simp_all add: eucl_inf eucl_sup inner_sum_left inner_Basis if_distrib comm_monoid_add_class.sum.delta
+ by (simp_all add: eucl_inf eucl_sup inner_sum_left inner_Basis if_distrib
cong: if_cong)
lemma inner_Basis_INF_left: "i \<in> Basis \<Longrightarrow> (INF x\<in>X. f x) \<bullet> i = (INF x\<in>X. f x \<bullet> i)"
and inner_Basis_SUP_left: "i \<in> Basis \<Longrightarrow> (SUP x\<in>X. f x) \<bullet> i = (SUP x\<in>X. f x \<bullet> i)"
using eucl_Sup [of "f ` X"] eucl_Inf [of "f ` X"] by (simp_all add: image_comp)
lemma abs_inner: "i \<in> Basis \<Longrightarrow> \<bar>x\<bar> \<bullet> i = \<bar>x \<bullet> i\<bar>"
by (auto simp: eucl_abs)
lemma
abs_scaleR: "\<bar>a *\<^sub>R b\<bar> = \<bar>a\<bar> *\<^sub>R \<bar>b\<bar>"
by (auto simp: eucl_abs abs_mult intro!: euclidean_eqI)
lemma interval_inner_leI:
assumes "x \<in> {a .. b}" "0 \<le> i"
shows "a\<bullet>i \<le> x\<bullet>i" "x\<bullet>i \<le> b\<bullet>i"
using assms
unfolding euclidean_inner[of a i] euclidean_inner[of x i] euclidean_inner[of b i]
by (auto intro!: ordered_comm_monoid_add_class.sum_mono mult_right_mono simp: eucl_le)
lemma inner_nonneg_nonneg:
shows "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> a \<bullet> b"
using interval_inner_leI[of a 0 a b]
by auto
lemma inner_Basis_mono:
shows "a \<le> b \<Longrightarrow> c \<in> Basis \<Longrightarrow> a \<bullet> c \<le> b \<bullet> c"
by (simp add: eucl_le)
lemma Basis_nonneg[intro, simp]: "i \<in> Basis \<Longrightarrow> 0 \<le> i"
by (auto simp: eucl_le inner_Basis)
lemma Sup_eq_maximum_componentwise:
fixes s::"'a set"
assumes i: "\<And>b. b \<in> Basis \<Longrightarrow> X \<bullet> b = i b \<bullet> b"
assumes sup: "\<And>b x. b \<in> Basis \<Longrightarrow> x \<in> s \<Longrightarrow> x \<bullet> b \<le> X \<bullet> b"
assumes i_s: "\<And>b. b \<in> Basis \<Longrightarrow> (i b \<bullet> b) \<in> (\<lambda>x. x \<bullet> b) ` s"
shows "Sup s = X"
using assms
unfolding eucl_Sup euclidean_representation_sum
by (auto intro!: conditionally_complete_lattice_class.cSup_eq_maximum)
lemma Inf_eq_minimum_componentwise:
assumes i: "\<And>b. b \<in> Basis \<Longrightarrow> X \<bullet> b = i b \<bullet> b"
assumes sup: "\<And>b x. b \<in> Basis \<Longrightarrow> x \<in> s \<Longrightarrow> X \<bullet> b \<le> x \<bullet> b"
assumes i_s: "\<And>b. b \<in> Basis \<Longrightarrow> (i b \<bullet> b) \<in> (\<lambda>x. x \<bullet> b) ` s"
shows "Inf s = X"
using assms
unfolding eucl_Inf euclidean_representation_sum
by (auto intro!: conditionally_complete_lattice_class.cInf_eq_minimum)
end
proposition compact_attains_Inf_componentwise:
fixes b::"'a::ordered_euclidean_space"
assumes "b \<in> Basis" assumes "X \<noteq> {}" "compact X"
obtains x where "x \<in> X" "x \<bullet> b = Inf X \<bullet> b" "\<And>y. y \<in> X \<Longrightarrow> x \<bullet> b \<le> y \<bullet> b"
proof atomize_elim
let ?proj = "(\<lambda>x. x \<bullet> b) ` X"
from assms have "compact ?proj" "?proj \<noteq> {}"
by (auto intro!: compact_continuous_image continuous_intros)
from compact_attains_inf[OF this]
obtain s x
where s: "s\<in>(\<lambda>x. x \<bullet> b) ` X" "\<And>t. t\<in>(\<lambda>x. x \<bullet> b) ` X \<Longrightarrow> s \<le> t"
and x: "x \<in> X" "s = x \<bullet> b" "\<And>y. y \<in> X \<Longrightarrow> x \<bullet> b \<le> y \<bullet> b"
by auto
hence "Inf ?proj = x \<bullet> b"
by (auto intro!: conditionally_complete_lattice_class.cInf_eq_minimum)
hence "x \<bullet> b = Inf X \<bullet> b"
- by (auto simp: eucl_Inf inner_sum_left inner_Basis if_distrib \<open>b \<in> Basis\<close> sum.delta
+ by (auto simp: eucl_Inf inner_sum_left inner_Basis if_distrib \<open>b \<in> Basis\<close>
cong: if_cong)
with x show "\<exists>x. x \<in> X \<and> x \<bullet> b = Inf X \<bullet> b \<and> (\<forall>y. y \<in> X \<longrightarrow> x \<bullet> b \<le> y \<bullet> b)" by blast
qed
proposition
compact_attains_Sup_componentwise:
fixes b::"'a::ordered_euclidean_space"
assumes "b \<in> Basis" assumes "X \<noteq> {}" "compact X"
obtains x where "x \<in> X" "x \<bullet> b = Sup X \<bullet> b" "\<And>y. y \<in> X \<Longrightarrow> y \<bullet> b \<le> x \<bullet> b"
proof atomize_elim
let ?proj = "(\<lambda>x. x \<bullet> b) ` X"
from assms have "compact ?proj" "?proj \<noteq> {}"
by (auto intro!: compact_continuous_image continuous_intros)
from compact_attains_sup[OF this]
obtain s x
where s: "s\<in>(\<lambda>x. x \<bullet> b) ` X" "\<And>t. t\<in>(\<lambda>x. x \<bullet> b) ` X \<Longrightarrow> t \<le> s"
and x: "x \<in> X" "s = x \<bullet> b" "\<And>y. y \<in> X \<Longrightarrow> y \<bullet> b \<le> x \<bullet> b"
by auto
hence "Sup ?proj = x \<bullet> b"
by (auto intro!: cSup_eq_maximum)
hence "x \<bullet> b = Sup X \<bullet> b"
- by (auto simp: eucl_Sup[where 'a='a] inner_sum_left inner_Basis if_distrib \<open>b \<in> Basis\<close> sum.delta
+ by (auto simp: eucl_Sup[where 'a='a] inner_sum_left inner_Basis if_distrib \<open>b \<in> Basis\<close>
cong: if_cong)
with x show "\<exists>x. x \<in> X \<and> x \<bullet> b = Sup X \<bullet> b \<and> (\<forall>y. y \<in> X \<longrightarrow> y \<bullet> b \<le> x \<bullet> b)" by blast
qed
lemma tendsto_sup[tendsto_intros]:
fixes X :: "'a \<Rightarrow> 'b::ordered_euclidean_space"
assumes "(X \<longlongrightarrow> x) net" "(Y \<longlongrightarrow> y) net"
shows "((\<lambda>i. sup (X i) (Y i)) \<longlongrightarrow> sup x y) net"
unfolding sup_max eucl_sup by (intro assms tendsto_intros)
lemma tendsto_inf[tendsto_intros]:
fixes X :: "'a \<Rightarrow> 'b::ordered_euclidean_space"
assumes "(X \<longlongrightarrow> x) net" "(Y \<longlongrightarrow> y) net"
shows "((\<lambda>i. inf (X i) (Y i)) \<longlongrightarrow> inf x y) net"
unfolding inf_min eucl_inf by (intro assms tendsto_intros)
lemma tendsto_componentwise_max:
assumes f: "(f \<longlongrightarrow> l) F" and g: "(g \<longlongrightarrow> m) F"
shows "((\<lambda>x. (\<Sum>i\<in>Basis. max (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i)) \<longlongrightarrow> (\<Sum>i\<in>Basis. max (l \<bullet> i) (m \<bullet> i) *\<^sub>R i)) F"
by (intro tendsto_intros assms)
lemma tendsto_componentwise_min:
assumes f: "(f \<longlongrightarrow> l) F" and g: "(g \<longlongrightarrow> m) F"
shows "((\<lambda>x. (\<Sum>i\<in>Basis. min (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i)) \<longlongrightarrow> (\<Sum>i\<in>Basis. min (l \<bullet> i) (m \<bullet> i) *\<^sub>R i)) F"
by (intro tendsto_intros assms)
lemma (in order) atLeastatMost_empty'[simp]:
"(\<not> a \<le> b) \<Longrightarrow> {a..b} = {}"
by (auto)
instance real :: ordered_euclidean_space
by standard auto
lemma in_Basis_prod_iff:
fixes i::"'a::euclidean_space*'b::euclidean_space"
shows "i \<in> Basis \<longleftrightarrow> fst i = 0 \<and> snd i \<in> Basis \<or> snd i = 0 \<and> fst i \<in> Basis"
by (cases i) (auto simp: Basis_prod_def)
instantiation\<^marker>\<open>tag unimportant\<close> prod :: (abs, abs) abs
begin
definition "\<bar>x\<bar> = (\<bar>fst x\<bar>, \<bar>snd x\<bar>)"
instance ..
end
instance prod :: (ordered_euclidean_space, ordered_euclidean_space) ordered_euclidean_space
by standard
(auto intro!: add_mono simp add: euclidean_representation_sum' Ball_def inner_prod_def
in_Basis_prod_iff inner_Basis_inf_left inner_Basis_sup_left inner_Basis_INF_left Inf_prod_def
inner_Basis_SUP_left Sup_prod_def less_prod_def less_eq_prod_def eucl_le[where 'a='a]
eucl_le[where 'a='b] abs_prod_def abs_inner)
text\<open>Instantiation for intervals on \<open>ordered_euclidean_space\<close>\<close>
proposition
fixes a :: "'a::ordered_euclidean_space"
shows cbox_interval: "cbox a b = {a..b}"
and interval_cbox: "{a..b} = cbox a b"
and eucl_le_atMost: "{x. \<forall>i\<in>Basis. x \<bullet> i <= a \<bullet> i} = {..a}"
and eucl_le_atLeast: "{x. \<forall>i\<in>Basis. a \<bullet> i <= x \<bullet> i} = {a..}"
by (auto simp: eucl_le[where 'a='a] eucl_less_def box_def cbox_def)
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)
lemma closed_eucl_atLeastAtMost[simp, intro]:
fixes a :: "'a::ordered_euclidean_space"
shows "closed {a..b}"
by (simp add: cbox_interval[symmetric] closed_cbox)
lemma closed_eucl_atMost[simp, intro]:
fixes a :: "'a::ordered_euclidean_space"
shows "closed {..a}"
by (simp add: closed_interval_left eucl_le_atMost[symmetric])
lemma closed_eucl_atLeast[simp, intro]:
fixes a :: "'a::ordered_euclidean_space"
shows "closed {a..}"
by (simp add: closed_interval_right eucl_le_atLeast[symmetric])
lemma bounded_closed_interval [simp]:
fixes a :: "'a::ordered_euclidean_space"
shows "bounded {a .. b}"
using bounded_cbox[of a b]
by (metis interval_cbox)
lemma convex_closed_interval [simp]:
fixes a :: "'a::ordered_euclidean_space"
shows "convex {a .. b}"
using convex_box[of a b]
by (metis interval_cbox)
lemma image_smult_interval:"(\<lambda>x. m *\<^sub>R (x::_::ordered_euclidean_space)) ` {a .. b} =
(if {a .. b} = {} then {} else if 0 \<le> m then {m *\<^sub>R a .. m *\<^sub>R b} else {m *\<^sub>R b .. m *\<^sub>R a})"
using image_smult_cbox[of m a b]
by (simp add: cbox_interval)
lemma [simp]:
fixes a b::"'a::ordered_euclidean_space"
shows is_interval_ic: "is_interval {..a}"
and is_interval_ci: "is_interval {a..}"
and is_interval_cc: "is_interval {b..a}"
by (force simp: is_interval_def eucl_le[where 'a='a])+
lemma connected_interval [simp]:
fixes a b::"'a::ordered_euclidean_space"
shows "connected {a..b}"
using is_interval_cc is_interval_connected by blast
lemma compact_interval [simp]:
fixes a b::"'a::ordered_euclidean_space"
shows "compact {a .. b}"
by (metis compact_cbox interval_cbox)
no_notation
eucl_less (infix "<e" 50)
lemma One_nonneg: "0 \<le> (\<Sum>Basis::'a::ordered_euclidean_space)"
by (auto intro: sum_nonneg)
lemma
fixes a b::"'a::ordered_euclidean_space"
shows bdd_above_cbox[intro, simp]: "bdd_above (cbox a b)"
and bdd_below_cbox[intro, simp]: "bdd_below (cbox a b)"
and bdd_above_box[intro, simp]: "bdd_above (box a b)"
and bdd_below_box[intro, simp]: "bdd_below (box a b)"
unfolding atomize_conj
by (metis bdd_above_Icc bdd_above_mono bdd_below_Icc bdd_below_mono bounded_box
bounded_subset_cbox_symmetric interval_cbox)
instantiation vec :: (ordered_euclidean_space, finite) ordered_euclidean_space
begin
definition\<^marker>\<open>tag important\<close> "inf x y = (\<chi> i. inf (x $ i) (y $ i))"
definition\<^marker>\<open>tag important\<close> "sup x y = (\<chi> i. sup (x $ i) (y $ i))"
definition\<^marker>\<open>tag important\<close> "Inf X = (\<chi> i. (INF x\<in>X. x $ i))"
definition\<^marker>\<open>tag important\<close> "Sup X = (\<chi> i. (SUP x\<in>X. x $ i))"
definition\<^marker>\<open>tag important\<close> "\<bar>x\<bar> = (\<chi> i. \<bar>x $ i\<bar>)"
instance
apply standard
unfolding euclidean_representation_sum'
apply (auto simp: less_eq_vec_def inf_vec_def sup_vec_def Inf_vec_def Sup_vec_def inner_axis
Basis_vec_def inner_Basis_inf_left inner_Basis_sup_left inner_Basis_INF_left
inner_Basis_SUP_left eucl_le[where 'a='a] less_le_not_le abs_vec_def abs_inner)
done
end
end
diff --git a/src/HOL/Analysis/Path_Connected.thy b/src/HOL/Analysis/Path_Connected.thy
--- a/src/HOL/Analysis/Path_Connected.thy
+++ b/src/HOL/Analysis/Path_Connected.thy
@@ -1,4175 +1,4175 @@
(* Title: HOL/Analysis/Path_Connected.thy
Authors: LC Paulson and Robert Himmelmann (TU Muenchen), based on material from HOL Light
*)
section \<open>Path-Connectedness\<close>
theory Path_Connected
imports
Starlike
T1_Spaces
begin
subsection \<open>Paths and Arcs\<close>
definition\<^marker>\<open>tag important\<close> path :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool"
where "path g \<longleftrightarrow> continuous_on {0..1} g"
definition\<^marker>\<open>tag important\<close> pathstart :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> 'a"
where "pathstart g = g 0"
definition\<^marker>\<open>tag important\<close> pathfinish :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> 'a"
where "pathfinish g = g 1"
definition\<^marker>\<open>tag important\<close> path_image :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> 'a set"
where "path_image g = g ` {0 .. 1}"
definition\<^marker>\<open>tag important\<close> reversepath :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> real \<Rightarrow> 'a"
where "reversepath g = (\<lambda>x. g(1 - x))"
definition\<^marker>\<open>tag important\<close> joinpaths :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> real \<Rightarrow> 'a"
(infixr "+++" 75)
where "g1 +++ g2 = (\<lambda>x. if x \<le> 1/2 then g1 (2 * x) else g2 (2 * x - 1))"
definition\<^marker>\<open>tag important\<close> simple_path :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool"
where "simple_path g \<longleftrightarrow>
path g \<and> (\<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. g x = g y \<longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0)"
definition\<^marker>\<open>tag important\<close> arc :: "(real \<Rightarrow> 'a :: topological_space) \<Rightarrow> bool"
where "arc g \<longleftrightarrow> path g \<and> inj_on g {0..1}"
subsection\<^marker>\<open>tag unimportant\<close>\<open>Invariance theorems\<close>
lemma path_eq: "path p \<Longrightarrow> (\<And>t. t \<in> {0..1} \<Longrightarrow> p t = q t) \<Longrightarrow> path q"
using continuous_on_eq path_def by blast
lemma path_continuous_image: "path g \<Longrightarrow> continuous_on (path_image g) f \<Longrightarrow> path(f \<circ> g)"
unfolding path_def path_image_def
using continuous_on_compose by blast
lemma path_translation_eq:
fixes g :: "real \<Rightarrow> 'a :: real_normed_vector"
shows "path((\<lambda>x. a + x) \<circ> g) = path g"
proof -
have g: "g = (\<lambda>x. -a + x) \<circ> ((\<lambda>x. a + x) \<circ> g)"
by (rule ext) simp
show ?thesis
unfolding path_def
apply safe
apply (subst g)
apply (rule continuous_on_compose)
apply (auto intro: continuous_intros)
done
qed
lemma path_linear_image_eq:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "linear f" "inj f"
shows "path(f \<circ> g) = path g"
proof -
from linear_injective_left_inverse [OF assms]
obtain h where h: "linear h" "h \<circ> f = id"
by blast
then have g: "g = h \<circ> (f \<circ> g)"
by (metis comp_assoc id_comp)
show ?thesis
unfolding path_def
using h assms
by (metis g continuous_on_compose linear_continuous_on linear_conv_bounded_linear)
qed
lemma pathstart_translation: "pathstart((\<lambda>x. a + x) \<circ> g) = a + pathstart g"
by (simp add: pathstart_def)
lemma pathstart_linear_image_eq: "linear f \<Longrightarrow> pathstart(f \<circ> g) = f(pathstart g)"
by (simp add: pathstart_def)
lemma pathfinish_translation: "pathfinish((\<lambda>x. a + x) \<circ> g) = a + pathfinish g"
by (simp add: pathfinish_def)
lemma pathfinish_linear_image: "linear f \<Longrightarrow> pathfinish(f \<circ> g) = f(pathfinish g)"
by (simp add: pathfinish_def)
lemma path_image_translation: "path_image((\<lambda>x. a + x) \<circ> g) = (\<lambda>x. a + x) ` (path_image g)"
by (simp add: image_comp path_image_def)
lemma path_image_linear_image: "linear f \<Longrightarrow> path_image(f \<circ> g) = f ` (path_image g)"
by (simp add: image_comp path_image_def)
lemma reversepath_translation: "reversepath((\<lambda>x. a + x) \<circ> g) = (\<lambda>x. a + x) \<circ> reversepath g"
by (rule ext) (simp add: reversepath_def)
lemma reversepath_linear_image: "linear f \<Longrightarrow> reversepath(f \<circ> g) = f \<circ> reversepath g"
by (rule ext) (simp add: reversepath_def)
lemma joinpaths_translation:
"((\<lambda>x. a + x) \<circ> g1) +++ ((\<lambda>x. a + x) \<circ> g2) = (\<lambda>x. a + x) \<circ> (g1 +++ g2)"
by (rule ext) (simp add: joinpaths_def)
lemma joinpaths_linear_image: "linear f \<Longrightarrow> (f \<circ> g1) +++ (f \<circ> g2) = f \<circ> (g1 +++ g2)"
by (rule ext) (simp add: joinpaths_def)
lemma simple_path_translation_eq:
fixes g :: "real \<Rightarrow> 'a::euclidean_space"
shows "simple_path((\<lambda>x. a + x) \<circ> g) = simple_path g"
by (simp add: simple_path_def path_translation_eq)
lemma simple_path_linear_image_eq:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "linear f" "inj f"
shows "simple_path(f \<circ> g) = simple_path g"
using assms inj_on_eq_iff [of f]
by (auto simp: path_linear_image_eq simple_path_def path_translation_eq)
lemma arc_translation_eq:
fixes g :: "real \<Rightarrow> 'a::euclidean_space"
shows "arc((\<lambda>x. a + x) \<circ> g) = arc g"
by (auto simp: arc_def inj_on_def path_translation_eq)
lemma arc_linear_image_eq:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "linear f" "inj f"
shows "arc(f \<circ> g) = arc g"
using assms inj_on_eq_iff [of f]
by (auto simp: arc_def inj_on_def path_linear_image_eq)
subsection\<^marker>\<open>tag unimportant\<close>\<open>Basic lemmas about paths\<close>
lemma pathin_iff_path_real [simp]: "pathin euclideanreal g \<longleftrightarrow> path g"
by (simp add: pathin_def path_def)
lemma continuous_on_path: "path f \<Longrightarrow> t \<subseteq> {0..1} \<Longrightarrow> continuous_on t f"
using continuous_on_subset path_def by blast
lemma arc_imp_simple_path: "arc g \<Longrightarrow> simple_path g"
by (simp add: arc_def inj_on_def simple_path_def)
lemma arc_imp_path: "arc g \<Longrightarrow> path g"
using arc_def by blast
lemma arc_imp_inj_on: "arc g \<Longrightarrow> inj_on g {0..1}"
by (auto simp: arc_def)
lemma simple_path_imp_path: "simple_path g \<Longrightarrow> path g"
using simple_path_def by blast
lemma simple_path_cases: "simple_path g \<Longrightarrow> arc g \<or> pathfinish g = pathstart g"
unfolding simple_path_def arc_def inj_on_def pathfinish_def pathstart_def
by force
lemma simple_path_imp_arc: "simple_path g \<Longrightarrow> pathfinish g \<noteq> pathstart g \<Longrightarrow> arc g"
using simple_path_cases by auto
lemma arc_distinct_ends: "arc g \<Longrightarrow> pathfinish g \<noteq> pathstart g"
unfolding arc_def inj_on_def pathfinish_def pathstart_def
by fastforce
lemma arc_simple_path: "arc g \<longleftrightarrow> simple_path g \<and> pathfinish g \<noteq> pathstart g"
using arc_distinct_ends arc_imp_simple_path simple_path_cases by blast
lemma simple_path_eq_arc: "pathfinish g \<noteq> pathstart g \<Longrightarrow> (simple_path g = arc g)"
by (simp add: arc_simple_path)
lemma path_image_const [simp]: "path_image (\<lambda>t. a) = {a}"
by (force simp: path_image_def)
lemma path_image_nonempty [simp]: "path_image g \<noteq> {}"
unfolding path_image_def image_is_empty box_eq_empty
by auto
lemma pathstart_in_path_image[intro]: "pathstart g \<in> path_image g"
unfolding pathstart_def path_image_def
by auto
lemma pathfinish_in_path_image[intro]: "pathfinish g \<in> path_image g"
unfolding pathfinish_def path_image_def
by auto
lemma connected_path_image[intro]: "path g \<Longrightarrow> connected (path_image g)"
unfolding path_def path_image_def
using connected_continuous_image connected_Icc by blast
lemma compact_path_image[intro]: "path g \<Longrightarrow> compact (path_image g)"
unfolding path_def path_image_def
using compact_continuous_image connected_Icc by blast
lemma reversepath_reversepath[simp]: "reversepath (reversepath g) = g"
unfolding reversepath_def
by auto
lemma pathstart_reversepath[simp]: "pathstart (reversepath g) = pathfinish g"
unfolding pathstart_def reversepath_def pathfinish_def
by auto
lemma pathfinish_reversepath[simp]: "pathfinish (reversepath g) = pathstart g"
unfolding pathstart_def reversepath_def pathfinish_def
by auto
lemma pathstart_join[simp]: "pathstart (g1 +++ g2) = pathstart g1"
unfolding pathstart_def joinpaths_def pathfinish_def
by auto
lemma pathfinish_join[simp]: "pathfinish (g1 +++ g2) = pathfinish g2"
unfolding pathstart_def joinpaths_def pathfinish_def
by auto
lemma path_image_reversepath[simp]: "path_image (reversepath g) = path_image g"
proof -
have *: "\<And>g. path_image (reversepath g) \<subseteq> path_image g"
unfolding path_image_def subset_eq reversepath_def Ball_def image_iff
by force
show ?thesis
using *[of g] *[of "reversepath g"]
unfolding reversepath_reversepath
by auto
qed
lemma path_reversepath [simp]: "path (reversepath g) \<longleftrightarrow> path g"
proof -
have *: "\<And>g. path g \<Longrightarrow> path (reversepath g)"
unfolding path_def reversepath_def
apply (rule continuous_on_compose[unfolded o_def, of _ "\<lambda>x. 1 - x"])
apply (auto intro: continuous_intros continuous_on_subset[of "{0..1}"])
done
show ?thesis
using *[of "reversepath g"] *[of g]
unfolding reversepath_reversepath
by (rule iffI)
qed
lemma arc_reversepath:
assumes "arc g" shows "arc(reversepath g)"
proof -
have injg: "inj_on g {0..1}"
using assms
by (simp add: arc_def)
have **: "\<And>x y::real. 1-x = 1-y \<Longrightarrow> x = y"
by simp
show ?thesis
using assms by (clarsimp simp: arc_def intro!: inj_onI) (simp add: inj_onD reversepath_def **)
qed
lemma simple_path_reversepath: "simple_path g \<Longrightarrow> simple_path (reversepath g)"
apply (simp add: simple_path_def)
apply (force simp: reversepath_def)
done
lemmas reversepath_simps =
path_reversepath path_image_reversepath pathstart_reversepath pathfinish_reversepath
lemma path_join[simp]:
assumes "pathfinish g1 = pathstart g2"
shows "path (g1 +++ g2) \<longleftrightarrow> path g1 \<and> path g2"
unfolding path_def pathfinish_def pathstart_def
proof safe
assume cont: "continuous_on {0..1} (g1 +++ g2)"
have g1: "continuous_on {0..1} g1 \<longleftrightarrow> continuous_on {0..1} ((g1 +++ g2) \<circ> (\<lambda>x. x / 2))"
by (intro continuous_on_cong refl) (auto simp: joinpaths_def)
have g2: "continuous_on {0..1} g2 \<longleftrightarrow> continuous_on {0..1} ((g1 +++ g2) \<circ> (\<lambda>x. x / 2 + 1/2))"
using assms
by (intro continuous_on_cong refl) (auto simp: joinpaths_def pathfinish_def pathstart_def)
show "continuous_on {0..1} g1" and "continuous_on {0..1} g2"
unfolding g1 g2
by (auto intro!: continuous_intros continuous_on_subset[OF cont] simp del: o_apply)
next
assume g1g2: "continuous_on {0..1} g1" "continuous_on {0..1} g2"
have 01: "{0 .. 1} = {0..1/2} \<union> {1/2 .. 1::real}"
by auto
{
fix x :: real
assume "0 \<le> x" and "x \<le> 1"
then have "x \<in> (\<lambda>x. x * 2) ` {0..1 / 2}"
by (intro image_eqI[where x="x/2"]) auto
}
note 1 = this
{
fix x :: real
assume "0 \<le> x" and "x \<le> 1"
then have "x \<in> (\<lambda>x. x * 2 - 1) ` {1 / 2..1}"
by (intro image_eqI[where x="x/2 + 1/2"]) auto
}
note 2 = this
show "continuous_on {0..1} (g1 +++ g2)"
using assms
unfolding joinpaths_def 01
apply (intro continuous_on_cases closed_atLeastAtMost g1g2[THEN continuous_on_compose2] continuous_intros)
apply (auto simp: field_simps pathfinish_def pathstart_def intro!: 1 2)
done
qed
subsection\<^marker>\<open>tag unimportant\<close> \<open>Path Images\<close>
lemma bounded_path_image: "path g \<Longrightarrow> bounded(path_image g)"
by (simp add: compact_imp_bounded compact_path_image)
lemma closed_path_image:
fixes g :: "real \<Rightarrow> 'a::t2_space"
shows "path g \<Longrightarrow> closed(path_image g)"
by (metis compact_path_image compact_imp_closed)
lemma connected_simple_path_image: "simple_path g \<Longrightarrow> connected(path_image g)"
by (metis connected_path_image simple_path_imp_path)
lemma compact_simple_path_image: "simple_path g \<Longrightarrow> compact(path_image g)"
by (metis compact_path_image simple_path_imp_path)
lemma bounded_simple_path_image: "simple_path g \<Longrightarrow> bounded(path_image g)"
by (metis bounded_path_image simple_path_imp_path)
lemma closed_simple_path_image:
fixes g :: "real \<Rightarrow> 'a::t2_space"
shows "simple_path g \<Longrightarrow> closed(path_image g)"
by (metis closed_path_image simple_path_imp_path)
lemma connected_arc_image: "arc g \<Longrightarrow> connected(path_image g)"
by (metis connected_path_image arc_imp_path)
lemma compact_arc_image: "arc g \<Longrightarrow> compact(path_image g)"
by (metis compact_path_image arc_imp_path)
lemma bounded_arc_image: "arc g \<Longrightarrow> bounded(path_image g)"
by (metis bounded_path_image arc_imp_path)
lemma closed_arc_image:
fixes g :: "real \<Rightarrow> 'a::t2_space"
shows "arc g \<Longrightarrow> closed(path_image g)"
by (metis closed_path_image arc_imp_path)
lemma path_image_join_subset: "path_image (g1 +++ g2) \<subseteq> path_image g1 \<union> path_image g2"
unfolding path_image_def joinpaths_def
by auto
lemma subset_path_image_join:
assumes "path_image g1 \<subseteq> s"
and "path_image g2 \<subseteq> s"
shows "path_image (g1 +++ g2) \<subseteq> s"
using path_image_join_subset[of g1 g2] and assms
by auto
lemma path_image_join:
"pathfinish g1 = pathstart g2 \<Longrightarrow> path_image(g1 +++ g2) = path_image g1 \<union> path_image g2"
apply (rule subset_antisym [OF path_image_join_subset])
apply (auto simp: pathfinish_def pathstart_def path_image_def joinpaths_def image_def)
apply (drule sym)
apply (rule_tac x="xa/2" in bexI, auto)
apply (rule ccontr)
apply (drule_tac x="(xa+1)/2" in bspec)
apply (auto simp: field_simps)
apply (drule_tac x="1/2" in bspec, auto)
done
lemma not_in_path_image_join:
assumes "x \<notin> path_image g1"
and "x \<notin> path_image g2"
shows "x \<notin> path_image (g1 +++ g2)"
using assms and path_image_join_subset[of g1 g2]
by auto
lemma pathstart_compose: "pathstart(f \<circ> p) = f(pathstart p)"
by (simp add: pathstart_def)
lemma pathfinish_compose: "pathfinish(f \<circ> p) = f(pathfinish p)"
by (simp add: pathfinish_def)
lemma path_image_compose: "path_image (f \<circ> p) = f ` (path_image p)"
by (simp add: image_comp path_image_def)
lemma path_compose_join: "f \<circ> (p +++ q) = (f \<circ> p) +++ (f \<circ> q)"
by (rule ext) (simp add: joinpaths_def)
lemma path_compose_reversepath: "f \<circ> reversepath p = reversepath(f \<circ> p)"
by (rule ext) (simp add: reversepath_def)
lemma joinpaths_eq:
"(\<And>t. t \<in> {0..1} \<Longrightarrow> p t = p' t) \<Longrightarrow>
(\<And>t. t \<in> {0..1} \<Longrightarrow> q t = q' t)
\<Longrightarrow> t \<in> {0..1} \<Longrightarrow> (p +++ q) t = (p' +++ q') t"
by (auto simp: joinpaths_def)
lemma simple_path_inj_on: "simple_path g \<Longrightarrow> inj_on g {0<..<1}"
by (auto simp: simple_path_def path_image_def inj_on_def less_eq_real_def Ball_def)
subsection\<^marker>\<open>tag unimportant\<close>\<open>Simple paths with the endpoints removed\<close>
lemma simple_path_endless:
"simple_path c \<Longrightarrow> path_image c - {pathstart c,pathfinish c} = c ` {0<..<1}"
apply (auto simp: simple_path_def path_image_def pathstart_def pathfinish_def Ball_def Bex_def image_def)
apply (metis eq_iff le_less_linear)
apply (metis leD linear)
using less_eq_real_def zero_le_one apply blast
using less_eq_real_def zero_le_one apply blast
done
lemma connected_simple_path_endless:
"simple_path c \<Longrightarrow> connected(path_image c - {pathstart c,pathfinish c})"
apply (simp add: simple_path_endless)
apply (rule connected_continuous_image)
apply (meson continuous_on_subset greaterThanLessThan_subseteq_atLeastAtMost_iff le_numeral_extra(3) le_numeral_extra(4) path_def simple_path_imp_path)
by auto
lemma nonempty_simple_path_endless:
"simple_path c \<Longrightarrow> path_image c - {pathstart c,pathfinish c} \<noteq> {}"
by (simp add: simple_path_endless)
subsection\<^marker>\<open>tag unimportant\<close>\<open>The operations on paths\<close>
lemma path_image_subset_reversepath: "path_image(reversepath g) \<le> path_image g"
by (auto simp: path_image_def reversepath_def)
lemma path_imp_reversepath: "path g \<Longrightarrow> path(reversepath g)"
apply (auto simp: path_def reversepath_def)
using continuous_on_compose [of "{0..1}" "\<lambda>x. 1 - x" g]
apply (auto simp: continuous_on_op_minus)
done
lemma half_bounded_equal: "1 \<le> x * 2 \<Longrightarrow> x * 2 \<le> 1 \<longleftrightarrow> x = (1/2::real)"
by simp
lemma continuous_on_joinpaths:
assumes "continuous_on {0..1} g1" "continuous_on {0..1} g2" "pathfinish g1 = pathstart g2"
shows "continuous_on {0..1} (g1 +++ g2)"
proof -
have *: "{0..1::real} = {0..1/2} \<union> {1/2..1}"
by auto
have gg: "g2 0 = g1 1"
by (metis assms(3) pathfinish_def pathstart_def)
have 1: "continuous_on {0..1/2} (g1 +++ g2)"
apply (rule continuous_on_eq [of _ "g1 \<circ> (\<lambda>x. 2*x)"])
apply (rule continuous_intros | simp add: joinpaths_def assms)+
done
have "continuous_on {1/2..1} (g2 \<circ> (\<lambda>x. 2*x-1))"
apply (rule continuous_on_subset [of "{1/2..1}"])
apply (rule continuous_intros | simp add: image_affinity_atLeastAtMost_diff assms)+
done
then have 2: "continuous_on {1/2..1} (g1 +++ g2)"
apply (rule continuous_on_eq [of "{1/2..1}" "g2 \<circ> (\<lambda>x. 2*x-1)"])
apply (rule assms continuous_intros | simp add: joinpaths_def mult.commute half_bounded_equal gg)+
done
show ?thesis
apply (subst *)
apply (rule continuous_on_closed_Un)
using 1 2
apply auto
done
qed
lemma path_join_imp: "\<lbrakk>path g1; path g2; pathfinish g1 = pathstart g2\<rbrakk> \<Longrightarrow> path(g1 +++ g2)"
by (simp)
lemma simple_path_join_loop:
assumes "arc g1" "arc g2"
"pathfinish g1 = pathstart g2" "pathfinish g2 = pathstart g1"
"path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g1, pathstart g2}"
shows "simple_path(g1 +++ g2)"
proof -
have injg1: "inj_on g1 {0..1}"
using assms
by (simp add: arc_def)
have injg2: "inj_on g2 {0..1}"
using assms
by (simp add: arc_def)
have g12: "g1 1 = g2 0"
and g21: "g2 1 = g1 0"
and sb: "g1 ` {0..1} \<inter> g2 ` {0..1} \<subseteq> {g1 0, g2 0}"
using assms
by (simp_all add: arc_def pathfinish_def pathstart_def path_image_def)
{ fix x and y::real
assume xyI: "x = 1 \<longrightarrow> y \<noteq> 0"
and xy: "x \<le> 1" "0 \<le> y" " y * 2 \<le> 1" "\<not> x * 2 \<le> 1" "g2 (2 * x - 1) = g1 (2 * y)"
have g1im: "g1 (2 * y) \<in> g1 ` {0..1} \<inter> g2 ` {0..1}"
using xy
apply simp
apply (rule_tac x="2 * x - 1" in image_eqI, auto)
done
have False
using subsetD [OF sb g1im] xy
apply auto
apply (drule inj_onD [OF injg1])
using g21 [symmetric] xyI
apply (auto dest: inj_onD [OF injg2])
done
} note * = this
{ fix x and y::real
assume xy: "y \<le> 1" "0 \<le> x" "\<not> y * 2 \<le> 1" "x * 2 \<le> 1" "g1 (2 * x) = g2 (2 * y - 1)"
have g1im: "g1 (2 * x) \<in> g1 ` {0..1} \<inter> g2 ` {0..1}"
using xy
apply simp
apply (rule_tac x="2 * x" in image_eqI, auto)
done
have "x = 0 \<and> y = 1"
using subsetD [OF sb g1im] xy
apply auto
apply (force dest: inj_onD [OF injg1])
using g21 [symmetric]
apply (auto dest: inj_onD [OF injg2])
done
} note ** = this
show ?thesis
using assms
- apply (simp add: arc_def simple_path_def path_join, clarify)
+ apply (simp add: arc_def simple_path_def, clarify)
apply (simp add: joinpaths_def split: if_split_asm)
apply (force dest: inj_onD [OF injg1])
apply (metis *)
apply (metis **)
apply (force dest: inj_onD [OF injg2])
done
qed
lemma arc_join:
assumes "arc g1" "arc g2"
"pathfinish g1 = pathstart g2"
"path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g2}"
shows "arc(g1 +++ g2)"
proof -
have injg1: "inj_on g1 {0..1}"
using assms
by (simp add: arc_def)
have injg2: "inj_on g2 {0..1}"
using assms
by (simp add: arc_def)
have g11: "g1 1 = g2 0"
and sb: "g1 ` {0..1} \<inter> g2 ` {0..1} \<subseteq> {g2 0}"
using assms
by (simp_all add: arc_def pathfinish_def pathstart_def path_image_def)
{ fix x and y::real
assume xy: "x \<le> 1" "0 \<le> y" " y * 2 \<le> 1" "\<not> x * 2 \<le> 1" "g2 (2 * x - 1) = g1 (2 * y)"
have g1im: "g1 (2 * y) \<in> g1 ` {0..1} \<inter> g2 ` {0..1}"
using xy
apply simp
apply (rule_tac x="2 * x - 1" in image_eqI, auto)
done
have False
using subsetD [OF sb g1im] xy
by (auto dest: inj_onD [OF injg2])
} note * = this
show ?thesis
apply (simp add: arc_def inj_on_def)
apply (clarsimp simp add: arc_imp_path assms)
apply (simp add: joinpaths_def split: if_split_asm)
apply (force dest: inj_onD [OF injg1])
apply (metis *)
apply (metis *)
apply (force dest: inj_onD [OF injg2])
done
qed
lemma reversepath_joinpaths:
"pathfinish g1 = pathstart g2 \<Longrightarrow> reversepath(g1 +++ g2) = reversepath g2 +++ reversepath g1"
unfolding reversepath_def pathfinish_def pathstart_def joinpaths_def
by (rule ext) (auto simp: mult.commute)
subsection\<^marker>\<open>tag unimportant\<close>\<open>Some reversed and "if and only if" versions of joining theorems\<close>
lemma path_join_path_ends:
fixes g1 :: "real \<Rightarrow> 'a::metric_space"
assumes "path(g1 +++ g2)" "path g2"
shows "pathfinish g1 = pathstart g2"
proof (rule ccontr)
define e where "e = dist (g1 1) (g2 0)"
assume Neg: "pathfinish g1 \<noteq> pathstart g2"
then have "0 < dist (pathfinish g1) (pathstart g2)"
by auto
then have "e > 0"
by (metis e_def pathfinish_def pathstart_def)
then obtain d1 where "d1 > 0"
and d1: "\<And>x'. \<lbrakk>x'\<in>{0..1}; norm x' < d1\<rbrakk> \<Longrightarrow> dist (g2 x') (g2 0) < e/2"
using assms(2) unfolding path_def continuous_on_iff
apply (drule_tac x=0 in bspec, simp)
by (metis half_gt_zero_iff norm_conv_dist)
obtain d2 where "d2 > 0"
and d2: "\<And>x'. \<lbrakk>x'\<in>{0..1}; dist x' (1/2) < d2\<rbrakk>
\<Longrightarrow> dist ((g1 +++ g2) x') (g1 1) < e/2"
using assms(1) \<open>e > 0\<close> unfolding path_def continuous_on_iff
apply (drule_tac x="1/2" in bspec, simp)
apply (drule_tac x="e/2" in spec)
apply (force simp: joinpaths_def)
done
have int01_1: "min (1/2) (min d1 d2) / 2 \<in> {0..1}"
using \<open>d1 > 0\<close> \<open>d2 > 0\<close> by (simp add: min_def)
have dist1: "norm (min (1 / 2) (min d1 d2) / 2) < d1"
using \<open>d1 > 0\<close> \<open>d2 > 0\<close> by (simp add: min_def dist_norm)
have int01_2: "1/2 + min (1/2) (min d1 d2) / 4 \<in> {0..1}"
using \<open>d1 > 0\<close> \<open>d2 > 0\<close> by (simp add: min_def)
have dist2: "dist (1 / 2 + min (1 / 2) (min d1 d2) / 4) (1 / 2) < d2"
using \<open>d1 > 0\<close> \<open>d2 > 0\<close> by (simp add: min_def dist_norm)
have [simp]: "\<not> min (1 / 2) (min d1 d2) \<le> 0"
using \<open>d1 > 0\<close> \<open>d2 > 0\<close> by (simp add: min_def)
have "dist (g2 (min (1 / 2) (min d1 d2) / 2)) (g1 1) < e/2"
"dist (g2 (min (1 / 2) (min d1 d2) / 2)) (g2 0) < e/2"
using d1 [OF int01_1 dist1] d2 [OF int01_2 dist2] by (simp_all add: joinpaths_def)
then have "dist (g1 1) (g2 0) < e/2 + e/2"
using dist_triangle_half_r e_def by blast
then show False
by (simp add: e_def [symmetric])
qed
lemma path_join_eq [simp]:
fixes g1 :: "real \<Rightarrow> 'a::metric_space"
assumes "path g1" "path g2"
shows "path(g1 +++ g2) \<longleftrightarrow> pathfinish g1 = pathstart g2"
using assms by (metis path_join_path_ends path_join_imp)
lemma simple_path_joinE:
assumes "simple_path(g1 +++ g2)" and "pathfinish g1 = pathstart g2"
obtains "arc g1" "arc g2"
"path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g1, pathstart g2}"
proof -
have *: "\<And>x y. \<lbrakk>0 \<le> x; x \<le> 1; 0 \<le> y; y \<le> 1; (g1 +++ g2) x = (g1 +++ g2) y\<rbrakk>
\<Longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0"
using assms by (simp add: simple_path_def)
have "path g1"
using assms path_join simple_path_imp_path by blast
moreover have "inj_on g1 {0..1}"
proof (clarsimp simp: inj_on_def)
fix x y
assume "g1 x = g1 y" "0 \<le> x" "x \<le> 1" "0 \<le> y" "y \<le> 1"
then show "x = y"
using * [of "x/2" "y/2"] by (simp add: joinpaths_def split_ifs)
qed
ultimately have "arc g1"
using assms by (simp add: arc_def)
have [simp]: "g2 0 = g1 1"
using assms by (metis pathfinish_def pathstart_def)
have "path g2"
using assms path_join simple_path_imp_path by blast
moreover have "inj_on g2 {0..1}"
proof (clarsimp simp: inj_on_def)
fix x y
assume "g2 x = g2 y" "0 \<le> x" "x \<le> 1" "0 \<le> y" "y \<le> 1"
then show "x = y"
using * [of "(x + 1) / 2" "(y + 1) / 2"]
by (force simp: joinpaths_def split_ifs field_split_simps)
qed
ultimately have "arc g2"
using assms by (simp add: arc_def)
have "g2 y = g1 0 \<or> g2 y = g1 1"
if "g1 x = g2 y" "0 \<le> x" "x \<le> 1" "0 \<le> y" "y \<le> 1" for x y
using * [of "x / 2" "(y + 1) / 2"] that
by (auto simp: joinpaths_def split_ifs field_split_simps)
then have "path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g1, pathstart g2}"
by (fastforce simp: pathstart_def pathfinish_def path_image_def)
with \<open>arc g1\<close> \<open>arc g2\<close> show ?thesis using that by blast
qed
lemma simple_path_join_loop_eq:
assumes "pathfinish g2 = pathstart g1" "pathfinish g1 = pathstart g2"
shows "simple_path(g1 +++ g2) \<longleftrightarrow>
arc g1 \<and> arc g2 \<and> path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g1, pathstart g2}"
by (metis assms simple_path_joinE simple_path_join_loop)
lemma arc_join_eq:
assumes "pathfinish g1 = pathstart g2"
shows "arc(g1 +++ g2) \<longleftrightarrow>
arc g1 \<and> arc g2 \<and> path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g2}"
(is "?lhs = ?rhs")
proof
assume ?lhs
then have "simple_path(g1 +++ g2)" by (rule arc_imp_simple_path)
then have *: "\<And>x y. \<lbrakk>0 \<le> x; x \<le> 1; 0 \<le> y; y \<le> 1; (g1 +++ g2) x = (g1 +++ g2) y\<rbrakk>
\<Longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0"
using assms by (simp add: simple_path_def)
have False if "g1 0 = g2 u" "0 \<le> u" "u \<le> 1" for u
using * [of 0 "(u + 1) / 2"] that assms arc_distinct_ends [OF \<open>?lhs\<close>]
by (auto simp: joinpaths_def pathstart_def pathfinish_def split_ifs field_split_simps)
then have n1: "pathstart g1 \<notin> path_image g2"
unfolding pathstart_def path_image_def
using atLeastAtMost_iff by blast
show ?rhs using \<open>?lhs\<close>
apply (rule simple_path_joinE [OF arc_imp_simple_path assms])
using n1 by force
next
assume ?rhs then show ?lhs
using assms
by (fastforce simp: pathfinish_def pathstart_def intro!: arc_join)
qed
lemma arc_join_eq_alt:
"pathfinish g1 = pathstart g2
\<Longrightarrow> (arc(g1 +++ g2) \<longleftrightarrow>
arc g1 \<and> arc g2 \<and>
path_image g1 \<inter> path_image g2 = {pathstart g2})"
using pathfinish_in_path_image by (fastforce simp: arc_join_eq)
subsection\<^marker>\<open>tag unimportant\<close>\<open>The joining of paths is associative\<close>
lemma path_assoc:
"\<lbrakk>pathfinish p = pathstart q; pathfinish q = pathstart r\<rbrakk>
\<Longrightarrow> path(p +++ (q +++ r)) \<longleftrightarrow> path((p +++ q) +++ r)"
by simp
lemma simple_path_assoc:
assumes "pathfinish p = pathstart q" "pathfinish q = pathstart r"
shows "simple_path (p +++ (q +++ r)) \<longleftrightarrow> simple_path ((p +++ q) +++ r)"
proof (cases "pathstart p = pathfinish r")
case True show ?thesis
proof
assume "simple_path (p +++ q +++ r)"
with assms True show "simple_path ((p +++ q) +++ r)"
by (fastforce simp add: simple_path_join_loop_eq arc_join_eq path_image_join
dest: arc_distinct_ends [of r])
next
assume 0: "simple_path ((p +++ q) +++ r)"
with assms True have q: "pathfinish r \<notin> path_image q"
using arc_distinct_ends
by (fastforce simp add: simple_path_join_loop_eq arc_join_eq path_image_join)
have "pathstart r \<notin> path_image p"
using assms
by (metis 0 IntI arc_distinct_ends arc_join_eq_alt empty_iff insert_iff
pathfinish_in_path_image pathfinish_join simple_path_joinE)
with assms 0 q True show "simple_path (p +++ q +++ r)"
by (auto simp: simple_path_join_loop_eq arc_join_eq path_image_join
dest!: subsetD [OF _ IntI])
qed
next
case False
{ fix x :: 'a
assume a: "path_image p \<inter> path_image q \<subseteq> {pathstart q}"
"(path_image p \<union> path_image q) \<inter> path_image r \<subseteq> {pathstart r}"
"x \<in> path_image p" "x \<in> path_image r"
have "pathstart r \<in> path_image q"
by (metis assms(2) pathfinish_in_path_image)
with a have "x = pathstart q"
by blast
}
with False assms show ?thesis
by (auto simp: simple_path_eq_arc simple_path_join_loop_eq arc_join_eq path_image_join)
qed
lemma arc_assoc:
"\<lbrakk>pathfinish p = pathstart q; pathfinish q = pathstart r\<rbrakk>
\<Longrightarrow> arc(p +++ (q +++ r)) \<longleftrightarrow> arc((p +++ q) +++ r)"
by (simp add: arc_simple_path simple_path_assoc)
subsubsection\<^marker>\<open>tag unimportant\<close>\<open>Symmetry and loops\<close>
lemma path_sym:
"\<lbrakk>pathfinish p = pathstart q; pathfinish q = pathstart p\<rbrakk> \<Longrightarrow> path(p +++ q) \<longleftrightarrow> path(q +++ p)"
by auto
lemma simple_path_sym:
"\<lbrakk>pathfinish p = pathstart q; pathfinish q = pathstart p\<rbrakk>
\<Longrightarrow> simple_path(p +++ q) \<longleftrightarrow> simple_path(q +++ p)"
by (metis (full_types) inf_commute insert_commute simple_path_joinE simple_path_join_loop)
lemma path_image_sym:
"\<lbrakk>pathfinish p = pathstart q; pathfinish q = pathstart p\<rbrakk>
\<Longrightarrow> path_image(p +++ q) = path_image(q +++ p)"
by (simp add: path_image_join sup_commute)
subsection\<open>Subpath\<close>
definition\<^marker>\<open>tag important\<close> subpath :: "real \<Rightarrow> real \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> real \<Rightarrow> 'a::real_normed_vector"
where "subpath a b g \<equiv> \<lambda>x. g((b - a) * x + a)"
lemma path_image_subpath_gen:
fixes g :: "_ \<Rightarrow> 'a::real_normed_vector"
shows "path_image(subpath u v g) = g ` (closed_segment u v)"
by (auto simp add: closed_segment_real_eq path_image_def subpath_def)
lemma path_image_subpath:
fixes g :: "real \<Rightarrow> 'a::real_normed_vector"
shows "path_image(subpath u v g) = (if u \<le> v then g ` {u..v} else g ` {v..u})"
by (simp add: path_image_subpath_gen closed_segment_eq_real_ivl)
lemma path_image_subpath_commute:
fixes g :: "real \<Rightarrow> 'a::real_normed_vector"
shows "path_image(subpath u v g) = path_image(subpath v u g)"
by (simp add: path_image_subpath_gen closed_segment_eq_real_ivl)
lemma path_subpath [simp]:
fixes g :: "real \<Rightarrow> 'a::real_normed_vector"
assumes "path g" "u \<in> {0..1}" "v \<in> {0..1}"
shows "path(subpath u v g)"
proof -
have "continuous_on {0..1} (g \<circ> (\<lambda>x. ((v-u) * x+ u)))"
apply (rule continuous_intros | simp)+
apply (simp add: image_affinity_atLeastAtMost [where c=u])
using assms
apply (auto simp: path_def continuous_on_subset)
done
then show ?thesis
by (simp add: path_def subpath_def)
qed
lemma pathstart_subpath [simp]: "pathstart(subpath u v g) = g(u)"
by (simp add: pathstart_def subpath_def)
lemma pathfinish_subpath [simp]: "pathfinish(subpath u v g) = g(v)"
by (simp add: pathfinish_def subpath_def)
lemma subpath_trivial [simp]: "subpath 0 1 g = g"
by (simp add: subpath_def)
lemma subpath_reversepath: "subpath 1 0 g = reversepath g"
by (simp add: reversepath_def subpath_def)
lemma reversepath_subpath: "reversepath(subpath u v g) = subpath v u g"
by (simp add: reversepath_def subpath_def algebra_simps)
lemma subpath_translation: "subpath u v ((\<lambda>x. a + x) \<circ> g) = (\<lambda>x. a + x) \<circ> subpath u v g"
by (rule ext) (simp add: subpath_def)
lemma subpath_image: "subpath u v (f \<circ> g) = f \<circ> subpath u v g"
by (rule ext) (simp add: subpath_def)
lemma affine_ineq:
fixes x :: "'a::linordered_idom"
assumes "x \<le> 1" "v \<le> u"
shows "v + x * u \<le> u + x * v"
proof -
have "(1-x)*(u-v) \<ge> 0"
using assms by auto
then show ?thesis
by (simp add: algebra_simps)
qed
lemma sum_le_prod1:
fixes a::real shows "\<lbrakk>a \<le> 1; b \<le> 1\<rbrakk> \<Longrightarrow> a + b \<le> 1 + a * b"
by (metis add.commute affine_ineq mult.right_neutral)
lemma simple_path_subpath_eq:
"simple_path(subpath u v g) \<longleftrightarrow>
path(subpath u v g) \<and> u\<noteq>v \<and>
(\<forall>x y. x \<in> closed_segment u v \<and> y \<in> closed_segment u v \<and> g x = g y
\<longrightarrow> x = y \<or> x = u \<and> y = v \<or> x = v \<and> y = u)"
(is "?lhs = ?rhs")
proof (rule iffI)
assume ?lhs
then have p: "path (\<lambda>x. g ((v - u) * x + u))"
and sim: "(\<And>x y. \<lbrakk>x\<in>{0..1}; y\<in>{0..1}; g ((v - u) * x + u) = g ((v - u) * y + u)\<rbrakk>
\<Longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0)"
by (auto simp: simple_path_def subpath_def)
{ fix x y
assume "x \<in> closed_segment u v" "y \<in> closed_segment u v" "g x = g y"
then have "x = y \<or> x = u \<and> y = v \<or> x = v \<and> y = u"
using sim [of "(x-u)/(v-u)" "(y-u)/(v-u)"] p
by (auto split: if_split_asm simp add: closed_segment_real_eq image_affinity_atLeastAtMost)
(simp_all add: field_split_simps)
} moreover
have "path(subpath u v g) \<and> u\<noteq>v"
using sim [of "1/3" "2/3"] p
by (auto simp: subpath_def)
ultimately show ?rhs
by metis
next
assume ?rhs
then
have d1: "\<And>x y. \<lbrakk>g x = g y; u \<le> x; x \<le> v; u \<le> y; y \<le> v\<rbrakk> \<Longrightarrow> x = y \<or> x = u \<and> y = v \<or> x = v \<and> y = u"
and d2: "\<And>x y. \<lbrakk>g x = g y; v \<le> x; x \<le> u; v \<le> y; y \<le> u\<rbrakk> \<Longrightarrow> x = y \<or> x = u \<and> y = v \<or> x = v \<and> y = u"
and ne: "u < v \<or> v < u"
and psp: "path (subpath u v g)"
by (auto simp: closed_segment_real_eq image_affinity_atLeastAtMost)
have [simp]: "\<And>x. u + x * v = v + x * u \<longleftrightarrow> u=v \<or> x=1"
by algebra
show ?lhs using psp ne
unfolding simple_path_def subpath_def
by (fastforce simp add: algebra_simps affine_ineq mult_left_mono crossproduct_eq dest: d1 d2)
qed
lemma arc_subpath_eq:
"arc(subpath u v g) \<longleftrightarrow> path(subpath u v g) \<and> u\<noteq>v \<and> inj_on g (closed_segment u v)"
(is "?lhs = ?rhs")
proof (rule iffI)
assume ?lhs
then have p: "path (\<lambda>x. g ((v - u) * x + u))"
and sim: "(\<And>x y. \<lbrakk>x\<in>{0..1}; y\<in>{0..1}; g ((v - u) * x + u) = g ((v - u) * y + u)\<rbrakk>
\<Longrightarrow> x = y)"
by (auto simp: arc_def inj_on_def subpath_def)
{ fix x y
assume "x \<in> closed_segment u v" "y \<in> closed_segment u v" "g x = g y"
then have "x = y"
using sim [of "(x-u)/(v-u)" "(y-u)/(v-u)"] p
by (cases "v = u")
(simp_all split: if_split_asm add: inj_on_def closed_segment_real_eq image_affinity_atLeastAtMost,
simp add: field_simps)
} moreover
have "path(subpath u v g) \<and> u\<noteq>v"
using sim [of "1/3" "2/3"] p
by (auto simp: subpath_def)
ultimately show ?rhs
unfolding inj_on_def
by metis
next
assume ?rhs
then
have d1: "\<And>x y. \<lbrakk>g x = g y; u \<le> x; x \<le> v; u \<le> y; y \<le> v\<rbrakk> \<Longrightarrow> x = y"
and d2: "\<And>x y. \<lbrakk>g x = g y; v \<le> x; x \<le> u; v \<le> y; y \<le> u\<rbrakk> \<Longrightarrow> x = y"
and ne: "u < v \<or> v < u"
and psp: "path (subpath u v g)"
by (auto simp: inj_on_def closed_segment_real_eq image_affinity_atLeastAtMost)
show ?lhs using psp ne
unfolding arc_def subpath_def inj_on_def
by (auto simp: algebra_simps affine_ineq mult_left_mono crossproduct_eq dest: d1 d2)
qed
lemma simple_path_subpath:
assumes "simple_path g" "u \<in> {0..1}" "v \<in> {0..1}" "u \<noteq> v"
shows "simple_path(subpath u v g)"
using assms
apply (simp add: simple_path_subpath_eq simple_path_imp_path)
apply (simp add: simple_path_def closed_segment_real_eq image_affinity_atLeastAtMost, fastforce)
done
lemma arc_simple_path_subpath:
"\<lbrakk>simple_path g; u \<in> {0..1}; v \<in> {0..1}; g u \<noteq> g v\<rbrakk> \<Longrightarrow> arc(subpath u v g)"
by (force intro: simple_path_subpath simple_path_imp_arc)
lemma arc_subpath_arc:
"\<lbrakk>arc g; u \<in> {0..1}; v \<in> {0..1}; u \<noteq> v\<rbrakk> \<Longrightarrow> arc(subpath u v g)"
by (meson arc_def arc_imp_simple_path arc_simple_path_subpath inj_onD)
lemma arc_simple_path_subpath_interior:
"\<lbrakk>simple_path g; u \<in> {0..1}; v \<in> {0..1}; u \<noteq> v; \<bar>u-v\<bar> < 1\<rbrakk> \<Longrightarrow> arc(subpath u v g)"
apply (rule arc_simple_path_subpath)
apply (force simp: simple_path_def)+
done
lemma path_image_subpath_subset:
"\<lbrakk>u \<in> {0..1}; v \<in> {0..1}\<rbrakk> \<Longrightarrow> path_image(subpath u v g) \<subseteq> path_image g"
apply (simp add: closed_segment_real_eq image_affinity_atLeastAtMost path_image_subpath)
apply (auto simp: path_image_def)
done
lemma join_subpaths_middle: "subpath (0) ((1 / 2)) p +++ subpath ((1 / 2)) 1 p = p"
by (rule ext) (simp add: joinpaths_def subpath_def field_split_simps)
subsection\<^marker>\<open>tag unimportant\<close>\<open>There is a subpath to the frontier\<close>
lemma subpath_to_frontier_explicit:
fixes S :: "'a::metric_space set"
assumes g: "path g" and "pathfinish g \<notin> S"
obtains u where "0 \<le> u" "u \<le> 1"
"\<And>x. 0 \<le> x \<and> x < u \<Longrightarrow> g x \<in> interior S"
"(g u \<notin> interior S)" "(u = 0 \<or> g u \<in> closure S)"
proof -
have gcon: "continuous_on {0..1} g" using g by (simp add: path_def)
then have com: "compact ({0..1} \<inter> {u. g u \<in> closure (- S)})"
apply (simp add: Int_commute [of "{0..1}"] compact_eq_bounded_closed closed_vimage_Int [unfolded vimage_def])
using compact_eq_bounded_closed apply fastforce
done
have "1 \<in> {u. g u \<in> closure (- S)}"
using assms by (simp add: pathfinish_def closure_def)
then have dis: "{0..1} \<inter> {u. g u \<in> closure (- S)} \<noteq> {}"
using atLeastAtMost_iff zero_le_one by blast
then obtain u where "0 \<le> u" "u \<le> 1" and gu: "g u \<in> closure (- S)"
and umin: "\<And>t. \<lbrakk>0 \<le> t; t \<le> 1; g t \<in> closure (- S)\<rbrakk> \<Longrightarrow> u \<le> t"
using compact_attains_inf [OF com dis] by fastforce
then have umin': "\<And>t. \<lbrakk>0 \<le> t; t \<le> 1; t < u\<rbrakk> \<Longrightarrow> g t \<in> S"
using closure_def by fastforce
{ assume "u \<noteq> 0"
then have "u > 0" using \<open>0 \<le> u\<close> by auto
{ fix e::real assume "e > 0"
obtain d where "d>0" and d: "\<And>x'. \<lbrakk>x' \<in> {0..1}; dist x' u \<le> d\<rbrakk> \<Longrightarrow> dist (g x') (g u) < e"
using continuous_onE [OF gcon _ \<open>e > 0\<close>] \<open>0 \<le> _\<close> \<open>_ \<le> 1\<close> atLeastAtMost_iff by auto
have *: "dist (max 0 (u - d / 2)) u \<le> d"
using \<open>0 \<le> u\<close> \<open>u \<le> 1\<close> \<open>d > 0\<close> by (simp add: dist_real_def)
have "\<exists>y\<in>S. dist y (g u) < e"
using \<open>0 < u\<close> \<open>u \<le> 1\<close> \<open>d > 0\<close>
by (force intro: d [OF _ *] umin')
}
then have "g u \<in> closure S"
by (simp add: frontier_def closure_approachable)
}
then show ?thesis
apply (rule_tac u=u in that)
apply (auto simp: \<open>0 \<le> u\<close> \<open>u \<le> 1\<close> gu interior_closure umin)
using \<open>_ \<le> 1\<close> interior_closure umin apply fastforce
done
qed
lemma subpath_to_frontier_strong:
assumes g: "path g" and "pathfinish g \<notin> S"
obtains u where "0 \<le> u" "u \<le> 1" "g u \<notin> interior S"
"u = 0 \<or> (\<forall>x. 0 \<le> x \<and> x < 1 \<longrightarrow> subpath 0 u g x \<in> interior S) \<and> g u \<in> closure S"
proof -
obtain u where "0 \<le> u" "u \<le> 1"
and gxin: "\<And>x. 0 \<le> x \<and> x < u \<Longrightarrow> g x \<in> interior S"
and gunot: "(g u \<notin> interior S)" and u0: "(u = 0 \<or> g u \<in> closure S)"
using subpath_to_frontier_explicit [OF assms] by blast
show ?thesis
apply (rule that [OF \<open>0 \<le> u\<close> \<open>u \<le> 1\<close>])
apply (simp add: gunot)
using \<open>0 \<le> u\<close> u0 by (force simp: subpath_def gxin)
qed
lemma subpath_to_frontier:
assumes g: "path g" and g0: "pathstart g \<in> closure S" and g1: "pathfinish g \<notin> S"
obtains u where "0 \<le> u" "u \<le> 1" "g u \<in> frontier S" "(path_image(subpath 0 u g) - {g u}) \<subseteq> interior S"
proof -
obtain u where "0 \<le> u" "u \<le> 1"
and notin: "g u \<notin> interior S"
and disj: "u = 0 \<or>
(\<forall>x. 0 \<le> x \<and> x < 1 \<longrightarrow> subpath 0 u g x \<in> interior S) \<and> g u \<in> closure S"
using subpath_to_frontier_strong [OF g g1] by blast
show ?thesis
apply (rule that [OF \<open>0 \<le> u\<close> \<open>u \<le> 1\<close>])
apply (metis DiffI disj frontier_def g0 notin pathstart_def)
using \<open>0 \<le> u\<close> g0 disj
apply (simp add: path_image_subpath_gen)
apply (auto simp: closed_segment_eq_real_ivl pathstart_def pathfinish_def subpath_def)
apply (rename_tac y)
apply (drule_tac x="y/u" in spec)
apply (auto split: if_split_asm)
done
qed
lemma exists_path_subpath_to_frontier:
fixes S :: "'a::real_normed_vector set"
assumes "path g" "pathstart g \<in> closure S" "pathfinish g \<notin> S"
obtains h where "path h" "pathstart h = pathstart g" "path_image h \<subseteq> path_image g"
"path_image h - {pathfinish h} \<subseteq> interior S"
"pathfinish h \<in> frontier S"
proof -
obtain u where u: "0 \<le> u" "u \<le> 1" "g u \<in> frontier S" "(path_image(subpath 0 u g) - {g u}) \<subseteq> interior S"
using subpath_to_frontier [OF assms] by blast
show ?thesis
apply (rule that [of "subpath 0 u g"])
using assms u
apply (simp_all add: path_image_subpath)
apply (simp add: pathstart_def)
apply (force simp: closed_segment_eq_real_ivl path_image_def)
done
qed
lemma exists_path_subpath_to_frontier_closed:
fixes S :: "'a::real_normed_vector set"
assumes S: "closed S" and g: "path g" and g0: "pathstart g \<in> S" and g1: "pathfinish g \<notin> S"
obtains h where "path h" "pathstart h = pathstart g" "path_image h \<subseteq> path_image g \<inter> S"
"pathfinish h \<in> frontier S"
proof -
obtain h where h: "path h" "pathstart h = pathstart g" "path_image h \<subseteq> path_image g"
"path_image h - {pathfinish h} \<subseteq> interior S"
"pathfinish h \<in> frontier S"
using exists_path_subpath_to_frontier [OF g _ g1] closure_closed [OF S] g0 by auto
show ?thesis
apply (rule that [OF \<open>path h\<close>])
using assms h
apply auto
apply (metis Diff_single_insert frontier_subset_eq insert_iff interior_subset subset_iff)
done
qed
subsection \<open>Shift Path to Start at Some Given Point\<close>
definition\<^marker>\<open>tag important\<close> shiftpath :: "real \<Rightarrow> (real \<Rightarrow> 'a::topological_space) \<Rightarrow> real \<Rightarrow> 'a"
where "shiftpath a f = (\<lambda>x. if (a + x) \<le> 1 then f (a + x) else f (a + x - 1))"
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 pathstart_shiftpath: "a \<le> 1 \<Longrightarrow> pathstart (shiftpath a g) = g a"
unfolding pathstart_def shiftpath_def by auto
lemma pathfinish_shiftpath:
assumes "0 \<le> a"
and "pathfinish g = pathstart g"
shows "pathfinish (shiftpath a g) = g a"
using assms
unfolding pathstart_def pathfinish_def shiftpath_def
by auto
lemma endpoints_shiftpath:
assumes "pathfinish g = pathstart g"
and "a \<in> {0 .. 1}"
shows "pathfinish (shiftpath a g) = g a"
and "pathstart (shiftpath a g) = g a"
using assms
by (auto intro!: pathfinish_shiftpath pathstart_shiftpath)
lemma closed_shiftpath:
assumes "pathfinish g = pathstart g"
and "a \<in> {0..1}"
shows "pathfinish (shiftpath a g) = pathstart (shiftpath a g)"
using endpoints_shiftpath[OF assms]
by auto
lemma path_shiftpath:
assumes "path g"
and "pathfinish g = pathstart g"
and "a \<in> {0..1}"
shows "path (shiftpath a g)"
proof -
have *: "{0 .. 1} = {0 .. 1-a} \<union> {1-a .. 1}"
using assms(3) by auto
have **: "\<And>x. x + a = 1 \<Longrightarrow> g (x + a - 1) = g (x + a)"
using assms(2)[unfolded pathfinish_def pathstart_def]
by auto
show ?thesis
unfolding path_def shiftpath_def *
proof (rule continuous_on_closed_Un)
have contg: "continuous_on {0..1} g"
using \<open>path g\<close> path_def by blast
show "continuous_on {0..1-a} (\<lambda>x. if a + x \<le> 1 then g (a + x) else g (a + x - 1))"
proof (rule continuous_on_eq)
show "continuous_on {0..1-a} (g \<circ> (+) a)"
by (intro continuous_intros continuous_on_subset [OF contg]) (use \<open>a \<in> {0..1}\<close> in auto)
qed auto
show "continuous_on {1-a..1} (\<lambda>x. if a + x \<le> 1 then g (a + x) else g (a + x - 1))"
proof (rule continuous_on_eq)
show "continuous_on {1-a..1} (g \<circ> (+) (a - 1))"
by (intro continuous_intros continuous_on_subset [OF contg]) (use \<open>a \<in> {0..1}\<close> in auto)
qed (auto simp: "**" add.commute add_diff_eq)
qed auto
qed
lemma shiftpath_shiftpath:
assumes "pathfinish g = pathstart g"
and "a \<in> {0..1}"
and "x \<in> {0..1}"
shows "shiftpath (1 - a) (shiftpath a g) x = g x"
using assms
unfolding pathfinish_def pathstart_def shiftpath_def
by auto
lemma path_image_shiftpath:
assumes a: "a \<in> {0..1}"
and "pathfinish g = pathstart g"
shows "path_image (shiftpath a g) = path_image g"
proof -
{ fix x
assume g: "g 1 = g 0" "x \<in> {0..1::real}" and gne: "\<And>y. y\<in>{0..1} \<inter> {x. \<not> a + x \<le> 1} \<Longrightarrow> g x \<noteq> g (a + y - 1)"
then have "\<exists>y\<in>{0..1} \<inter> {x. a + x \<le> 1}. g x = g (a + y)"
proof (cases "a \<le> x")
case False
then show ?thesis
apply (rule_tac x="1 + x - a" in bexI)
using g gne[of "1 + x - a"] a
apply (force simp: field_simps)+
done
next
case True
then show ?thesis
using g a by (rule_tac x="x - a" in bexI) (auto simp: field_simps)
qed
}
then show ?thesis
using assms
unfolding shiftpath_def path_image_def pathfinish_def pathstart_def
by (auto simp: image_iff)
qed
lemma simple_path_shiftpath:
assumes "simple_path g" "pathfinish g = pathstart g" and a: "0 \<le> a" "a \<le> 1"
shows "simple_path (shiftpath a g)"
unfolding simple_path_def
proof (intro conjI impI ballI)
show "path (shiftpath a g)"
by (simp add: assms path_shiftpath simple_path_imp_path)
have *: "\<And>x y. \<lbrakk>g x = g y; x \<in> {0..1}; y \<in> {0..1}\<rbrakk> \<Longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0"
using assms by (simp add: simple_path_def)
show "x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0"
if "x \<in> {0..1}" "y \<in> {0..1}" "shiftpath a g x = shiftpath a g y" for x y
using that a unfolding shiftpath_def
by (force split: if_split_asm dest!: *)
qed
subsection \<open>Straight-Line Paths\<close>
definition\<^marker>\<open>tag important\<close> linepath :: "'a::real_normed_vector \<Rightarrow> 'a \<Rightarrow> real \<Rightarrow> 'a"
where "linepath a b = (\<lambda>x. (1 - x) *\<^sub>R a + x *\<^sub>R b)"
lemma pathstart_linepath[simp]: "pathstart (linepath a b) = a"
unfolding pathstart_def linepath_def
by auto
lemma pathfinish_linepath[simp]: "pathfinish (linepath a b) = b"
unfolding pathfinish_def linepath_def
by auto
lemma linepath_inner: "linepath a b x \<bullet> v = linepath (a \<bullet> v) (b \<bullet> v) x"
by (simp add: linepath_def algebra_simps)
lemma Re_linepath': "Re (linepath a b x) = linepath (Re a) (Re b) x"
by (simp add: linepath_def)
lemma Im_linepath': "Im (linepath a b x) = linepath (Im a) (Im b) x"
by (simp add: linepath_def)
lemma linepath_0': "linepath a b 0 = a"
by (simp add: linepath_def)
lemma linepath_1': "linepath a b 1 = b"
by (simp add: linepath_def)
lemma continuous_linepath_at[intro]: "continuous (at x) (linepath a b)"
unfolding linepath_def
by (intro continuous_intros)
lemma continuous_on_linepath [intro,continuous_intros]: "continuous_on s (linepath a b)"
using continuous_linepath_at
by (auto intro!: continuous_at_imp_continuous_on)
lemma path_linepath[iff]: "path (linepath a b)"
unfolding path_def
by (rule continuous_on_linepath)
lemma path_image_linepath[simp]: "path_image (linepath a b) = closed_segment a b"
unfolding path_image_def segment linepath_def
by auto
lemma reversepath_linepath[simp]: "reversepath (linepath a b) = linepath b a"
unfolding reversepath_def linepath_def
by auto
lemma linepath_0 [simp]: "linepath 0 b x = x *\<^sub>R b"
by (simp add: linepath_def)
lemma linepath_cnj: "cnj (linepath a b x) = linepath (cnj a) (cnj b) x"
by (simp add: linepath_def)
lemma arc_linepath:
assumes "a \<noteq> b" shows [simp]: "arc (linepath a b)"
proof -
{
fix x y :: "real"
assume "x *\<^sub>R b + y *\<^sub>R a = x *\<^sub>R a + y *\<^sub>R b"
then have "(x - y) *\<^sub>R a = (x - y) *\<^sub>R b"
by (simp add: algebra_simps)
with assms have "x = y"
by simp
}
then show ?thesis
unfolding arc_def inj_on_def
by (fastforce simp: algebra_simps linepath_def)
qed
lemma simple_path_linepath[intro]: "a \<noteq> b \<Longrightarrow> simple_path (linepath a b)"
by (simp add: arc_imp_simple_path)
lemma linepath_trivial [simp]: "linepath a a x = a"
by (simp add: linepath_def real_vector.scale_left_diff_distrib)
lemma linepath_refl: "linepath a a = (\<lambda>x. a)"
by auto
lemma subpath_refl: "subpath a a g = linepath (g a) (g a)"
by (simp add: subpath_def linepath_def algebra_simps)
lemma linepath_of_real: "(linepath (of_real a) (of_real b) x) = of_real ((1 - x)*a + x*b)"
by (simp add: scaleR_conv_of_real linepath_def)
lemma of_real_linepath: "of_real (linepath a b x) = linepath (of_real a) (of_real b) x"
by (metis linepath_of_real mult.right_neutral of_real_def real_scaleR_def)
lemma inj_on_linepath:
assumes "a \<noteq> b" shows "inj_on (linepath a b) {0..1}"
proof (clarsimp simp: inj_on_def linepath_def)
fix x y
assume "(1 - x) *\<^sub>R a + x *\<^sub>R b = (1 - y) *\<^sub>R a + y *\<^sub>R b" "0 \<le> x" "x \<le> 1" "0 \<le> y" "y \<le> 1"
then have "x *\<^sub>R (a - b) = y *\<^sub>R (a - b)"
by (auto simp: algebra_simps)
then show "x=y"
using assms by auto
qed
lemma linepath_le_1:
fixes a::"'a::linordered_idom" shows "\<lbrakk>a \<le> 1; b \<le> 1; 0 \<le> u; u \<le> 1\<rbrakk> \<Longrightarrow> (1 - u) * a + u * b \<le> 1"
using mult_left_le [of a "1-u"] mult_left_le [of b u] by auto
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 bounded_linear_linepath:
assumes "bounded_linear f"
shows "f (linepath a b x) = linepath (f a) (f b) x"
proof -
interpret f: bounded_linear f by fact
show ?thesis by (simp add: linepath_def f.add f.scale)
qed
lemma bounded_linear_linepath':
assumes "bounded_linear f"
shows "f \<circ> linepath a b = linepath (f a) (f b)"
using bounded_linear_linepath[OF assms] by (simp add: fun_eq_iff)
lemma linepath_cnj': "cnj \<circ> linepath a b = linepath (cnj a) (cnj b)"
by (simp add: linepath_def fun_eq_iff)
lemma differentiable_linepath [intro]: "linepath a b differentiable at x within A"
by (auto simp: linepath_def)
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
subsection\<^marker>\<open>tag unimportant\<close>\<open>Segments via convex hulls\<close>
lemma segments_subset_convex_hull:
"closed_segment a b \<subseteq> (convex hull {a,b,c})"
"closed_segment a c \<subseteq> (convex hull {a,b,c})"
"closed_segment b c \<subseteq> (convex hull {a,b,c})"
"closed_segment b a \<subseteq> (convex hull {a,b,c})"
"closed_segment c a \<subseteq> (convex hull {a,b,c})"
"closed_segment c b \<subseteq> (convex hull {a,b,c})"
by (auto simp: segment_convex_hull linepath_of_real elim!: rev_subsetD [OF _ hull_mono])
lemma midpoints_in_convex_hull:
assumes "x \<in> convex hull s" "y \<in> convex hull s"
shows "midpoint x y \<in> convex hull s"
proof -
have "(1 - inverse(2)) *\<^sub>R x + inverse(2) *\<^sub>R y \<in> convex hull s"
by (rule convexD_alt) (use assms in auto)
then show ?thesis
by (simp add: midpoint_def algebra_simps)
qed
lemma not_in_interior_convex_hull_3:
fixes a :: "complex"
shows "a \<notin> interior(convex hull {a,b,c})"
"b \<notin> interior(convex hull {a,b,c})"
"c \<notin> interior(convex hull {a,b,c})"
by (auto simp: card_insert_le_m1 not_in_interior_convex_hull)
lemma midpoint_in_closed_segment [simp]: "midpoint a b \<in> closed_segment a b"
using midpoints_in_convex_hull segment_convex_hull by blast
lemma midpoint_in_open_segment [simp]: "midpoint a b \<in> open_segment a b \<longleftrightarrow> a \<noteq> b"
by (simp add: open_segment_def)
lemma continuous_IVT_local_extremum:
fixes f :: "'a::euclidean_space \<Rightarrow> real"
assumes contf: "continuous_on (closed_segment a b) f"
and "a \<noteq> b" "f a = f b"
obtains z where "z \<in> open_segment a b"
"(\<forall>w \<in> closed_segment a b. (f w) \<le> (f z)) \<or>
(\<forall>w \<in> closed_segment a b. (f z) \<le> (f w))"
proof -
obtain c where "c \<in> closed_segment a b" and c: "\<And>y. y \<in> closed_segment a b \<Longrightarrow> f y \<le> f c"
using continuous_attains_sup [of "closed_segment a b" f] contf by auto
obtain d where "d \<in> closed_segment a b" and d: "\<And>y. y \<in> closed_segment a b \<Longrightarrow> f d \<le> f y"
using continuous_attains_inf [of "closed_segment a b" f] contf by auto
show ?thesis
proof (cases "c \<in> open_segment a b \<or> d \<in> open_segment a b")
case True
then show ?thesis
using c d that by blast
next
case False
then have "(c = a \<or> c = b) \<and> (d = a \<or> d = b)"
by (simp add: \<open>c \<in> closed_segment a b\<close> \<open>d \<in> closed_segment a b\<close> open_segment_def)
with \<open>a \<noteq> b\<close> \<open>f a = f b\<close> c d show ?thesis
by (rule_tac z = "midpoint a b" in that) (fastforce+)
qed
qed
text\<open>An injective map into R is also an open map w.r.T. the universe, and conversely. \<close>
proposition injective_eq_1d_open_map_UNIV:
fixes f :: "real \<Rightarrow> real"
assumes contf: "continuous_on S f" and S: "is_interval S"
shows "inj_on f S \<longleftrightarrow> (\<forall>T. open T \<and> T \<subseteq> S \<longrightarrow> open(f ` T))"
(is "?lhs = ?rhs")
proof safe
fix T
assume injf: ?lhs and "open T" and "T \<subseteq> S"
have "\<exists>U. open U \<and> f x \<in> U \<and> U \<subseteq> f ` T" if "x \<in> T" for x
proof -
obtain \<delta> where "\<delta> > 0" and \<delta>: "cball x \<delta> \<subseteq> T"
using \<open>open T\<close> \<open>x \<in> T\<close> open_contains_cball_eq by blast
show ?thesis
proof (intro exI conjI)
have "closed_segment (x-\<delta>) (x+\<delta>) = {x-\<delta>..x+\<delta>}"
using \<open>0 < \<delta>\<close> by (auto simp: closed_segment_eq_real_ivl)
also have "\<dots> \<subseteq> S"
using \<delta> \<open>T \<subseteq> S\<close> by (auto simp: dist_norm subset_eq)
finally have "f ` (open_segment (x-\<delta>) (x+\<delta>)) = open_segment (f (x-\<delta>)) (f (x+\<delta>))"
using continuous_injective_image_open_segment_1
by (metis continuous_on_subset [OF contf] inj_on_subset [OF injf])
then show "open (f ` {x-\<delta><..<x+\<delta>})"
using \<open>0 < \<delta>\<close> by (simp add: open_segment_eq_real_ivl)
show "f x \<in> f ` {x - \<delta><..<x + \<delta>}"
by (auto simp: \<open>\<delta> > 0\<close>)
show "f ` {x - \<delta><..<x + \<delta>} \<subseteq> f ` T"
using \<delta> by (auto simp: dist_norm subset_iff)
qed
qed
with open_subopen show "open (f ` T)"
by blast
next
assume R: ?rhs
have False if xy: "x \<in> S" "y \<in> S" and "f x = f y" "x \<noteq> y" for x y
proof -
have "open (f ` open_segment x y)"
using R
by (metis S convex_contains_open_segment is_interval_convex open_greaterThanLessThan open_segment_eq_real_ivl xy)
moreover
have "continuous_on (closed_segment x y) f"
by (meson S closed_segment_subset contf continuous_on_subset is_interval_convex that)
then obtain \<xi> where "\<xi> \<in> open_segment x y"
and \<xi>: "(\<forall>w \<in> closed_segment x y. (f w) \<le> (f \<xi>)) \<or>
(\<forall>w \<in> closed_segment x y. (f \<xi>) \<le> (f w))"
using continuous_IVT_local_extremum [of x y f] \<open>f x = f y\<close> \<open>x \<noteq> y\<close> by blast
ultimately obtain e where "e>0" and e: "\<And>u. dist u (f \<xi>) < e \<Longrightarrow> u \<in> f ` open_segment x y"
using open_dist by (metis image_eqI)
have fin: "f \<xi> + (e/2) \<in> f ` open_segment x y" "f \<xi> - (e/2) \<in> f ` open_segment x y"
using e [of "f \<xi> + (e/2)"] e [of "f \<xi> - (e/2)"] \<open>e > 0\<close> by (auto simp: dist_norm)
show ?thesis
using \<xi> \<open>0 < e\<close> fin open_closed_segment by fastforce
qed
then show ?lhs
by (force simp: inj_on_def)
qed
subsection\<^marker>\<open>tag unimportant\<close> \<open>Bounding a point away from a path\<close>
lemma not_on_path_ball:
fixes g :: "real \<Rightarrow> 'a::heine_borel"
assumes "path g"
and z: "z \<notin> path_image g"
shows "\<exists>e > 0. ball z e \<inter> path_image g = {}"
proof -
have "closed (path_image g)"
by (simp add: \<open>path g\<close> closed_path_image)
then obtain a where "a \<in> path_image g" "\<forall>y \<in> path_image g. dist z a \<le> dist z y"
by (auto intro: distance_attains_inf[OF _ path_image_nonempty, of g z])
then show ?thesis
by (rule_tac x="dist z a" in exI) (use dist_commute z in auto)
qed
lemma not_on_path_cball:
fixes g :: "real \<Rightarrow> 'a::heine_borel"
assumes "path g"
and "z \<notin> path_image g"
shows "\<exists>e>0. cball z e \<inter> (path_image g) = {}"
proof -
obtain e where "ball z e \<inter> path_image g = {}" "e > 0"
using not_on_path_ball[OF assms] by auto
moreover have "cball z (e/2) \<subseteq> ball z e"
using \<open>e > 0\<close> by auto
ultimately show ?thesis
by (rule_tac x="e/2" in exI) auto
qed
subsection \<open>Path component\<close>
text \<open>Original formalization by Tom Hales\<close>
definition\<^marker>\<open>tag important\<close> "path_component s x y \<longleftrightarrow>
(\<exists>g. path g \<and> path_image g \<subseteq> s \<and> pathstart g = x \<and> pathfinish g = y)"
abbreviation\<^marker>\<open>tag important\<close>
"path_component_set s x \<equiv> Collect (path_component s x)"
lemmas path_defs = path_def pathstart_def pathfinish_def path_image_def path_component_def
lemma path_component_mem:
assumes "path_component s x y"
shows "x \<in> s" and "y \<in> s"
using assms
unfolding path_defs
by auto
lemma path_component_refl:
assumes "x \<in> s"
shows "path_component s x x"
unfolding path_defs
apply (rule_tac x="\<lambda>u. x" in exI)
using assms
apply (auto intro!: continuous_intros)
done
lemma path_component_refl_eq: "path_component s x x \<longleftrightarrow> x \<in> s"
by (auto intro!: path_component_mem path_component_refl)
lemma path_component_sym: "path_component s x y \<Longrightarrow> path_component s y x"
unfolding path_component_def
apply (erule exE)
apply (rule_tac x="reversepath g" in exI, auto)
done
lemma path_component_trans:
assumes "path_component s x y" and "path_component s y z"
shows "path_component s x z"
using assms
unfolding path_component_def
apply (elim exE)
apply (rule_tac x="g +++ ga" in exI)
apply (auto simp: path_image_join)
done
lemma path_component_of_subset: "s \<subseteq> t \<Longrightarrow> path_component s x y \<Longrightarrow> path_component t x y"
unfolding path_component_def by auto
lemma path_component_linepath:
fixes s :: "'a::real_normed_vector set"
shows "closed_segment a b \<subseteq> s \<Longrightarrow> path_component s a b"
unfolding path_component_def
by (rule_tac x="linepath a b" in exI, auto)
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Path components as sets\<close>
lemma path_component_set:
"path_component_set s x =
{y. (\<exists>g. path g \<and> path_image g \<subseteq> s \<and> pathstart g = x \<and> pathfinish g = y)}"
by (auto simp: path_component_def)
lemma path_component_subset: "path_component_set s x \<subseteq> s"
by (auto simp: path_component_mem(2))
lemma path_component_eq_empty: "path_component_set s x = {} \<longleftrightarrow> x \<notin> s"
using path_component_mem path_component_refl_eq
by fastforce
lemma path_component_mono:
"s \<subseteq> t \<Longrightarrow> (path_component_set s x) \<subseteq> (path_component_set t x)"
by (simp add: Collect_mono path_component_of_subset)
lemma path_component_eq:
"y \<in> path_component_set s x \<Longrightarrow> path_component_set s y = path_component_set s x"
by (metis (no_types, lifting) Collect_cong mem_Collect_eq path_component_sym path_component_trans)
subsection \<open>Path connectedness of a space\<close>
definition\<^marker>\<open>tag important\<close> "path_connected s \<longleftrightarrow>
(\<forall>x\<in>s. \<forall>y\<in>s. \<exists>g. path g \<and> path_image g \<subseteq> s \<and> pathstart g = x \<and> pathfinish g = y)"
lemma path_connectedin_iff_path_connected_real [simp]:
"path_connectedin euclideanreal S \<longleftrightarrow> path_connected S"
by (simp add: path_connectedin path_connected_def path_defs)
lemma path_connected_component: "path_connected s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. path_component s x y)"
unfolding path_connected_def path_component_def by auto
lemma path_connected_component_set: "path_connected s \<longleftrightarrow> (\<forall>x\<in>s. path_component_set s x = s)"
unfolding path_connected_component path_component_subset
using path_component_mem by blast
lemma path_component_maximal:
"\<lbrakk>x \<in> t; path_connected t; t \<subseteq> s\<rbrakk> \<Longrightarrow> t \<subseteq> (path_component_set s x)"
by (metis path_component_mono path_connected_component_set)
lemma convex_imp_path_connected:
fixes s :: "'a::real_normed_vector set"
assumes "convex s"
shows "path_connected s"
unfolding path_connected_def
using assms convex_contains_segment by fastforce
lemma path_connected_UNIV [iff]: "path_connected (UNIV :: 'a::real_normed_vector set)"
by (simp add: convex_imp_path_connected)
lemma path_component_UNIV: "path_component_set UNIV x = (UNIV :: 'a::real_normed_vector set)"
using path_connected_component_set by auto
lemma path_connected_imp_connected:
assumes "path_connected S"
shows "connected S"
proof (rule connectedI)
fix e1 e2
assume as: "open e1" "open e2" "S \<subseteq> e1 \<union> e2" "e1 \<inter> e2 \<inter> S = {}" "e1 \<inter> S \<noteq> {}" "e2 \<inter> S \<noteq> {}"
then obtain x1 x2 where obt:"x1 \<in> e1 \<inter> S" "x2 \<in> e2 \<inter> S"
by auto
then obtain g where g: "path g" "path_image g \<subseteq> S" "pathstart g = x1" "pathfinish g = x2"
using assms[unfolded path_connected_def,rule_format,of x1 x2] by auto
have *: "connected {0..1::real}"
by (auto intro!: convex_connected)
have "{0..1} \<subseteq> {x \<in> {0..1}. g x \<in> e1} \<union> {x \<in> {0..1}. g x \<in> e2}"
using as(3) g(2)[unfolded path_defs] by blast
moreover have "{x \<in> {0..1}. g x \<in> e1} \<inter> {x \<in> {0..1}. g x \<in> e2} = {}"
using as(4) g(2)[unfolded path_defs]
unfolding subset_eq
by auto
moreover have "{x \<in> {0..1}. g x \<in> e1} \<noteq> {} \<and> {x \<in> {0..1}. g x \<in> e2} \<noteq> {}"
using g(3,4)[unfolded path_defs]
using obt
by (simp add: ex_in_conv [symmetric], metis zero_le_one order_refl)
ultimately show False
using *[unfolded connected_local not_ex, rule_format,
of "{0..1} \<inter> g -` e1" "{0..1} \<inter> g -` e2"]
using continuous_openin_preimage_gen[OF g(1)[unfolded path_def] as(1)]
using continuous_openin_preimage_gen[OF g(1)[unfolded path_def] as(2)]
by auto
qed
lemma open_path_component:
fixes S :: "'a::real_normed_vector set"
assumes "open S"
shows "open (path_component_set S x)"
unfolding open_contains_ball
proof
fix y
assume as: "y \<in> path_component_set S x"
then have "y \<in> S"
by (simp add: path_component_mem(2))
then obtain e where e: "e > 0" "ball y e \<subseteq> S"
using assms[unfolded open_contains_ball]
by auto
have "\<And>u. dist y u < e \<Longrightarrow> path_component S x u"
by (metis (full_types) as centre_in_ball convex_ball convex_imp_path_connected e mem_Collect_eq mem_ball path_component_eq path_component_of_subset path_connected_component)
then show "\<exists>e > 0. ball y e \<subseteq> path_component_set S x"
using \<open>e>0\<close> by auto
qed
lemma open_non_path_component:
fixes S :: "'a::real_normed_vector set"
assumes "open S"
shows "open (S - path_component_set S x)"
unfolding open_contains_ball
proof
fix y
assume y: "y \<in> S - path_component_set S x"
then obtain e where e: "e > 0" "ball y e \<subseteq> S"
using assms openE by auto
show "\<exists>e>0. ball y e \<subseteq> S - path_component_set S x"
proof (intro exI conjI subsetI DiffI notI)
show "\<And>x. x \<in> ball y e \<Longrightarrow> x \<in> S"
using e by blast
show False if "z \<in> ball y e" "z \<in> path_component_set S x" for z
proof -
have "y \<in> path_component_set S z"
by (meson assms convex_ball convex_imp_path_connected e open_contains_ball_eq open_path_component path_component_maximal that(1))
then have "y \<in> path_component_set S x"
using path_component_eq that(2) by blast
then show False
using y by blast
qed
qed (use e in auto)
qed
lemma connected_open_path_connected:
fixes S :: "'a::real_normed_vector set"
assumes "open S"
and "connected S"
shows "path_connected S"
unfolding path_connected_component_set
proof (rule, rule, rule path_component_subset, rule)
fix x y
assume "x \<in> S" and "y \<in> S"
show "y \<in> path_component_set S x"
proof (rule ccontr)
assume "\<not> ?thesis"
moreover have "path_component_set S x \<inter> S \<noteq> {}"
using \<open>x \<in> S\<close> path_component_eq_empty path_component_subset[of S x]
by auto
ultimately
show False
using \<open>y \<in> S\<close> open_non_path_component[OF assms(1)] open_path_component[OF assms(1)]
using assms(2)[unfolded connected_def not_ex, rule_format,
of "path_component_set S x" "S - path_component_set S x"]
by auto
qed
qed
lemma path_connected_continuous_image:
assumes "continuous_on S f"
and "path_connected S"
shows "path_connected (f ` S)"
unfolding path_connected_def
proof (rule, rule)
fix x' y'
assume "x' \<in> f ` S" "y' \<in> f ` S"
then obtain x y where x: "x \<in> S" and y: "y \<in> S" and x': "x' = f x" and y': "y' = f y"
by auto
from x y obtain g where "path g \<and> path_image g \<subseteq> S \<and> pathstart g = x \<and> pathfinish g = y"
using assms(2)[unfolded path_connected_def] by fast
then show "\<exists>g. path g \<and> path_image g \<subseteq> f ` S \<and> pathstart g = x' \<and> pathfinish g = y'"
unfolding x' y'
apply (rule_tac x="f \<circ> g" in exI)
unfolding path_defs
apply (intro conjI continuous_on_compose continuous_on_subset[OF assms(1)])
apply auto
done
qed
lemma path_connected_translationI:
fixes a :: "'a :: topological_group_add"
assumes "path_connected S" shows "path_connected ((\<lambda>x. a + x) ` S)"
by (intro path_connected_continuous_image assms continuous_intros)
lemma path_connected_translation:
fixes a :: "'a :: topological_group_add"
shows "path_connected ((\<lambda>x. a + x) ` S) = path_connected S"
proof -
have "\<forall>x y. (+) (x::'a) ` (+) (0 - x) ` y = y"
by (simp add: image_image)
then show ?thesis
by (metis (no_types) path_connected_translationI)
qed
lemma path_connected_segment [simp]:
fixes a :: "'a::real_normed_vector"
shows "path_connected (closed_segment a b)"
by (simp add: convex_imp_path_connected)
lemma path_connected_open_segment [simp]:
fixes a :: "'a::real_normed_vector"
shows "path_connected (open_segment a b)"
by (simp add: convex_imp_path_connected)
lemma homeomorphic_path_connectedness:
"S homeomorphic T \<Longrightarrow> path_connected S \<longleftrightarrow> path_connected T"
unfolding homeomorphic_def homeomorphism_def by (metis path_connected_continuous_image)
lemma path_connected_empty [simp]: "path_connected {}"
unfolding path_connected_def by auto
lemma path_connected_singleton [simp]: "path_connected {a}"
unfolding path_connected_def pathstart_def pathfinish_def path_image_def
apply clarify
apply (rule_tac x="\<lambda>x. a" in exI)
apply (simp add: image_constant_conv)
apply (simp add: path_def)
done
lemma path_connected_Un:
assumes "path_connected S"
and "path_connected T"
and "S \<inter> T \<noteq> {}"
shows "path_connected (S \<union> T)"
unfolding path_connected_component
proof (intro ballI)
fix x y
assume x: "x \<in> S \<union> T" and y: "y \<in> S \<union> T"
from assms obtain z where z: "z \<in> S" "z \<in> T"
by auto
show "path_component (S \<union> T) x y"
using x y
proof safe
assume "x \<in> S" "y \<in> S"
then show "path_component (S \<union> T) x y"
by (meson Un_upper1 \<open>path_connected S\<close> path_component_of_subset path_connected_component)
next
assume "x \<in> S" "y \<in> T"
then show "path_component (S \<union> T) x y"
by (metis z assms(1-2) le_sup_iff order_refl path_component_of_subset path_component_trans path_connected_component)
next
assume "x \<in> T" "y \<in> S"
then show "path_component (S \<union> T) x y"
by (metis z assms(1-2) le_sup_iff order_refl path_component_of_subset path_component_trans path_connected_component)
next
assume "x \<in> T" "y \<in> T"
then show "path_component (S \<union> T) x y"
by (metis Un_upper1 assms(2) path_component_of_subset path_connected_component sup_commute)
qed
qed
lemma path_connected_UNION:
assumes "\<And>i. i \<in> A \<Longrightarrow> path_connected (S i)"
and "\<And>i. i \<in> A \<Longrightarrow> z \<in> S i"
shows "path_connected (\<Union>i\<in>A. S i)"
unfolding path_connected_component
proof clarify
fix x i y j
assume *: "i \<in> A" "x \<in> S i" "j \<in> A" "y \<in> S j"
then have "path_component (S i) x z" and "path_component (S j) z y"
using assms by (simp_all add: path_connected_component)
then have "path_component (\<Union>i\<in>A. S i) x z" and "path_component (\<Union>i\<in>A. S i) z y"
using *(1,3) by (auto elim!: path_component_of_subset [rotated])
then show "path_component (\<Union>i\<in>A. S i) x y"
by (rule path_component_trans)
qed
lemma path_component_path_image_pathstart:
assumes p: "path p" and x: "x \<in> path_image p"
shows "path_component (path_image p) (pathstart p) x"
proof -
obtain y where x: "x = p y" and y: "0 \<le> y" "y \<le> 1"
using x by (auto simp: path_image_def)
show ?thesis
unfolding path_component_def
proof (intro exI conjI)
have "continuous_on {0..1} (p \<circ> ((*) y))"
apply (rule continuous_intros)+
using p [unfolded path_def] y
apply (auto simp: mult_le_one intro: continuous_on_subset [of _ p])
done
then show "path (\<lambda>u. p (y * u))"
by (simp add: path_def)
show "path_image (\<lambda>u. p (y * u)) \<subseteq> path_image p"
using y mult_le_one by (fastforce simp: path_image_def image_iff)
qed (auto simp: pathstart_def pathfinish_def x)
qed
lemma path_connected_path_image: "path p \<Longrightarrow> path_connected(path_image p)"
unfolding path_connected_component
by (meson path_component_path_image_pathstart path_component_sym path_component_trans)
lemma path_connected_path_component [simp]:
"path_connected (path_component_set s x)"
proof -
{ fix y z
assume pa: "path_component s x y" "path_component s x z"
then have pae: "path_component_set s x = path_component_set s y"
using path_component_eq by auto
have yz: "path_component s y z"
using pa path_component_sym path_component_trans by blast
then have "\<exists>g. path g \<and> path_image g \<subseteq> path_component_set s x \<and> pathstart g = y \<and> pathfinish g = z"
apply (simp add: path_component_def, clarify)
apply (rule_tac x=g in exI)
by (simp add: pae path_component_maximal path_connected_path_image pathstart_in_path_image)
}
then show ?thesis
by (simp add: path_connected_def)
qed
lemma path_component: "path_component S x y \<longleftrightarrow> (\<exists>t. path_connected t \<and> t \<subseteq> S \<and> x \<in> t \<and> y \<in> t)"
apply (intro iffI)
apply (metis path_connected_path_image path_defs(5) pathfinish_in_path_image pathstart_in_path_image)
using path_component_of_subset path_connected_component by blast
lemma path_component_path_component [simp]:
"path_component_set (path_component_set S x) x = path_component_set S x"
proof (cases "x \<in> S")
case True show ?thesis
apply (rule subset_antisym)
apply (simp add: path_component_subset)
by (simp add: True path_component_maximal path_component_refl)
next
case False then show ?thesis
by (metis False empty_iff path_component_eq_empty)
qed
lemma path_component_subset_connected_component:
"(path_component_set S x) \<subseteq> (connected_component_set S x)"
proof (cases "x \<in> S")
case True show ?thesis
apply (rule connected_component_maximal)
apply (auto simp: True path_component_subset path_component_refl path_connected_imp_connected)
done
next
case False then show ?thesis
using path_component_eq_empty by auto
qed
subsection\<^marker>\<open>tag unimportant\<close>\<open>Lemmas about path-connectedness\<close>
lemma path_connected_linear_image:
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
assumes "path_connected S" "bounded_linear f"
shows "path_connected(f ` S)"
by (auto simp: linear_continuous_on assms path_connected_continuous_image)
lemma is_interval_path_connected: "is_interval S \<Longrightarrow> path_connected S"
by (simp add: convex_imp_path_connected is_interval_convex)
lemma path_connected_Ioi[simp]: "path_connected {a<..}" for a :: real
by (simp add: convex_imp_path_connected)
lemma path_connected_Ici[simp]: "path_connected {a..}" for a :: real
by (simp add: convex_imp_path_connected)
lemma path_connected_Iio[simp]: "path_connected {..<a}" for a :: real
by (simp add: convex_imp_path_connected)
lemma path_connected_Iic[simp]: "path_connected {..a}" for a :: real
by (simp add: convex_imp_path_connected)
lemma path_connected_Ioo[simp]: "path_connected {a<..<b}" for a b :: real
by (simp add: convex_imp_path_connected)
lemma path_connected_Ioc[simp]: "path_connected {a<..b}" for a b :: real
by (simp add: convex_imp_path_connected)
lemma path_connected_Ico[simp]: "path_connected {a..<b}" for a b :: real
by (simp add: convex_imp_path_connected)
lemma path_connectedin_path_image:
assumes "pathin X g" shows "path_connectedin X (g ` ({0..1}))"
unfolding pathin_def
proof (rule path_connectedin_continuous_map_image)
show "continuous_map (subtopology euclideanreal {0..1}) X g"
using assms pathin_def by blast
qed (auto simp: is_interval_1 is_interval_path_connected)
lemma path_connected_space_subconnected:
"path_connected_space X \<longleftrightarrow>
(\<forall>x \<in> topspace X. \<forall>y \<in> topspace X. \<exists>S. path_connectedin X S \<and> x \<in> S \<and> y \<in> S)"
unfolding path_connected_space_def Ball_def
apply (intro all_cong1 imp_cong refl, safe)
using path_connectedin_path_image apply fastforce
by (meson path_connectedin)
lemma connectedin_path_image: "pathin X g \<Longrightarrow> connectedin X (g ` ({0..1}))"
by (simp add: path_connectedin_imp_connectedin path_connectedin_path_image)
lemma compactin_path_image: "pathin X g \<Longrightarrow> compactin X (g ` ({0..1}))"
unfolding pathin_def
by (rule image_compactin [of "top_of_set {0..1}"]) auto
lemma linear_homeomorphism_image:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "linear f" "inj f"
obtains g where "homeomorphism (f ` S) S g f"
using linear_injective_left_inverse [OF assms]
apply clarify
apply (rule_tac g=g in that)
using assms
apply (auto simp: homeomorphism_def eq_id_iff [symmetric] image_comp comp_def linear_conv_bounded_linear linear_continuous_on)
done
lemma linear_homeomorphic_image:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "linear f" "inj f"
shows "S homeomorphic f ` S"
by (meson homeomorphic_def homeomorphic_sym linear_homeomorphism_image [OF assms])
lemma path_connected_Times:
assumes "path_connected s" "path_connected t"
shows "path_connected (s \<times> t)"
proof (simp add: path_connected_def Sigma_def, clarify)
fix x1 y1 x2 y2
assume "x1 \<in> s" "y1 \<in> t" "x2 \<in> s" "y2 \<in> t"
obtain g where "path g" and g: "path_image g \<subseteq> s" and gs: "pathstart g = x1" and gf: "pathfinish g = x2"
using \<open>x1 \<in> s\<close> \<open>x2 \<in> s\<close> assms by (force simp: path_connected_def)
obtain h where "path h" and h: "path_image h \<subseteq> t" and hs: "pathstart h = y1" and hf: "pathfinish h = y2"
using \<open>y1 \<in> t\<close> \<open>y2 \<in> t\<close> assms by (force simp: path_connected_def)
have "path (\<lambda>z. (x1, h z))"
using \<open>path h\<close>
apply (simp add: path_def)
apply (rule continuous_on_compose2 [where f = h])
apply (rule continuous_intros | force)+
done
moreover have "path (\<lambda>z. (g z, y2))"
using \<open>path g\<close>
apply (simp add: path_def)
apply (rule continuous_on_compose2 [where f = g])
apply (rule continuous_intros | force)+
done
ultimately have 1: "path ((\<lambda>z. (x1, h z)) +++ (\<lambda>z. (g z, y2)))"
by (metis hf gs path_join_imp pathstart_def pathfinish_def)
have "path_image ((\<lambda>z. (x1, h z)) +++ (\<lambda>z. (g z, y2))) \<subseteq> path_image (\<lambda>z. (x1, h z)) \<union> path_image (\<lambda>z. (g z, y2))"
by (rule Path_Connected.path_image_join_subset)
also have "\<dots> \<subseteq> (\<Union>x\<in>s. \<Union>x1\<in>t. {(x, x1)})"
using g h \<open>x1 \<in> s\<close> \<open>y2 \<in> t\<close> by (force simp: path_image_def)
finally have 2: "path_image ((\<lambda>z. (x1, h z)) +++ (\<lambda>z. (g z, y2))) \<subseteq> (\<Union>x\<in>s. \<Union>x1\<in>t. {(x, x1)})" .
show "\<exists>g. path g \<and> path_image g \<subseteq> (\<Union>x\<in>s. \<Union>x1\<in>t. {(x, x1)}) \<and>
pathstart g = (x1, y1) \<and> pathfinish g = (x2, y2)"
apply (intro exI conjI)
apply (rule 1)
apply (rule 2)
apply (metis hs pathstart_def pathstart_join)
by (metis gf pathfinish_def pathfinish_join)
qed
lemma is_interval_path_connected_1:
fixes s :: "real set"
shows "is_interval s \<longleftrightarrow> path_connected s"
using is_interval_connected_1 is_interval_path_connected path_connected_imp_connected by blast
subsection\<^marker>\<open>tag unimportant\<close>\<open>Path components\<close>
lemma Union_path_component [simp]:
"Union {path_component_set S x |x. x \<in> S} = S"
apply (rule subset_antisym)
using path_component_subset apply force
using path_component_refl by auto
lemma path_component_disjoint:
"disjnt (path_component_set S a) (path_component_set S b) \<longleftrightarrow>
(a \<notin> path_component_set S b)"
apply (auto simp: disjnt_def)
using path_component_eq apply fastforce
using path_component_sym path_component_trans by blast
lemma path_component_eq_eq:
"path_component S x = path_component S y \<longleftrightarrow>
(x \<notin> S) \<and> (y \<notin> S) \<or> x \<in> S \<and> y \<in> S \<and> path_component S x y"
apply (rule iffI, metis (no_types) path_component_mem(1) path_component_refl)
apply (erule disjE, metis Collect_empty_eq_bot path_component_eq_empty)
apply (rule ext)
apply (metis path_component_trans path_component_sym)
done
lemma path_component_unique:
assumes "x \<in> c" "c \<subseteq> S" "path_connected c"
"\<And>c'. \<lbrakk>x \<in> c'; c' \<subseteq> S; path_connected c'\<rbrakk> \<Longrightarrow> c' \<subseteq> c"
shows "path_component_set S x = c"
apply (rule subset_antisym)
using assms
apply (metis mem_Collect_eq subsetCE path_component_eq_eq path_component_subset path_connected_path_component)
by (simp add: assms path_component_maximal)
lemma path_component_intermediate_subset:
"path_component_set u a \<subseteq> t \<and> t \<subseteq> u
\<Longrightarrow> path_component_set t a = path_component_set u a"
by (metis (no_types) path_component_mono path_component_path_component subset_antisym)
lemma complement_path_component_Union:
fixes x :: "'a :: topological_space"
shows "S - path_component_set S x =
\<Union>({path_component_set S y| y. y \<in> S} - {path_component_set S x})"
proof -
have *: "(\<And>x. x \<in> S - {a} \<Longrightarrow> disjnt a x) \<Longrightarrow> \<Union>S - a = \<Union>(S - {a})"
for a::"'a set" and S
by (auto simp: disjnt_def)
have "\<And>y. y \<in> {path_component_set S x |x. x \<in> S} - {path_component_set S x}
\<Longrightarrow> disjnt (path_component_set S x) y"
using path_component_disjoint path_component_eq by fastforce
then have "\<Union>{path_component_set S x |x. x \<in> S} - path_component_set S x =
\<Union>({path_component_set S y |y. y \<in> S} - {path_component_set S x})"
by (meson *)
then show ?thesis by simp
qed
subsection\<open>Path components\<close>
definition path_component_of
where "path_component_of X x y \<equiv> \<exists>g. pathin X g \<and> g 0 = x \<and> g 1 = y"
abbreviation path_component_of_set
where "path_component_of_set X x \<equiv> Collect (path_component_of X x)"
definition path_components_of :: "'a topology \<Rightarrow> 'a set set"
where "path_components_of X \<equiv> path_component_of_set X ` topspace X"
lemma pathin_canon_iff: "pathin (top_of_set T) g \<longleftrightarrow> path g \<and> g ` {0..1} \<subseteq> T"
by (simp add: path_def pathin_def)
lemma path_component_of_canon_iff [simp]:
"path_component_of (top_of_set T) a b \<longleftrightarrow> path_component T a b"
by (simp add: path_component_of_def pathin_canon_iff path_defs)
lemma path_component_in_topspace:
"path_component_of X x y \<Longrightarrow> x \<in> topspace X \<and> y \<in> topspace X"
by (auto simp: path_component_of_def pathin_def continuous_map_def)
lemma path_component_of_refl:
"path_component_of X x x \<longleftrightarrow> x \<in> topspace X"
apply (auto simp: path_component_in_topspace)
apply (force simp: path_component_of_def pathin_const)
done
lemma path_component_of_sym:
assumes "path_component_of X x y"
shows "path_component_of X y x"
using assms
apply (clarsimp simp: path_component_of_def pathin_def)
apply (rule_tac x="g \<circ> (\<lambda>t. 1 - t)" in exI)
apply (auto intro!: continuous_map_compose)
apply (force simp: continuous_map_in_subtopology continuous_on_op_minus)
done
lemma path_component_of_sym_iff:
"path_component_of X x y \<longleftrightarrow> path_component_of X y x"
by (metis path_component_of_sym)
lemma continuous_map_cases_le:
assumes contp: "continuous_map X euclideanreal p"
and contq: "continuous_map X euclideanreal q"
and contf: "continuous_map (subtopology X {x. x \<in> topspace X \<and> p x \<le> q x}) Y f"
and contg: "continuous_map (subtopology X {x. x \<in> topspace X \<and> q x \<le> p x}) Y g"
and fg: "\<And>x. \<lbrakk>x \<in> topspace X; p x = q x\<rbrakk> \<Longrightarrow> f x = g x"
shows "continuous_map X Y (\<lambda>x. if p x \<le> q x then f x else g x)"
proof -
have "continuous_map X Y (\<lambda>x. if q x - p x \<in> {0..} then f x else g x)"
proof (rule continuous_map_cases_function)
show "continuous_map X euclideanreal (\<lambda>x. q x - p x)"
by (intro contp contq continuous_intros)
show "continuous_map (subtopology X {x \<in> topspace X. q x - p x \<in> euclideanreal closure_of {0..}}) Y f"
by (simp add: contf)
show "continuous_map (subtopology X {x \<in> topspace X. q x - p x \<in> euclideanreal closure_of (topspace euclideanreal - {0..})}) Y g"
by (simp add: contg flip: Compl_eq_Diff_UNIV)
qed (auto simp: fg)
then show ?thesis
by simp
qed
lemma continuous_map_cases_lt:
assumes contp: "continuous_map X euclideanreal p"
and contq: "continuous_map X euclideanreal q"
and contf: "continuous_map (subtopology X {x. x \<in> topspace X \<and> p x \<le> q x}) Y f"
and contg: "continuous_map (subtopology X {x. x \<in> topspace X \<and> q x \<le> p x}) Y g"
and fg: "\<And>x. \<lbrakk>x \<in> topspace X; p x = q x\<rbrakk> \<Longrightarrow> f x = g x"
shows "continuous_map X Y (\<lambda>x. if p x < q x then f x else g x)"
proof -
have "continuous_map X Y (\<lambda>x. if q x - p x \<in> {0<..} then f x else g x)"
proof (rule continuous_map_cases_function)
show "continuous_map X euclideanreal (\<lambda>x. q x - p x)"
by (intro contp contq continuous_intros)
show "continuous_map (subtopology X {x \<in> topspace X. q x - p x \<in> euclideanreal closure_of {0<..}}) Y f"
by (simp add: contf)
show "continuous_map (subtopology X {x \<in> topspace X. q x - p x \<in> euclideanreal closure_of (topspace euclideanreal - {0<..})}) Y g"
by (simp add: contg flip: Compl_eq_Diff_UNIV)
qed (auto simp: fg)
then show ?thesis
by simp
qed
lemma path_component_of_trans:
assumes "path_component_of X x y" and "path_component_of X y z"
shows "path_component_of X x z"
unfolding path_component_of_def pathin_def
proof -
let ?T01 = "top_of_set {0..1::real}"
obtain g1 g2 where g1: "continuous_map ?T01 X g1" "x = g1 0" "y = g1 1"
and g2: "continuous_map ?T01 X g2" "g2 0 = g1 1" "z = g2 1"
using assms unfolding path_component_of_def pathin_def by blast
let ?g = "\<lambda>x. if x \<le> 1/2 then (g1 \<circ> (\<lambda>t. 2 * t)) x else (g2 \<circ> (\<lambda>t. 2 * t -1)) x"
show "\<exists>g. continuous_map ?T01 X g \<and> g 0 = x \<and> g 1 = z"
proof (intro exI conjI)
show "continuous_map (subtopology euclideanreal {0..1}) X ?g"
proof (intro continuous_map_cases_le continuous_map_compose, force, force)
show "continuous_map (subtopology ?T01 {x \<in> topspace ?T01. x \<le> 1/2}) ?T01 ((*) 2)"
by (auto simp: continuous_map_in_subtopology continuous_map_from_subtopology)
have "continuous_map
(subtopology (top_of_set {0..1}) {x. 0 \<le> x \<and> x \<le> 1 \<and> 1 \<le> x * 2})
euclideanreal (\<lambda>t. 2 * t - 1)"
by (intro continuous_intros) (force intro: continuous_map_from_subtopology)
then show "continuous_map (subtopology ?T01 {x \<in> topspace ?T01. 1/2 \<le> x}) ?T01 (\<lambda>t. 2 * t - 1)"
by (force simp: continuous_map_in_subtopology)
show "(g1 \<circ> (*) 2) x = (g2 \<circ> (\<lambda>t. 2 * t - 1)) x" if "x \<in> topspace ?T01" "x = 1/2" for x
using that by (simp add: g2(2) mult.commute continuous_map_from_subtopology)
qed (auto simp: g1 g2)
qed (auto simp: g1 g2)
qed
lemma path_component_of_mono:
"\<lbrakk>path_component_of (subtopology X S) x y; S \<subseteq> T\<rbrakk> \<Longrightarrow> path_component_of (subtopology X T) x y"
unfolding path_component_of_def
by (metis subsetD pathin_subtopology)
lemma path_component_of:
"path_component_of X x y \<longleftrightarrow> (\<exists>T. path_connectedin X T \<and> x \<in> T \<and> y \<in> T)"
apply (auto simp: path_component_of_def)
using path_connectedin_path_image apply fastforce
apply (metis path_connectedin)
done
lemma path_component_of_set:
"path_component_of X x y \<longleftrightarrow> (\<exists>g. pathin X g \<and> g 0 = x \<and> g 1 = y)"
by (auto simp: path_component_of_def)
lemma path_component_of_subset_topspace:
"Collect(path_component_of X x) \<subseteq> topspace X"
using path_component_in_topspace by fastforce
lemma path_component_of_eq_empty:
"Collect(path_component_of X x) = {} \<longleftrightarrow> (x \<notin> topspace X)"
using path_component_in_topspace path_component_of_refl by fastforce
lemma path_connected_space_iff_path_component:
"path_connected_space X \<longleftrightarrow> (\<forall>x \<in> topspace X. \<forall>y \<in> topspace X. path_component_of X x y)"
by (simp add: path_component_of path_connected_space_subconnected)
lemma path_connected_space_imp_path_component_of:
"\<lbrakk>path_connected_space X; a \<in> topspace X; b \<in> topspace X\<rbrakk>
\<Longrightarrow> path_component_of X a b"
by (simp add: path_connected_space_iff_path_component)
lemma path_connected_space_path_component_set:
"path_connected_space X \<longleftrightarrow> (\<forall>x \<in> topspace X. Collect(path_component_of X x) = topspace X)"
using path_component_of_subset_topspace path_connected_space_iff_path_component by fastforce
lemma path_component_of_maximal:
"\<lbrakk>path_connectedin X s; x \<in> s\<rbrakk> \<Longrightarrow> s \<subseteq> Collect(path_component_of X x)"
using path_component_of by fastforce
lemma path_component_of_equiv:
"path_component_of X x y \<longleftrightarrow> x \<in> topspace X \<and> y \<in> topspace X \<and> path_component_of X x = path_component_of X y"
(is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
apply (simp add: fun_eq_iff path_component_in_topspace)
apply (meson path_component_of_sym path_component_of_trans)
done
qed (simp add: path_component_of_refl)
lemma path_component_of_disjoint:
"disjnt (Collect (path_component_of X x)) (Collect (path_component_of X y)) \<longleftrightarrow>
~(path_component_of X x y)"
by (force simp: disjnt_def path_component_of_eq_empty path_component_of_equiv)
lemma path_component_of_eq:
"path_component_of X x = path_component_of X y \<longleftrightarrow>
(x \<notin> topspace X) \<and> (y \<notin> topspace X) \<or>
x \<in> topspace X \<and> y \<in> topspace X \<and> path_component_of X x y"
by (metis Collect_empty_eq_bot path_component_of_eq_empty path_component_of_equiv)
lemma path_connectedin_path_component_of:
"path_connectedin X (Collect (path_component_of X x))"
proof -
have "\<And>y. path_component_of X x y
\<Longrightarrow> path_component_of (subtopology X (Collect (path_component_of X x))) x y"
by (meson path_component_of path_component_of_maximal path_connectedin_subtopology)
then show ?thesis
apply (simp add: path_connectedin_def path_component_of_subset_topspace path_connected_space_iff_path_component)
by (metis Int_absorb1 mem_Collect_eq path_component_of_equiv path_component_of_subset_topspace topspace_subtopology)
qed
lemma path_connectedin_euclidean [simp]:
"path_connectedin euclidean S \<longleftrightarrow> path_connected S"
by (auto simp: path_connectedin_def path_connected_space_iff_path_component path_connected_component)
lemma path_connected_space_euclidean_subtopology [simp]:
"path_connected_space(subtopology euclidean S) \<longleftrightarrow> path_connected S"
using path_connectedin_topspace by force
lemma Union_path_components_of:
"\<Union>(path_components_of X) = topspace X"
by (auto simp: path_components_of_def path_component_of_equiv)
lemma path_components_of_maximal:
"\<lbrakk>C \<in> path_components_of X; path_connectedin X S; ~disjnt C S\<rbrakk> \<Longrightarrow> S \<subseteq> C"
apply (auto simp: path_components_of_def path_component_of_equiv)
using path_component_of_maximal path_connectedin_def apply fastforce
by (meson disjnt_subset2 path_component_of_disjoint path_component_of_equiv path_component_of_maximal)
lemma pairwise_disjoint_path_components_of:
"pairwise disjnt (path_components_of X)"
by (auto simp: path_components_of_def pairwise_def path_component_of_disjoint path_component_of_equiv)
lemma complement_path_components_of_Union:
"C \<in> path_components_of X
\<Longrightarrow> topspace X - C = \<Union>(path_components_of X - {C})"
by (metis Diff_cancel Diff_subset Union_path_components_of cSup_singleton diff_Union_pairwise_disjoint insert_subset pairwise_disjoint_path_components_of)
lemma nonempty_path_components_of:
"C \<in> path_components_of X \<Longrightarrow> (C \<noteq> {})"
apply (clarsimp simp: path_components_of_def path_component_of_eq_empty)
by (meson path_component_of_refl)
lemma path_components_of_subset: "C \<in> path_components_of X \<Longrightarrow> C \<subseteq> topspace X"
by (auto simp: path_components_of_def path_component_of_equiv)
lemma path_connectedin_path_components_of:
"C \<in> path_components_of X \<Longrightarrow> path_connectedin X C"
by (auto simp: path_components_of_def path_connectedin_path_component_of)
lemma path_component_in_path_components_of:
"Collect (path_component_of X a) \<in> path_components_of X \<longleftrightarrow> a \<in> topspace X"
apply (rule iffI)
using nonempty_path_components_of path_component_of_eq_empty apply fastforce
by (simp add: path_components_of_def)
lemma path_connectedin_Union:
assumes \<A>: "\<And>S. S \<in> \<A> \<Longrightarrow> path_connectedin X S" "\<Inter>\<A> \<noteq> {}"
shows "path_connectedin X (\<Union>\<A>)"
proof -
obtain a where "\<And>S. S \<in> \<A> \<Longrightarrow> a \<in> S"
using assms by blast
then have "\<And>x. x \<in> topspace (subtopology X (\<Union>\<A>)) \<Longrightarrow> path_component_of (subtopology X (\<Union>\<A>)) a x"
apply (simp)
by (meson Union_upper \<A> path_component_of path_connectedin_subtopology)
then show ?thesis
using \<A> unfolding path_connectedin_def
by (metis Sup_le_iff path_component_of_equiv path_connected_space_iff_path_component)
qed
lemma path_connectedin_Un:
"\<lbrakk>path_connectedin X S; path_connectedin X T; S \<inter> T \<noteq> {}\<rbrakk>
\<Longrightarrow> path_connectedin X (S \<union> T)"
by (blast intro: path_connectedin_Union [of "{S,T}", simplified])
lemma path_connected_space_iff_components_eq:
"path_connected_space X \<longleftrightarrow>
(\<forall>C \<in> path_components_of X. \<forall>C' \<in> path_components_of X. C = C')"
unfolding path_components_of_def
proof (intro iffI ballI)
assume "\<forall>C \<in> path_component_of_set X ` topspace X.
\<forall>C' \<in> path_component_of_set X ` topspace X. C = C'"
then show "path_connected_space X"
using path_component_of_refl path_connected_space_iff_path_component by fastforce
qed (auto simp: path_connected_space_path_component_set)
lemma path_components_of_eq_empty:
"path_components_of X = {} \<longleftrightarrow> topspace X = {}"
using Union_path_components_of nonempty_path_components_of by fastforce
lemma path_components_of_empty_space:
"topspace X = {} \<Longrightarrow> path_components_of X = {}"
by (simp add: path_components_of_eq_empty)
lemma path_components_of_subset_singleton:
"path_components_of X \<subseteq> {S} \<longleftrightarrow>
path_connected_space X \<and> (topspace X = {} \<or> topspace X = S)"
proof (cases "topspace X = {}")
case True
then show ?thesis
by (auto simp: path_components_of_empty_space path_connected_space_topspace_empty)
next
case False
have "(path_components_of X = {S}) \<longleftrightarrow> (path_connected_space X \<and> topspace X = S)"
proof (intro iffI conjI)
assume L: "path_components_of X = {S}"
then show "path_connected_space X"
by (simp add: path_connected_space_iff_components_eq)
show "topspace X = S"
by (metis L ccpo_Sup_singleton [of S] Union_path_components_of)
next
assume R: "path_connected_space X \<and> topspace X = S"
then show "path_components_of X = {S}"
using ccpo_Sup_singleton [of S]
by (metis False all_not_in_conv insert_iff mk_disjoint_insert path_component_in_path_components_of path_connected_space_iff_components_eq path_connected_space_path_component_set)
qed
with False show ?thesis
by (simp add: path_components_of_eq_empty subset_singleton_iff)
qed
lemma path_connected_space_iff_components_subset_singleton:
"path_connected_space X \<longleftrightarrow> (\<exists>a. path_components_of X \<subseteq> {a})"
by (simp add: path_components_of_subset_singleton)
lemma path_components_of_eq_singleton:
"path_components_of X = {S} \<longleftrightarrow> path_connected_space X \<and> topspace X \<noteq> {} \<and> S = topspace X"
by (metis cSup_singleton insert_not_empty path_components_of_subset_singleton subset_singleton_iff)
lemma path_components_of_path_connected_space:
"path_connected_space X \<Longrightarrow> path_components_of X = (if topspace X = {} then {} else {topspace X})"
by (simp add: path_components_of_eq_empty path_components_of_eq_singleton)
lemma path_component_subset_connected_component_of:
"path_component_of_set X x \<subseteq> connected_component_of_set X x"
proof (cases "x \<in> topspace X")
case True
then show ?thesis
by (simp add: connected_component_of_maximal path_component_of_refl path_connectedin_imp_connectedin path_connectedin_path_component_of)
next
case False
then show ?thesis
using path_component_of_eq_empty by fastforce
qed
lemma exists_path_component_of_superset:
assumes S: "path_connectedin X S" and ne: "topspace X \<noteq> {}"
obtains C where "C \<in> path_components_of X" "S \<subseteq> C"
proof (cases "S = {}")
case True
then show ?thesis
using ne path_components_of_eq_empty that by fastforce
next
case False
then obtain a where "a \<in> S"
by blast
show ?thesis
proof
show "Collect (path_component_of X a) \<in> path_components_of X"
by (meson \<open>a \<in> S\<close> S subsetD path_component_in_path_components_of path_connectedin_subset_topspace)
show "S \<subseteq> Collect (path_component_of X a)"
by (simp add: S \<open>a \<in> S\<close> path_component_of_maximal)
qed
qed
lemma path_component_of_eq_overlap:
"path_component_of X x = path_component_of X y \<longleftrightarrow>
(x \<notin> topspace X) \<and> (y \<notin> topspace X) \<or>
Collect (path_component_of X x) \<inter> Collect (path_component_of X y) \<noteq> {}"
by (metis disjnt_def empty_iff inf_bot_right mem_Collect_eq path_component_of_disjoint path_component_of_eq path_component_of_eq_empty)
lemma path_component_of_nonoverlap:
"Collect (path_component_of X x) \<inter> Collect (path_component_of X y) = {} \<longleftrightarrow>
(x \<notin> topspace X) \<or> (y \<notin> topspace X) \<or>
path_component_of X x \<noteq> path_component_of X y"
by (metis inf.idem path_component_of_eq_empty path_component_of_eq_overlap)
lemma path_component_of_overlap:
"Collect (path_component_of X x) \<inter> Collect (path_component_of X y) \<noteq> {} \<longleftrightarrow>
x \<in> topspace X \<and> y \<in> topspace X \<and> path_component_of X x = path_component_of X y"
by (meson path_component_of_nonoverlap)
lemma path_components_of_disjoint:
"\<lbrakk>C \<in> path_components_of X; C' \<in> path_components_of X\<rbrakk> \<Longrightarrow> disjnt C C' \<longleftrightarrow> C \<noteq> C'"
by (auto simp: path_components_of_def path_component_of_disjoint path_component_of_equiv)
lemma path_components_of_overlap:
"\<lbrakk>C \<in> path_components_of X; C' \<in> path_components_of X\<rbrakk> \<Longrightarrow> C \<inter> C' \<noteq> {} \<longleftrightarrow> C = C'"
by (auto simp: path_components_of_def path_component_of_equiv)
lemma path_component_of_unique:
"\<lbrakk>x \<in> C; path_connectedin X C; \<And>C'. \<lbrakk>x \<in> C'; path_connectedin X C'\<rbrakk> \<Longrightarrow> C' \<subseteq> C\<rbrakk>
\<Longrightarrow> Collect (path_component_of X x) = C"
by (meson subsetD eq_iff path_component_of_maximal path_connectedin_path_component_of)
lemma path_component_of_discrete_topology [simp]:
"Collect (path_component_of (discrete_topology U) x) = (if x \<in> U then {x} else {})"
proof -
have "\<And>C'. \<lbrakk>x \<in> C'; path_connectedin (discrete_topology U) C'\<rbrakk> \<Longrightarrow> C' \<subseteq> {x}"
by (metis path_connectedin_discrete_topology subsetD singletonD)
then have "x \<in> U \<Longrightarrow> Collect (path_component_of (discrete_topology U) x) = {x}"
by (simp add: path_component_of_unique)
then show ?thesis
using path_component_in_topspace by fastforce
qed
lemma path_component_of_discrete_topology_iff [simp]:
"path_component_of (discrete_topology U) x y \<longleftrightarrow> x \<in> U \<and> y=x"
by (metis empty_iff insertI1 mem_Collect_eq path_component_of_discrete_topology singletonD)
lemma path_components_of_discrete_topology [simp]:
"path_components_of (discrete_topology U) = (\<lambda>x. {x}) ` U"
by (auto simp: path_components_of_def image_def fun_eq_iff)
lemma homeomorphic_map_path_component_of:
assumes f: "homeomorphic_map X Y f" and x: "x \<in> topspace X"
shows "Collect (path_component_of Y (f x)) = f ` Collect(path_component_of X x)"
proof -
obtain g where g: "homeomorphic_maps X Y f g"
using f homeomorphic_map_maps by blast
show ?thesis
proof
have "Collect (path_component_of Y (f x)) \<subseteq> topspace Y"
by (simp add: path_component_of_subset_topspace)
moreover have "g ` Collect(path_component_of Y (f x)) \<subseteq> Collect (path_component_of X (g (f x)))"
using g x unfolding homeomorphic_maps_def
by (metis f homeomorphic_imp_surjective_map imageI mem_Collect_eq path_component_of_maximal path_component_of_refl path_connectedin_continuous_map_image path_connectedin_path_component_of)
ultimately show "Collect (path_component_of Y (f x)) \<subseteq> f ` Collect (path_component_of X x)"
using g x unfolding homeomorphic_maps_def continuous_map_def image_iff subset_iff
by metis
show "f ` Collect (path_component_of X x) \<subseteq> Collect (path_component_of Y (f x))"
proof (rule path_component_of_maximal)
show "path_connectedin Y (f ` Collect (path_component_of X x))"
by (meson f homeomorphic_map_path_connectedness_eq path_connectedin_path_component_of)
qed (simp add: path_component_of_refl x)
qed
qed
lemma homeomorphic_map_path_components_of:
assumes "homeomorphic_map X Y f"
shows "path_components_of Y = (image f) ` (path_components_of X)"
unfolding path_components_of_def homeomorphic_imp_surjective_map [OF assms, symmetric]
apply safe
using assms homeomorphic_map_path_component_of apply fastforce
by (metis assms homeomorphic_map_path_component_of imageI)
subsection \<open>Sphere is path-connected\<close>
lemma path_connected_punctured_universe:
assumes "2 \<le> DIM('a::euclidean_space)"
shows "path_connected (- {a::'a})"
proof -
let ?A = "{x::'a. \<exists>i\<in>Basis. x \<bullet> i < a \<bullet> i}"
let ?B = "{x::'a. \<exists>i\<in>Basis. a \<bullet> i < x \<bullet> i}"
have A: "path_connected ?A"
unfolding Collect_bex_eq
proof (rule path_connected_UNION)
fix i :: 'a
assume "i \<in> Basis"
then show "(\<Sum>i\<in>Basis. (a \<bullet> i - 1)*\<^sub>R i) \<in> {x::'a. x \<bullet> i < a \<bullet> i}"
by simp
show "path_connected {x. x \<bullet> i < a \<bullet> i}"
using convex_imp_path_connected [OF convex_halfspace_lt, of i "a \<bullet> i"]
by (simp add: inner_commute)
qed
have B: "path_connected ?B"
unfolding Collect_bex_eq
proof (rule path_connected_UNION)
fix i :: 'a
assume "i \<in> Basis"
then show "(\<Sum>i\<in>Basis. (a \<bullet> i + 1) *\<^sub>R i) \<in> {x::'a. a \<bullet> i < x \<bullet> i}"
by simp
show "path_connected {x. a \<bullet> i < x \<bullet> i}"
using convex_imp_path_connected [OF convex_halfspace_gt, of "a \<bullet> i" i]
by (simp add: inner_commute)
qed
obtain S :: "'a set" where "S \<subseteq> Basis" and "card S = Suc (Suc 0)"
using ex_card[OF assms]
by auto
then obtain b0 b1 :: 'a where "b0 \<in> Basis" and "b1 \<in> Basis" and "b0 \<noteq> b1"
unfolding card_Suc_eq by auto
then have "a + b0 - b1 \<in> ?A \<inter> ?B"
by (auto simp: inner_simps inner_Basis)
then have "?A \<inter> ?B \<noteq> {}"
by fast
with A B have "path_connected (?A \<union> ?B)"
by (rule path_connected_Un)
also have "?A \<union> ?B = {x. \<exists>i\<in>Basis. x \<bullet> i \<noteq> a \<bullet> i}"
unfolding neq_iff bex_disj_distrib Collect_disj_eq ..
also have "\<dots> = {x. x \<noteq> a}"
unfolding euclidean_eq_iff [where 'a='a]
by (simp add: Bex_def)
also have "\<dots> = - {a}"
by auto
finally show ?thesis .
qed
corollary connected_punctured_universe:
"2 \<le> DIM('N::euclidean_space) \<Longrightarrow> connected(- {a::'N})"
by (simp add: path_connected_punctured_universe path_connected_imp_connected)
proposition path_connected_sphere:
fixes a :: "'a :: euclidean_space"
assumes "2 \<le> DIM('a)"
shows "path_connected(sphere a r)"
proof (cases r "0::real" rule: linorder_cases)
case less
then show ?thesis
by (simp)
next
case equal
then show ?thesis
by (simp)
next
case greater
then have eq: "(sphere (0::'a) r) = (\<lambda>x. (r / norm x) *\<^sub>R x) ` (- {0::'a})"
by (force simp: image_iff split: if_split_asm)
have "continuous_on (- {0::'a}) (\<lambda>x. (r / norm x) *\<^sub>R x)"
by (intro continuous_intros) auto
then have "path_connected ((\<lambda>x. (r / norm x) *\<^sub>R x) ` (- {0::'a}))"
by (intro path_connected_continuous_image path_connected_punctured_universe assms)
with eq have "path_connected (sphere (0::'a) r)"
by auto
then have "path_connected((+) a ` (sphere (0::'a) r))"
by (simp add: path_connected_translation)
then show ?thesis
by (metis add.right_neutral sphere_translation)
qed
lemma connected_sphere:
fixes a :: "'a :: euclidean_space"
assumes "2 \<le> DIM('a)"
shows "connected(sphere a r)"
using path_connected_sphere [OF assms]
by (simp add: path_connected_imp_connected)
corollary path_connected_complement_bounded_convex:
fixes s :: "'a :: euclidean_space set"
assumes "bounded s" "convex s" and 2: "2 \<le> DIM('a)"
shows "path_connected (- s)"
proof (cases "s = {}")
case True then show ?thesis
using convex_imp_path_connected by auto
next
case False
then obtain a where "a \<in> s" by auto
{ fix x y assume "x \<notin> s" "y \<notin> s"
then have "x \<noteq> a" "y \<noteq> a" using \<open>a \<in> s\<close> by auto
then have bxy: "bounded(insert x (insert y s))"
by (simp add: \<open>bounded s\<close>)
then obtain B::real where B: "0 < B" and Bx: "norm (a - x) < B" and By: "norm (a - y) < B"
and "s \<subseteq> ball a B"
using bounded_subset_ballD [OF bxy, of a] by (auto simp: dist_norm)
define C where "C = B / norm(x - a)"
{ fix u
assume u: "(1 - u) *\<^sub>R x + u *\<^sub>R (a + C *\<^sub>R (x - a)) \<in> s" and "0 \<le> u" "u \<le> 1"
have CC: "1 \<le> 1 + (C - 1) * u"
using \<open>x \<noteq> a\<close> \<open>0 \<le> u\<close> Bx
by (auto simp add: C_def norm_minus_commute)
have *: "\<And>v. (1 - u) *\<^sub>R x + u *\<^sub>R (a + v *\<^sub>R (x - a)) = a + (1 + (v - 1) * u) *\<^sub>R (x - a)"
by (simp add: algebra_simps)
have "a + ((1 / (1 + C * u - u)) *\<^sub>R x + ((u / (1 + C * u - u)) *\<^sub>R a + (C * u / (1 + C * u - u)) *\<^sub>R x)) =
(1 + (u / (1 + C * u - u))) *\<^sub>R a + ((1 / (1 + C * u - u)) + (C * u / (1 + C * u - u))) *\<^sub>R x"
by (simp add: algebra_simps)
also have "\<dots> = (1 + (u / (1 + C * u - u))) *\<^sub>R a + (1 + (u / (1 + C * u - u))) *\<^sub>R x"
using CC by (simp add: field_simps)
also have "\<dots> = x + (1 + (u / (1 + C * u - u))) *\<^sub>R a + (u / (1 + C * u - u)) *\<^sub>R x"
by (simp add: algebra_simps)
also have "\<dots> = x + ((1 / (1 + C * u - u)) *\<^sub>R a +
((u / (1 + C * u - u)) *\<^sub>R x + (C * u / (1 + C * u - u)) *\<^sub>R a))"
using CC by (simp add: field_simps) (simp add: add_divide_distrib scaleR_add_left)
finally have xeq: "(1 - 1 / (1 + (C - 1) * u)) *\<^sub>R a + (1 / (1 + (C - 1) * u)) *\<^sub>R (a + (1 + (C - 1) * u) *\<^sub>R (x - a)) = x"
by (simp add: algebra_simps)
have False
using \<open>convex s\<close>
apply (simp add: convex_alt)
apply (drule_tac x=a in bspec)
apply (rule \<open>a \<in> s\<close>)
apply (drule_tac x="a + (1 + (C - 1) * u) *\<^sub>R (x - a)" in bspec)
using u apply (simp add: *)
apply (drule_tac x="1 / (1 + (C - 1) * u)" in spec)
using \<open>x \<noteq> a\<close> \<open>x \<notin> s\<close> \<open>0 \<le> u\<close> CC
apply (auto simp: xeq)
done
}
then have pcx: "path_component (- s) x (a + C *\<^sub>R (x - a))"
by (force simp: closed_segment_def intro!: path_component_linepath)
define D where "D = B / norm(y - a)" \<comment> \<open>massive duplication with the proof above\<close>
{ fix u
assume u: "(1 - u) *\<^sub>R y + u *\<^sub>R (a + D *\<^sub>R (y - a)) \<in> s" and "0 \<le> u" "u \<le> 1"
have DD: "1 \<le> 1 + (D - 1) * u"
using \<open>y \<noteq> a\<close> \<open>0 \<le> u\<close> By
by (auto simp add: D_def norm_minus_commute)
have *: "\<And>v. (1 - u) *\<^sub>R y + u *\<^sub>R (a + v *\<^sub>R (y - a)) = a + (1 + (v - 1) * u) *\<^sub>R (y - a)"
by (simp add: algebra_simps)
have "a + ((1 / (1 + D * u - u)) *\<^sub>R y + ((u / (1 + D * u - u)) *\<^sub>R a + (D * u / (1 + D * u - u)) *\<^sub>R y)) =
(1 + (u / (1 + D * u - u))) *\<^sub>R a + ((1 / (1 + D * u - u)) + (D * u / (1 + D * u - u))) *\<^sub>R y"
by (simp add: algebra_simps)
also have "\<dots> = (1 + (u / (1 + D * u - u))) *\<^sub>R a + (1 + (u / (1 + D * u - u))) *\<^sub>R y"
using DD by (simp add: field_simps)
also have "\<dots> = y + (1 + (u / (1 + D * u - u))) *\<^sub>R a + (u / (1 + D * u - u)) *\<^sub>R y"
by (simp add: algebra_simps)
also have "\<dots> = y + ((1 / (1 + D * u - u)) *\<^sub>R a +
((u / (1 + D * u - u)) *\<^sub>R y + (D * u / (1 + D * u - u)) *\<^sub>R a))"
using DD by (simp add: field_simps) (simp add: add_divide_distrib scaleR_add_left)
finally have xeq: "(1 - 1 / (1 + (D - 1) * u)) *\<^sub>R a + (1 / (1 + (D - 1) * u)) *\<^sub>R (a + (1 + (D - 1) * u) *\<^sub>R (y - a)) = y"
by (simp add: algebra_simps)
have False
using \<open>convex s\<close>
apply (simp add: convex_alt)
apply (drule_tac x=a in bspec)
apply (rule \<open>a \<in> s\<close>)
apply (drule_tac x="a + (1 + (D - 1) * u) *\<^sub>R (y - a)" in bspec)
using u apply (simp add: *)
apply (drule_tac x="1 / (1 + (D - 1) * u)" in spec)
using \<open>y \<noteq> a\<close> \<open>y \<notin> s\<close> \<open>0 \<le> u\<close> DD
apply (auto simp: xeq)
done
}
then have pdy: "path_component (- s) y (a + D *\<^sub>R (y - a))"
by (force simp: closed_segment_def intro!: path_component_linepath)
have pyx: "path_component (- s) (a + D *\<^sub>R (y - a)) (a + C *\<^sub>R (x - a))"
apply (rule path_component_of_subset [of "sphere a B"])
using \<open>s \<subseteq> ball a B\<close>
apply (force simp: ball_def dist_norm norm_minus_commute)
apply (rule path_connected_sphere [OF 2, of a B, simplified path_connected_component, rule_format])
using \<open>x \<noteq> a\<close> using \<open>y \<noteq> a\<close> B apply (auto simp: dist_norm C_def D_def)
done
have "path_component (- s) x y"
by (metis path_component_trans path_component_sym pcx pdy pyx)
}
then show ?thesis
by (auto simp: path_connected_component)
qed
lemma connected_complement_bounded_convex:
fixes s :: "'a :: euclidean_space set"
assumes "bounded s" "convex s" "2 \<le> DIM('a)"
shows "connected (- s)"
using path_connected_complement_bounded_convex [OF assms] path_connected_imp_connected by blast
lemma connected_diff_ball:
fixes s :: "'a :: euclidean_space set"
assumes "connected s" "cball a r \<subseteq> s" "2 \<le> DIM('a)"
shows "connected (s - ball a r)"
apply (rule connected_diff_open_from_closed [OF ball_subset_cball])
using assms connected_sphere
apply (auto simp: cball_diff_eq_sphere dist_norm)
done
proposition connected_open_delete:
assumes "open S" "connected S" and 2: "2 \<le> DIM('N::euclidean_space)"
shows "connected(S - {a::'N})"
proof (cases "a \<in> S")
case True
with \<open>open S\<close> obtain \<epsilon> where "\<epsilon> > 0" and \<epsilon>: "cball a \<epsilon> \<subseteq> S"
using open_contains_cball_eq by blast
have "dist a (a + \<epsilon> *\<^sub>R (SOME i. i \<in> Basis)) = \<epsilon>"
by (simp add: dist_norm SOME_Basis \<open>0 < \<epsilon>\<close> less_imp_le)
with \<epsilon> have "\<Inter>{S - ball a r |r. 0 < r \<and> r < \<epsilon>} \<subseteq> {} \<Longrightarrow> False"
apply (drule_tac c="a + scaleR (\<epsilon>) ((SOME i. i \<in> Basis))" in subsetD)
by auto
then have nonemp: "(\<Inter>{S - ball a r |r. 0 < r \<and> r < \<epsilon>}) = {} \<Longrightarrow> False"
by auto
have con: "\<And>r. r < \<epsilon> \<Longrightarrow> connected (S - ball a r)"
using \<epsilon> by (force intro: connected_diff_ball [OF \<open>connected S\<close> _ 2])
have "x \<in> \<Union>{S - ball a r |r. 0 < r \<and> r < \<epsilon>}" if "x \<in> S - {a}" for x
apply (rule UnionI [of "S - ball a (min \<epsilon> (dist a x) / 2)"])
using that \<open>0 < \<epsilon>\<close> apply simp_all
apply (rule_tac x="min \<epsilon> (dist a x) / 2" in exI)
apply auto
done
then have "S - {a} = \<Union>{S - ball a r | r. 0 < r \<and> r < \<epsilon>}"
by auto
then show ?thesis
by (auto intro: connected_Union con dest!: nonemp)
next
case False then show ?thesis
by (simp add: \<open>connected S\<close>)
qed
corollary path_connected_open_delete:
assumes "open S" "connected S" and 2: "2 \<le> DIM('N::euclidean_space)"
shows "path_connected(S - {a::'N})"
by (simp add: assms connected_open_delete connected_open_path_connected open_delete)
corollary path_connected_punctured_ball:
"2 \<le> DIM('N::euclidean_space) \<Longrightarrow> path_connected(ball a r - {a::'N})"
by (simp add: path_connected_open_delete)
corollary connected_punctured_ball:
"2 \<le> DIM('N::euclidean_space) \<Longrightarrow> connected(ball a r - {a::'N})"
by (simp add: connected_open_delete)
corollary connected_open_delete_finite:
fixes S T::"'a::euclidean_space set"
assumes S: "open S" "connected S" and 2: "2 \<le> DIM('a)" and "finite T"
shows "connected(S - T)"
using \<open>finite T\<close> S
proof (induct T)
case empty
show ?case using \<open>connected S\<close> by simp
next
case (insert x F)
then have "connected (S-F)" by auto
moreover have "open (S - F)" using finite_imp_closed[OF \<open>finite F\<close>] \<open>open S\<close> by auto
ultimately have "connected (S - F - {x})" using connected_open_delete[OF _ _ 2] by auto
thus ?case by (metis Diff_insert)
qed
lemma sphere_1D_doubleton_zero:
assumes 1: "DIM('a) = 1" and "r > 0"
obtains x y::"'a::euclidean_space"
where "sphere 0 r = {x,y} \<and> dist x y = 2*r"
proof -
obtain b::'a where b: "Basis = {b}"
using 1 card_1_singletonE by blast
show ?thesis
proof (intro that conjI)
have "x = norm x *\<^sub>R b \<or> x = - norm x *\<^sub>R b" if "r = norm x" for x
proof -
have xb: "(x \<bullet> b) *\<^sub>R b = x"
using euclidean_representation [of x, unfolded b] by force
then have "norm ((x \<bullet> b) *\<^sub>R b) = norm x"
by simp
with b have "\<bar>x \<bullet> b\<bar> = norm x"
using norm_Basis by (simp add: b)
with xb show ?thesis
apply (simp add: abs_if split: if_split_asm)
apply (metis add.inverse_inverse real_vector.scale_minus_left xb)
done
qed
with \<open>r > 0\<close> b show "sphere 0 r = {r *\<^sub>R b, - r *\<^sub>R b}"
by (force simp: sphere_def dist_norm)
have "dist (r *\<^sub>R b) (- r *\<^sub>R b) = norm (r *\<^sub>R b + r *\<^sub>R b)"
by (simp add: dist_norm)
also have "\<dots> = norm ((2*r) *\<^sub>R b)"
by (metis mult_2 scaleR_add_left)
also have "\<dots> = 2*r"
using \<open>r > 0\<close> b norm_Basis by fastforce
finally show "dist (r *\<^sub>R b) (- r *\<^sub>R b) = 2*r" .
qed
qed
lemma sphere_1D_doubleton:
fixes a :: "'a :: euclidean_space"
assumes "DIM('a) = 1" and "r > 0"
obtains x y where "sphere a r = {x,y} \<and> dist x y = 2*r"
proof -
have "sphere a r = (+) a ` sphere 0 r"
by (metis add.right_neutral sphere_translation)
then show ?thesis
using sphere_1D_doubleton_zero [OF assms]
by (metis (mono_tags, lifting) dist_add_cancel image_empty image_insert that)
qed
lemma psubset_sphere_Compl_connected:
fixes S :: "'a::euclidean_space set"
assumes S: "S \<subset> sphere a r" and "0 < r" and 2: "2 \<le> DIM('a)"
shows "connected(- S)"
proof -
have "S \<subseteq> sphere a r"
using S by blast
obtain b where "dist a b = r" and "b \<notin> S"
using S mem_sphere by blast
have CS: "- S = {x. dist a x \<le> r \<and> (x \<notin> S)} \<union> {x. r \<le> dist a x \<and> (x \<notin> S)}"
by auto
have "{x. dist a x \<le> r \<and> x \<notin> S} \<inter> {x. r \<le> dist a x \<and> x \<notin> S} \<noteq> {}"
using \<open>b \<notin> S\<close> \<open>dist a b = r\<close> by blast
moreover have "connected {x. dist a x \<le> r \<and> x \<notin> S}"
apply (rule connected_intermediate_closure [of "ball a r"])
using assms by auto
moreover
have "connected {x. r \<le> dist a x \<and> x \<notin> S}"
apply (rule connected_intermediate_closure [of "- cball a r"])
using assms apply (auto intro: connected_complement_bounded_convex)
apply (metis ComplI interior_cball interior_closure mem_ball not_less)
done
ultimately show ?thesis
by (simp add: CS connected_Un)
qed
subsection\<open>Every annulus is a connected set\<close>
lemma path_connected_2DIM_I:
fixes a :: "'N::euclidean_space"
assumes 2: "2 \<le> DIM('N)" and pc: "path_connected {r. 0 \<le> r \<and> P r}"
shows "path_connected {x. P(norm(x - a))}"
proof -
have "{x. P(norm(x - a))} = (+) a ` {x. P(norm x)}"
by force
moreover have "path_connected {x::'N. P(norm x)}"
proof -
let ?D = "{x. 0 \<le> x \<and> P x} \<times> sphere (0::'N) 1"
have "x \<in> (\<lambda>z. fst z *\<^sub>R snd z) ` ?D"
if "P (norm x)" for x::'N
proof (cases "x=0")
case True
with that show ?thesis
apply (simp add: image_iff)
apply (rule_tac x=0 in exI, simp)
using vector_choose_size zero_le_one by blast
next
case False
with that show ?thesis
by (rule_tac x="(norm x, x /\<^sub>R norm x)" in image_eqI) auto
qed
then have *: "{x::'N. P(norm x)} = (\<lambda>z. fst z *\<^sub>R snd z) ` ?D"
by auto
have "continuous_on ?D (\<lambda>z:: real\<times>'N. fst z *\<^sub>R snd z)"
by (intro continuous_intros)
moreover have "path_connected ?D"
by (metis path_connected_Times [OF pc] path_connected_sphere 2)
ultimately show ?thesis
apply (subst *)
apply (rule path_connected_continuous_image, auto)
done
qed
ultimately show ?thesis
using path_connected_translation by metis
qed
proposition path_connected_annulus:
fixes a :: "'N::euclidean_space"
assumes "2 \<le> DIM('N)"
shows "path_connected {x. r1 < norm(x - a) \<and> norm(x - a) < r2}"
"path_connected {x. r1 < norm(x - a) \<and> norm(x - a) \<le> r2}"
"path_connected {x. r1 \<le> norm(x - a) \<and> norm(x - a) < r2}"
"path_connected {x. r1 \<le> norm(x - a) \<and> norm(x - a) \<le> r2}"
by (auto simp: is_interval_def intro!: is_interval_convex convex_imp_path_connected path_connected_2DIM_I [OF assms])
proposition connected_annulus:
fixes a :: "'N::euclidean_space"
assumes "2 \<le> DIM('N::euclidean_space)"
shows "connected {x. r1 < norm(x - a) \<and> norm(x - a) < r2}"
"connected {x. r1 < norm(x - a) \<and> norm(x - a) \<le> r2}"
"connected {x. r1 \<le> norm(x - a) \<and> norm(x - a) < r2}"
"connected {x. r1 \<le> norm(x - a) \<and> norm(x - a) \<le> r2}"
by (auto simp: path_connected_annulus [OF assms] path_connected_imp_connected)
subsection\<^marker>\<open>tag unimportant\<close>\<open>Relations between components and path components\<close>
lemma open_connected_component:
fixes s :: "'a::real_normed_vector set"
shows "open s \<Longrightarrow> open (connected_component_set s x)"
apply (simp add: open_contains_ball, clarify)
apply (rename_tac y)
apply (drule_tac x=y in bspec)
apply (simp add: connected_component_in, clarify)
apply (rule_tac x=e in exI)
by (metis mem_Collect_eq connected_component_eq connected_component_maximal centre_in_ball connected_ball)
corollary open_components:
fixes s :: "'a::real_normed_vector set"
shows "\<lbrakk>open u; s \<in> components u\<rbrakk> \<Longrightarrow> open s"
by (simp add: components_iff) (metis open_connected_component)
lemma in_closure_connected_component:
fixes s :: "'a::real_normed_vector set"
assumes x: "x \<in> s" and s: "open s"
shows "x \<in> closure (connected_component_set s y) \<longleftrightarrow> x \<in> connected_component_set s y"
proof -
{ assume "x \<in> closure (connected_component_set s y)"
moreover have "x \<in> connected_component_set s x"
using x by simp
ultimately have "x \<in> connected_component_set s y"
using s by (meson Compl_disjoint closure_iff_nhds_not_empty connected_component_disjoint disjoint_eq_subset_Compl open_connected_component)
}
then show ?thesis
by (auto simp: closure_def)
qed
lemma connected_disjoint_Union_open_pick:
assumes "pairwise disjnt B"
"\<And>S. S \<in> A \<Longrightarrow> connected S \<and> S \<noteq> {}"
"\<And>S. S \<in> B \<Longrightarrow> open S"
"\<Union>A \<subseteq> \<Union>B"
"S \<in> A"
obtains T where "T \<in> B" "S \<subseteq> T" "S \<inter> \<Union>(B - {T}) = {}"
proof -
have "S \<subseteq> \<Union>B" "connected S" "S \<noteq> {}"
using assms \<open>S \<in> A\<close> by blast+
then obtain T where "T \<in> B" "S \<inter> T \<noteq> {}"
by (metis Sup_inf_eq_bot_iff inf.absorb_iff2 inf_commute)
have 1: "open T" by (simp add: \<open>T \<in> B\<close> assms)
have 2: "open (\<Union>(B-{T}))" using assms by blast
have 3: "S \<subseteq> T \<union> \<Union>(B - {T})" using \<open>S \<subseteq> \<Union>B\<close> by blast
have "T \<inter> \<Union>(B - {T}) = {}" using \<open>T \<in> B\<close> \<open>pairwise disjnt B\<close>
by (auto simp: pairwise_def disjnt_def)
then have 4: "T \<inter> \<Union>(B - {T}) \<inter> S = {}" by auto
from connectedD [OF \<open>connected S\<close> 1 2 4 3]
have "S \<inter> \<Union>(B-{T}) = {}"
by (auto simp: Int_commute \<open>S \<inter> T \<noteq> {}\<close>)
with \<open>T \<in> B\<close> have "S \<subseteq> T"
using "3" by auto
show ?thesis
using \<open>S \<inter> \<Union>(B - {T}) = {}\<close> \<open>S \<subseteq> T\<close> \<open>T \<in> B\<close> that by auto
qed
lemma connected_disjoint_Union_open_subset:
assumes A: "pairwise disjnt A" and B: "pairwise disjnt B"
and SA: "\<And>S. S \<in> A \<Longrightarrow> open S \<and> connected S \<and> S \<noteq> {}"
and SB: "\<And>S. S \<in> B \<Longrightarrow> open S \<and> connected S \<and> S \<noteq> {}"
and eq [simp]: "\<Union>A = \<Union>B"
shows "A \<subseteq> B"
proof
fix S
assume "S \<in> A"
obtain T where "T \<in> B" "S \<subseteq> T" "S \<inter> \<Union>(B - {T}) = {}"
apply (rule connected_disjoint_Union_open_pick [OF B, of A])
using SA SB \<open>S \<in> A\<close> by auto
moreover obtain S' where "S' \<in> A" "T \<subseteq> S'" "T \<inter> \<Union>(A - {S'}) = {}"
apply (rule connected_disjoint_Union_open_pick [OF A, of B])
using SA SB \<open>T \<in> B\<close> by auto
ultimately have "S' = S"
by (metis A Int_subset_iff SA \<open>S \<in> A\<close> disjnt_def inf.orderE pairwise_def)
with \<open>T \<subseteq> S'\<close> have "T \<subseteq> S" by simp
with \<open>S \<subseteq> T\<close> have "S = T" by blast
with \<open>T \<in> B\<close> show "S \<in> B" by simp
qed
lemma connected_disjoint_Union_open_unique:
assumes A: "pairwise disjnt A" and B: "pairwise disjnt B"
and SA: "\<And>S. S \<in> A \<Longrightarrow> open S \<and> connected S \<and> S \<noteq> {}"
and SB: "\<And>S. S \<in> B \<Longrightarrow> open S \<and> connected S \<and> S \<noteq> {}"
and eq [simp]: "\<Union>A = \<Union>B"
shows "A = B"
by (rule subset_antisym; metis connected_disjoint_Union_open_subset assms)
proposition components_open_unique:
fixes S :: "'a::real_normed_vector set"
assumes "pairwise disjnt A" "\<Union>A = S"
"\<And>X. X \<in> A \<Longrightarrow> open X \<and> connected X \<and> X \<noteq> {}"
shows "components S = A"
proof -
have "open S" using assms by blast
show ?thesis
apply (rule connected_disjoint_Union_open_unique)
apply (simp add: components_eq disjnt_def pairwise_def)
using \<open>open S\<close>
apply (simp_all add: assms open_components in_components_connected in_components_nonempty)
done
qed
subsection\<^marker>\<open>tag unimportant\<close>\<open>Existence of unbounded components\<close>
lemma cobounded_unbounded_component:
fixes s :: "'a :: euclidean_space set"
assumes "bounded (-s)"
shows "\<exists>x. x \<in> s \<and> \<not> bounded (connected_component_set s x)"
proof -
obtain i::'a where i: "i \<in> Basis"
using nonempty_Basis by blast
obtain B where B: "B>0" "-s \<subseteq> ball 0 B"
using bounded_subset_ballD [OF assms, of 0] by auto
then have *: "\<And>x. B \<le> norm x \<Longrightarrow> x \<in> s"
by (force simp: ball_def dist_norm)
have unbounded_inner: "\<not> bounded {x. inner i x \<ge> B}"
apply (auto simp: bounded_def dist_norm)
apply (rule_tac x="x + (max B e + 1 + \<bar>i \<bullet> x\<bar>) *\<^sub>R i" in exI)
apply simp
using i
apply (auto simp: algebra_simps)
done
have **: "{x. B \<le> i \<bullet> x} \<subseteq> connected_component_set s (B *\<^sub>R i)"
apply (rule connected_component_maximal)
apply (auto simp: i intro: convex_connected convex_halfspace_ge [of B])
apply (rule *)
apply (rule order_trans [OF _ Basis_le_norm [OF i]])
by (simp add: inner_commute)
have "B *\<^sub>R i \<in> s"
by (rule *) (simp add: norm_Basis [OF i])
then show ?thesis
apply (rule_tac x="B *\<^sub>R i" in exI, clarify)
apply (frule bounded_subset [of _ "{x. B \<le> i \<bullet> x}", OF _ **])
using unbounded_inner apply blast
done
qed
lemma cobounded_unique_unbounded_component:
fixes s :: "'a :: euclidean_space set"
assumes bs: "bounded (-s)" and "2 \<le> DIM('a)"
and bo: "\<not> bounded(connected_component_set s x)"
"\<not> bounded(connected_component_set s y)"
shows "connected_component_set s x = connected_component_set s y"
proof -
obtain i::'a where i: "i \<in> Basis"
using nonempty_Basis by blast
obtain B where B: "B>0" "-s \<subseteq> ball 0 B"
using bounded_subset_ballD [OF bs, of 0] by auto
then have *: "\<And>x. B \<le> norm x \<Longrightarrow> x \<in> s"
by (force simp: ball_def dist_norm)
have ccb: "connected (- ball 0 B :: 'a set)"
using assms by (auto intro: connected_complement_bounded_convex)
obtain x' where x': "connected_component s x x'" "norm x' > B"
using bo [unfolded bounded_def dist_norm, simplified, rule_format]
by (metis diff_zero norm_minus_commute not_less)
obtain y' where y': "connected_component s y y'" "norm y' > B"
using bo [unfolded bounded_def dist_norm, simplified, rule_format]
by (metis diff_zero norm_minus_commute not_less)
have x'y': "connected_component s x' y'"
apply (simp add: connected_component_def)
apply (rule_tac x="- ball 0 B" in exI)
using x' y'
apply (auto simp: ccb dist_norm *)
done
show ?thesis
apply (rule connected_component_eq)
using x' y' x'y'
by (metis (no_types, lifting) connected_component_eq_empty connected_component_eq_eq connected_component_idemp connected_component_in)
qed
lemma cobounded_unbounded_components:
fixes s :: "'a :: euclidean_space set"
shows "bounded (-s) \<Longrightarrow> \<exists>c. c \<in> components s \<and> \<not>bounded c"
by (metis cobounded_unbounded_component components_def imageI)
lemma cobounded_unique_unbounded_components:
fixes s :: "'a :: euclidean_space set"
shows "\<lbrakk>bounded (- s); c \<in> components s; \<not> bounded c; c' \<in> components s; \<not> bounded c'; 2 \<le> DIM('a)\<rbrakk> \<Longrightarrow> c' = c"
unfolding components_iff
by (metis cobounded_unique_unbounded_component)
lemma cobounded_has_bounded_component:
fixes S :: "'a :: euclidean_space set"
assumes "bounded (- S)" "\<not> connected S" "2 \<le> DIM('a)"
obtains C where "C \<in> components S" "bounded C"
by (meson cobounded_unique_unbounded_components connected_eq_connected_components_eq assms)
subsection\<open>The \<open>inside\<close> and \<open>outside\<close> of a Set\<close>
text\<^marker>\<open>tag important\<close>\<open>The inside comprises the points in a bounded connected component of the set's complement.
The outside comprises the points in unbounded connected component of the complement.\<close>
definition\<^marker>\<open>tag important\<close> inside where
"inside S \<equiv> {x. (x \<notin> S) \<and> bounded(connected_component_set ( - S) x)}"
definition\<^marker>\<open>tag important\<close> outside where
"outside S \<equiv> -S \<inter> {x. \<not> bounded(connected_component_set (- S) x)}"
lemma outside: "outside S = {x. \<not> bounded(connected_component_set (- S) x)}"
by (auto simp: outside_def) (metis Compl_iff bounded_empty connected_component_eq_empty)
lemma inside_no_overlap [simp]: "inside S \<inter> S = {}"
by (auto simp: inside_def)
lemma outside_no_overlap [simp]:
"outside S \<inter> S = {}"
by (auto simp: outside_def)
lemma inside_Int_outside [simp]: "inside S \<inter> outside S = {}"
by (auto simp: inside_def outside_def)
lemma inside_Un_outside [simp]: "inside S \<union> outside S = (- S)"
by (auto simp: inside_def outside_def)
lemma inside_eq_outside:
"inside S = outside S \<longleftrightarrow> S = UNIV"
by (auto simp: inside_def outside_def)
lemma inside_outside: "inside S = (- (S \<union> outside S))"
by (force simp: inside_def outside)
lemma outside_inside: "outside S = (- (S \<union> inside S))"
by (auto simp: inside_outside) (metis IntI equals0D outside_no_overlap)
lemma union_with_inside: "S \<union> inside S = - outside S"
by (auto simp: inside_outside) (simp add: outside_inside)
lemma union_with_outside: "S \<union> outside S = - inside S"
by (simp add: inside_outside)
lemma outside_mono: "S \<subseteq> T \<Longrightarrow> outside T \<subseteq> outside S"
by (auto simp: outside bounded_subset connected_component_mono)
lemma inside_mono: "S \<subseteq> T \<Longrightarrow> inside S - T \<subseteq> inside T"
by (auto simp: inside_def bounded_subset connected_component_mono)
lemma segment_bound_lemma:
fixes u::real
assumes "x \<ge> B" "y \<ge> B" "0 \<le> u" "u \<le> 1"
shows "(1 - u) * x + u * y \<ge> B"
proof -
obtain dx dy where "dx \<ge> 0" "dy \<ge> 0" "x = B + dx" "y = B + dy"
using assms by auto (metis add.commute diff_add_cancel)
with \<open>0 \<le> u\<close> \<open>u \<le> 1\<close> show ?thesis
by (simp add: add_increasing2 mult_left_le field_simps)
qed
lemma cobounded_outside:
fixes S :: "'a :: real_normed_vector set"
assumes "bounded S" shows "bounded (- outside S)"
proof -
obtain B where B: "B>0" "S \<subseteq> ball 0 B"
using bounded_subset_ballD [OF assms, of 0] by auto
{ fix x::'a and C::real
assume Bno: "B \<le> norm x" and C: "0 < C"
have "\<exists>y. connected_component (- S) x y \<and> norm y > C"
proof (cases "x = 0")
case True with B Bno show ?thesis by force
next
case False
have "closed_segment x (((B + C) / norm x) *\<^sub>R x) \<subseteq> - ball 0 B"
proof
fix w
assume "w \<in> closed_segment x (((B + C) / norm x) *\<^sub>R x)"
then obtain u where
w: "w = (1 - u + u * (B + C) / norm x) *\<^sub>R x" "0 \<le> u" "u \<le> 1"
by (auto simp add: closed_segment_def real_vector_class.scaleR_add_left [symmetric])
with False B C have "B \<le> (1 - u) * norm x + u * (B + C)"
using segment_bound_lemma [of B "norm x" "B + C" u] Bno
by simp
with False B C show "w \<in> - ball 0 B"
using distrib_right [of _ _ "norm x"]
by (simp add: ball_def w not_less)
qed
also have "... \<subseteq> -S"
by (simp add: B)
finally have "\<exists>T. connected T \<and> T \<subseteq> - S \<and> x \<in> T \<and> ((B + C) / norm x) *\<^sub>R x \<in> T"
by (rule_tac x="closed_segment x (((B+C)/norm x) *\<^sub>R x)" in exI) simp
with False B
show ?thesis
by (rule_tac x="((B+C)/norm x) *\<^sub>R x" in exI) (simp add: connected_component_def)
qed
}
then show ?thesis
apply (simp add: outside_def assms)
apply (rule bounded_subset [OF bounded_ball [of 0 B]])
apply (force simp: dist_norm not_less bounded_pos)
done
qed
lemma unbounded_outside:
fixes S :: "'a::{real_normed_vector, perfect_space} set"
shows "bounded S \<Longrightarrow> \<not> bounded(outside S)"
using cobounded_imp_unbounded cobounded_outside by blast
lemma bounded_inside:
fixes S :: "'a::{real_normed_vector, perfect_space} set"
shows "bounded S \<Longrightarrow> bounded(inside S)"
by (simp add: bounded_Int cobounded_outside inside_outside)
lemma connected_outside:
fixes S :: "'a::euclidean_space set"
assumes "bounded S" "2 \<le> DIM('a)"
shows "connected(outside S)"
apply (clarsimp simp add: connected_iff_connected_component outside)
apply (rule_tac s="connected_component_set (- S) x" in connected_component_of_subset)
apply (metis (no_types) assms cobounded_unbounded_component cobounded_unique_unbounded_component connected_component_eq_eq connected_component_idemp double_complement mem_Collect_eq)
apply clarify
apply (metis connected_component_eq_eq connected_component_in)
done
lemma outside_connected_component_lt:
"outside S = {x. \<forall>B. \<exists>y. B < norm(y) \<and> connected_component (- S) x y}"
apply (auto simp: outside bounded_def dist_norm)
apply (metis diff_0 norm_minus_cancel not_less)
by (metis less_diff_eq norm_minus_commute norm_triangle_ineq2 order.trans pinf(6))
lemma outside_connected_component_le:
"outside S =
{x. \<forall>B. \<exists>y. B \<le> norm(y) \<and>
connected_component (- S) x y}"
apply (simp add: outside_connected_component_lt)
apply (simp add: Set.set_eq_iff)
by (meson gt_ex leD le_less_linear less_imp_le order.trans)
lemma not_outside_connected_component_lt:
fixes S :: "'a::euclidean_space set"
assumes S: "bounded S" and "2 \<le> DIM('a)"
shows "- (outside S) = {x. \<forall>B. \<exists>y. B < norm(y) \<and> \<not> connected_component (- S) x y}"
proof -
obtain B::real where B: "0 < B" and Bno: "\<And>x. x \<in> S \<Longrightarrow> norm x \<le> B"
using S [simplified bounded_pos] by auto
{ fix y::'a and z::'a
assume yz: "B < norm z" "B < norm y"
have "connected_component (- cball 0 B) y z"
apply (rule connected_componentI [OF _ subset_refl])
apply (rule connected_complement_bounded_convex)
using assms yz
by (auto simp: dist_norm)
then have "connected_component (- S) y z"
apply (rule connected_component_of_subset)
apply (metis Bno Compl_anti_mono mem_cball_0 subset_iff)
done
} note cyz = this
show ?thesis
apply (auto simp: outside)
apply (metis Compl_iff bounded_iff cobounded_imp_unbounded mem_Collect_eq not_le)
apply (simp add: bounded_pos)
by (metis B connected_component_trans cyz not_le)
qed
lemma not_outside_connected_component_le:
fixes S :: "'a::euclidean_space set"
assumes S: "bounded S" "2 \<le> DIM('a)"
shows "- (outside S) = {x. \<forall>B. \<exists>y. B \<le> norm(y) \<and> \<not> connected_component (- S) x y}"
apply (auto intro: less_imp_le simp: not_outside_connected_component_lt [OF assms])
by (meson gt_ex less_le_trans)
lemma inside_connected_component_lt:
fixes S :: "'a::euclidean_space set"
assumes S: "bounded S" "2 \<le> DIM('a)"
shows "inside S = {x. (x \<notin> S) \<and> (\<forall>B. \<exists>y. B < norm(y) \<and> \<not> connected_component (- S) x y)}"
by (auto simp: inside_outside not_outside_connected_component_lt [OF assms])
lemma inside_connected_component_le:
fixes S :: "'a::euclidean_space set"
assumes S: "bounded S" "2 \<le> DIM('a)"
shows "inside S = {x. (x \<notin> S) \<and> (\<forall>B. \<exists>y. B \<le> norm(y) \<and> \<not> connected_component (- S) x y)}"
by (auto simp: inside_outside not_outside_connected_component_le [OF assms])
lemma inside_subset:
assumes "connected U" and "\<not> bounded U" and "T \<union> U = - S"
shows "inside S \<subseteq> T"
apply (auto simp: inside_def)
by (metis bounded_subset [of "connected_component_set (- S) _"] connected_component_maximal
Compl_iff Un_iff assms subsetI)
lemma frontier_not_empty:
fixes S :: "'a :: real_normed_vector set"
shows "\<lbrakk>S \<noteq> {}; S \<noteq> UNIV\<rbrakk> \<Longrightarrow> frontier S \<noteq> {}"
using connected_Int_frontier [of UNIV S] by auto
lemma frontier_eq_empty:
fixes S :: "'a :: real_normed_vector set"
shows "frontier S = {} \<longleftrightarrow> S = {} \<or> S = UNIV"
using frontier_UNIV frontier_empty frontier_not_empty by blast
lemma frontier_of_connected_component_subset:
fixes S :: "'a::real_normed_vector set"
shows "frontier(connected_component_set S x) \<subseteq> frontier S"
proof -
{ fix y
assume y1: "y \<in> closure (connected_component_set S x)"
and y2: "y \<notin> interior (connected_component_set S x)"
have "y \<in> closure S"
using y1 closure_mono connected_component_subset by blast
moreover have "z \<in> interior (connected_component_set S x)"
if "0 < e" "ball y e \<subseteq> interior S" "dist y z < e" for e z
proof -
have "ball y e \<subseteq> connected_component_set S y"
apply (rule connected_component_maximal)
using that interior_subset mem_ball apply auto
done
then show ?thesis
using y1 apply (simp add: closure_approachable open_contains_ball_eq [OF open_interior])
by (metis connected_component_eq dist_commute mem_Collect_eq mem_ball mem_interior subsetD \<open>0 < e\<close> y2)
qed
then have "y \<notin> interior S"
using y2 by (force simp: open_contains_ball_eq [OF open_interior])
ultimately have "y \<in> frontier S"
by (auto simp: frontier_def)
}
then show ?thesis by (auto simp: frontier_def)
qed
lemma frontier_Union_subset_closure:
fixes F :: "'a::real_normed_vector set set"
shows "frontier(\<Union>F) \<subseteq> closure(\<Union>t \<in> F. frontier t)"
proof -
have "\<exists>y\<in>F. \<exists>y\<in>frontier y. dist y x < e"
if "T \<in> F" "y \<in> T" "dist y x < e"
"x \<notin> interior (\<Union>F)" "0 < e" for x y e T
proof (cases "x \<in> T")
case True with that show ?thesis
by (metis Diff_iff Sup_upper closure_subset contra_subsetD dist_self frontier_def interior_mono)
next
case False
have 1: "closed_segment x y \<inter> T \<noteq> {}" using \<open>y \<in> T\<close> by blast
have 2: "closed_segment x y - T \<noteq> {}"
using False by blast
obtain c where "c \<in> closed_segment x y" "c \<in> frontier T"
using False connected_Int_frontier [OF connected_segment 1 2] by auto
then show ?thesis
proof -
have "norm (y - x) < e"
by (metis dist_norm \<open>dist y x < e\<close>)
moreover have "norm (c - x) \<le> norm (y - x)"
by (simp add: \<open>c \<in> closed_segment x y\<close> segment_bound(1))
ultimately have "norm (c - x) < e"
by linarith
then show ?thesis
by (metis (no_types) \<open>c \<in> frontier T\<close> dist_norm that(1))
qed
qed
then show ?thesis
by (fastforce simp add: frontier_def closure_approachable)
qed
lemma frontier_Union_subset:
fixes F :: "'a::real_normed_vector set set"
shows "finite F \<Longrightarrow> frontier(\<Union>F) \<subseteq> (\<Union>t \<in> F. frontier t)"
by (rule order_trans [OF frontier_Union_subset_closure])
(auto simp: closure_subset_eq)
lemma frontier_of_components_subset:
fixes S :: "'a::real_normed_vector set"
shows "C \<in> components S \<Longrightarrow> frontier C \<subseteq> frontier S"
by (metis Path_Connected.frontier_of_connected_component_subset components_iff)
lemma frontier_of_components_closed_complement:
fixes S :: "'a::real_normed_vector set"
shows "\<lbrakk>closed S; C \<in> components (- S)\<rbrakk> \<Longrightarrow> frontier C \<subseteq> S"
using frontier_complement frontier_of_components_subset frontier_subset_eq by blast
lemma frontier_minimal_separating_closed:
fixes S :: "'a::real_normed_vector set"
assumes "closed S"
and nconn: "\<not> connected(- S)"
and C: "C \<in> components (- S)"
and conn: "\<And>T. \<lbrakk>closed T; T \<subset> S\<rbrakk> \<Longrightarrow> connected(- T)"
shows "frontier C = S"
proof (rule ccontr)
assume "frontier C \<noteq> S"
then have "frontier C \<subset> S"
using frontier_of_components_closed_complement [OF \<open>closed S\<close> C] by blast
then have "connected(- (frontier C))"
by (simp add: conn)
have "\<not> connected(- (frontier C))"
unfolding connected_def not_not
proof (intro exI conjI)
show "open C"
using C \<open>closed S\<close> open_components by blast
show "open (- closure C)"
by blast
show "C \<inter> - closure C \<inter> - frontier C = {}"
using closure_subset by blast
show "C \<inter> - frontier C \<noteq> {}"
using C \<open>open C\<close> components_eq frontier_disjoint_eq by fastforce
show "- frontier C \<subseteq> C \<union> - closure C"
by (simp add: \<open>open C\<close> closed_Compl frontier_closures)
then show "- closure C \<inter> - frontier C \<noteq> {}"
by (metis (no_types, lifting) C Compl_subset_Compl_iff \<open>frontier C \<subset> S\<close> compl_sup frontier_closures in_components_subset psubsetE sup.absorb_iff2 sup.boundedE sup_bot.right_neutral sup_inf_absorb)
qed
then show False
using \<open>connected (- frontier C)\<close> by blast
qed
lemma connected_component_UNIV [simp]:
fixes x :: "'a::real_normed_vector"
shows "connected_component_set UNIV x = UNIV"
using connected_iff_eq_connected_component_set [of "UNIV::'a set"] connected_UNIV
by auto
lemma connected_component_eq_UNIV:
fixes x :: "'a::real_normed_vector"
shows "connected_component_set s x = UNIV \<longleftrightarrow> s = UNIV"
using connected_component_in connected_component_UNIV by blast
lemma components_UNIV [simp]: "components UNIV = {UNIV :: 'a::real_normed_vector set}"
by (auto simp: components_eq_sing_iff)
lemma interior_inside_frontier:
fixes s :: "'a::real_normed_vector set"
assumes "bounded s"
shows "interior s \<subseteq> inside (frontier s)"
proof -
{ fix x y
assume x: "x \<in> interior s" and y: "y \<notin> s"
and cc: "connected_component (- frontier s) x y"
have "connected_component_set (- frontier s) x \<inter> frontier s \<noteq> {}"
apply (rule connected_Int_frontier, simp)
apply (metis IntI cc connected_component_in connected_component_refl empty_iff interiorE mem_Collect_eq rev_subsetD x)
using y cc
by blast
then have "bounded (connected_component_set (- frontier s) x)"
using connected_component_in by auto
}
then show ?thesis
apply (auto simp: inside_def frontier_def)
apply (rule classical)
apply (rule bounded_subset [OF assms], blast)
done
qed
lemma inside_empty [simp]: "inside {} = ({} :: 'a :: {real_normed_vector, perfect_space} set)"
by (simp add: inside_def)
lemma outside_empty [simp]: "outside {} = (UNIV :: 'a :: {real_normed_vector, perfect_space} set)"
using inside_empty inside_Un_outside by blast
lemma inside_same_component:
"\<lbrakk>connected_component (- s) x y; x \<in> inside s\<rbrakk> \<Longrightarrow> y \<in> inside s"
using connected_component_eq connected_component_in
by (fastforce simp add: inside_def)
lemma outside_same_component:
"\<lbrakk>connected_component (- s) x y; x \<in> outside s\<rbrakk> \<Longrightarrow> y \<in> outside s"
using connected_component_eq connected_component_in
by (fastforce simp add: outside_def)
lemma convex_in_outside:
fixes s :: "'a :: {real_normed_vector, perfect_space} set"
assumes s: "convex s" and z: "z \<notin> s"
shows "z \<in> outside s"
proof (cases "s={}")
case True then show ?thesis by simp
next
case False then obtain a where "a \<in> s" by blast
with z have zna: "z \<noteq> a" by auto
{ assume "bounded (connected_component_set (- s) z)"
with bounded_pos_less obtain B where "B>0" and B: "\<And>x. connected_component (- s) z x \<Longrightarrow> norm x < B"
by (metis mem_Collect_eq)
define C where "C = (B + 1 + norm z) / norm (z-a)"
have "C > 0"
using \<open>0 < B\<close> zna by (simp add: C_def field_split_simps add_strict_increasing)
have "\<bar>norm (z + C *\<^sub>R (z-a)) - norm (C *\<^sub>R (z-a))\<bar> \<le> norm z"
by (metis add_diff_cancel norm_triangle_ineq3)
moreover have "norm (C *\<^sub>R (z-a)) > norm z + B"
using zna \<open>B>0\<close> by (simp add: C_def le_max_iff_disj)
ultimately have C: "norm (z + C *\<^sub>R (z-a)) > B" by linarith
{ fix u::real
assume u: "0\<le>u" "u\<le>1" and ins: "(1 - u) *\<^sub>R z + u *\<^sub>R (z + C *\<^sub>R (z - a)) \<in> s"
then have Cpos: "1 + u * C > 0"
by (meson \<open>0 < C\<close> add_pos_nonneg less_eq_real_def zero_le_mult_iff zero_less_one)
then have *: "(1 / (1 + u * C)) *\<^sub>R z + (u * C / (1 + u * C)) *\<^sub>R z = z"
by (simp add: scaleR_add_left [symmetric] field_split_simps)
then have False
using convexD_alt [OF s \<open>a \<in> s\<close> ins, of "1/(u*C + 1)"] \<open>C>0\<close> \<open>z \<notin> s\<close> Cpos u
by (simp add: * field_split_simps)
} note contra = this
have "connected_component (- s) z (z + C *\<^sub>R (z-a))"
apply (rule connected_componentI [OF connected_segment [of z "z + C *\<^sub>R (z-a)"]])
apply (simp add: closed_segment_def)
using contra
apply auto
done
then have False
using zna B [of "z + C *\<^sub>R (z-a)"] C
by (auto simp: field_split_simps max_mult_distrib_right)
}
then show ?thesis
by (auto simp: outside_def z)
qed
lemma outside_convex:
fixes s :: "'a :: {real_normed_vector, perfect_space} set"
assumes "convex s"
shows "outside s = - s"
by (metis ComplD assms convex_in_outside equalityI inside_Un_outside subsetI sup.cobounded2)
lemma outside_singleton [simp]:
fixes x :: "'a :: {real_normed_vector, perfect_space}"
shows "outside {x} = -{x}"
by (auto simp: outside_convex)
lemma inside_convex:
fixes s :: "'a :: {real_normed_vector, perfect_space} set"
shows "convex s \<Longrightarrow> inside s = {}"
by (simp add: inside_outside outside_convex)
lemma inside_singleton [simp]:
fixes x :: "'a :: {real_normed_vector, perfect_space}"
shows "inside {x} = {}"
by (auto simp: inside_convex)
lemma outside_subset_convex:
fixes s :: "'a :: {real_normed_vector, perfect_space} set"
shows "\<lbrakk>convex t; s \<subseteq> t\<rbrakk> \<Longrightarrow> - t \<subseteq> outside s"
using outside_convex outside_mono by blast
lemma outside_Un_outside_Un:
fixes S :: "'a::real_normed_vector set"
assumes "S \<inter> outside(T \<union> U) = {}"
shows "outside(T \<union> U) \<subseteq> outside(T \<union> S)"
proof
fix x
assume x: "x \<in> outside (T \<union> U)"
have "Y \<subseteq> - S" if "connected Y" "Y \<subseteq> - T" "Y \<subseteq> - U" "x \<in> Y" "u \<in> Y" for u Y
proof -
have "Y \<subseteq> connected_component_set (- (T \<union> U)) x"
by (simp add: connected_component_maximal that)
also have "\<dots> \<subseteq> outside(T \<union> U)"
by (metis (mono_tags, lifting) Collect_mono mem_Collect_eq outside outside_same_component x)
finally have "Y \<subseteq> outside(T \<union> U)" .
with assms show ?thesis by auto
qed
with x show "x \<in> outside (T \<union> S)"
by (simp add: outside_connected_component_lt connected_component_def) meson
qed
lemma outside_frontier_misses_closure:
fixes s :: "'a::real_normed_vector set"
assumes "bounded s"
shows "outside(frontier s) \<subseteq> - closure s"
unfolding outside_inside Lattices.boolean_algebra_class.compl_le_compl_iff
proof -
{ assume "interior s \<subseteq> inside (frontier s)"
hence "interior s \<union> inside (frontier s) = inside (frontier s)"
by (simp add: subset_Un_eq)
then have "closure s \<subseteq> frontier s \<union> inside (frontier s)"
using frontier_def by auto
}
then show "closure s \<subseteq> frontier s \<union> inside (frontier s)"
using interior_inside_frontier [OF assms] by blast
qed
lemma outside_frontier_eq_complement_closure:
fixes s :: "'a :: {real_normed_vector, perfect_space} set"
assumes "bounded s" "convex s"
shows "outside(frontier s) = - closure s"
by (metis Diff_subset assms convex_closure frontier_def outside_frontier_misses_closure
outside_subset_convex subset_antisym)
lemma inside_frontier_eq_interior:
fixes s :: "'a :: {real_normed_vector, perfect_space} set"
shows "\<lbrakk>bounded s; convex s\<rbrakk> \<Longrightarrow> inside(frontier s) = interior s"
apply (simp add: inside_outside outside_frontier_eq_complement_closure)
using closure_subset interior_subset
apply (auto simp: frontier_def)
done
lemma open_inside:
fixes s :: "'a::real_normed_vector set"
assumes "closed s"
shows "open (inside s)"
proof -
{ fix x assume x: "x \<in> inside s"
have "open (connected_component_set (- s) x)"
using assms open_connected_component by blast
then obtain e where e: "e>0" and e: "\<And>y. dist y x < e \<longrightarrow> connected_component (- s) x y"
using dist_not_less_zero
apply (simp add: open_dist)
by (metis (no_types, lifting) Compl_iff connected_component_refl_eq inside_def mem_Collect_eq x)
then have "\<exists>e>0. ball x e \<subseteq> inside s"
by (metis e dist_commute inside_same_component mem_ball subsetI x)
}
then show ?thesis
by (simp add: open_contains_ball)
qed
lemma open_outside:
fixes s :: "'a::real_normed_vector set"
assumes "closed s"
shows "open (outside s)"
proof -
{ fix x assume x: "x \<in> outside s"
have "open (connected_component_set (- s) x)"
using assms open_connected_component by blast
then obtain e where e: "e>0" and e: "\<And>y. dist y x < e \<longrightarrow> connected_component (- s) x y"
using dist_not_less_zero
apply (simp add: open_dist)
by (metis Int_iff outside_def connected_component_refl_eq x)
then have "\<exists>e>0. ball x e \<subseteq> outside s"
by (metis e dist_commute outside_same_component mem_ball subsetI x)
}
then show ?thesis
by (simp add: open_contains_ball)
qed
lemma closure_inside_subset:
fixes s :: "'a::real_normed_vector set"
assumes "closed s"
shows "closure(inside s) \<subseteq> s \<union> inside s"
by (metis assms closure_minimal open_closed open_outside sup.cobounded2 union_with_inside)
lemma frontier_inside_subset:
fixes s :: "'a::real_normed_vector set"
assumes "closed s"
shows "frontier(inside s) \<subseteq> s"
proof -
have "closure (inside s) \<inter> - inside s = closure (inside s) - interior (inside s)"
by (metis (no_types) Diff_Compl assms closure_closed interior_closure open_closed open_inside)
moreover have "- inside s \<inter> - outside s = s"
by (metis (no_types) compl_sup double_compl inside_Un_outside)
moreover have "closure (inside s) \<subseteq> - outside s"
by (metis (no_types) assms closure_inside_subset union_with_inside)
ultimately have "closure (inside s) - interior (inside s) \<subseteq> s"
by blast
then show ?thesis
by (simp add: frontier_def open_inside interior_open)
qed
lemma closure_outside_subset:
fixes s :: "'a::real_normed_vector set"
assumes "closed s"
shows "closure(outside s) \<subseteq> s \<union> outside s"
apply (rule closure_minimal, simp)
by (metis assms closed_open inside_outside open_inside)
lemma frontier_outside_subset:
fixes s :: "'a::real_normed_vector set"
assumes "closed s"
shows "frontier(outside s) \<subseteq> s"
apply (simp add: frontier_def open_outside interior_open)
by (metis Diff_subset_conv assms closure_outside_subset interior_eq open_outside sup.commute)
lemma inside_complement_unbounded_connected_empty:
"\<lbrakk>connected (- s); \<not> bounded (- s)\<rbrakk> \<Longrightarrow> inside s = {}"
apply (simp add: inside_def)
by (meson Compl_iff bounded_subset connected_component_maximal order_refl)
lemma inside_bounded_complement_connected_empty:
fixes s :: "'a::{real_normed_vector, perfect_space} set"
shows "\<lbrakk>connected (- s); bounded s\<rbrakk> \<Longrightarrow> inside s = {}"
by (metis inside_complement_unbounded_connected_empty cobounded_imp_unbounded)
lemma inside_inside:
assumes "s \<subseteq> inside t"
shows "inside s - t \<subseteq> inside t"
unfolding inside_def
proof clarify
fix x
assume x: "x \<notin> t" "x \<notin> s" and bo: "bounded (connected_component_set (- s) x)"
show "bounded (connected_component_set (- t) x)"
proof (cases "s \<inter> connected_component_set (- t) x = {}")
case True show ?thesis
apply (rule bounded_subset [OF bo])
apply (rule connected_component_maximal)
using x True apply auto
done
next
case False then show ?thesis
using assms [unfolded inside_def] x
apply (simp add: disjoint_iff_not_equal, clarify)
apply (drule subsetD, assumption, auto)
by (metis (no_types, hide_lams) ComplI connected_component_eq_eq)
qed
qed
lemma inside_inside_subset: "inside(inside s) \<subseteq> s"
using inside_inside union_with_outside by fastforce
lemma inside_outside_intersect_connected:
"\<lbrakk>connected t; inside s \<inter> t \<noteq> {}; outside s \<inter> t \<noteq> {}\<rbrakk> \<Longrightarrow> s \<inter> t \<noteq> {}"
apply (simp add: inside_def outside_def ex_in_conv [symmetric] disjoint_eq_subset_Compl, clarify)
by (metis (no_types, hide_lams) Compl_anti_mono connected_component_eq connected_component_maximal contra_subsetD double_compl)
lemma outside_bounded_nonempty:
fixes s :: "'a :: {real_normed_vector, perfect_space} set"
assumes "bounded s" shows "outside s \<noteq> {}"
by (metis (no_types, lifting) Collect_empty_eq Collect_mem_eq Compl_eq_Diff_UNIV Diff_cancel
Diff_disjoint UNIV_I assms ball_eq_empty bounded_diff cobounded_outside convex_ball
double_complement order_refl outside_convex outside_def)
lemma outside_compact_in_open:
fixes s :: "'a :: {real_normed_vector,perfect_space} set"
assumes s: "compact s" and t: "open t" and "s \<subseteq> t" "t \<noteq> {}"
shows "outside s \<inter> t \<noteq> {}"
proof -
have "outside s \<noteq> {}"
by (simp add: compact_imp_bounded outside_bounded_nonempty s)
with assms obtain a b where a: "a \<in> outside s" and b: "b \<in> t" by auto
show ?thesis
proof (cases "a \<in> t")
case True with a show ?thesis by blast
next
case False
have front: "frontier t \<subseteq> - s"
using \<open>s \<subseteq> t\<close> frontier_disjoint_eq t by auto
{ fix \<gamma>
assume "path \<gamma>" and pimg_sbs: "path_image \<gamma> - {pathfinish \<gamma>} \<subseteq> interior (- t)"
and pf: "pathfinish \<gamma> \<in> frontier t" and ps: "pathstart \<gamma> = a"
define c where "c = pathfinish \<gamma>"
have "c \<in> -s" unfolding c_def using front pf by blast
moreover have "open (-s)" using s compact_imp_closed by blast
ultimately obtain \<epsilon>::real where "\<epsilon> > 0" and \<epsilon>: "cball c \<epsilon> \<subseteq> -s"
using open_contains_cball[of "-s"] s by blast
then obtain d where "d \<in> t" and d: "dist d c < \<epsilon>"
using closure_approachable [of c t] pf unfolding c_def
by (metis Diff_iff frontier_def)
then have "d \<in> -s" using \<epsilon>
using dist_commute by (metis contra_subsetD mem_cball not_le not_less_iff_gr_or_eq)
have pimg_sbs_cos: "path_image \<gamma> \<subseteq> -s"
using pimg_sbs apply (auto simp: path_image_def)
apply (drule subsetD)
using \<open>c \<in> - s\<close> \<open>s \<subseteq> t\<close> interior_subset apply (auto simp: c_def)
done
have "closed_segment c d \<le> cball c \<epsilon>"
apply (simp add: segment_convex_hull)
apply (rule hull_minimal)
using \<open>\<epsilon> > 0\<close> d apply (auto simp: dist_commute)
done
with \<epsilon> have "closed_segment c d \<subseteq> -s" by blast
moreover have con_gcd: "connected (path_image \<gamma> \<union> closed_segment c d)"
by (rule connected_Un) (auto simp: c_def \<open>path \<gamma>\<close> connected_path_image)
ultimately have "connected_component (- s) a d"
unfolding connected_component_def using pimg_sbs_cos ps by blast
then have "outside s \<inter> t \<noteq> {}"
using outside_same_component [OF _ a] by (metis IntI \<open>d \<in> t\<close> empty_iff)
} note * = this
have pal: "pathstart (linepath a b) \<in> closure (- t)"
by (auto simp: False closure_def)
show ?thesis
by (rule exists_path_subpath_to_frontier [OF path_linepath pal _ *]) (auto simp: b)
qed
qed
lemma inside_inside_compact_connected:
fixes s :: "'a :: euclidean_space set"
assumes s: "closed s" and t: "compact t" and "connected t" "s \<subseteq> inside t"
shows "inside s \<subseteq> inside t"
proof (cases "inside t = {}")
case True with assms show ?thesis by auto
next
case False
consider "DIM('a) = 1" | "DIM('a) \<ge> 2"
using antisym not_less_eq_eq by fastforce
then show ?thesis
proof cases
case 1 then show ?thesis
using connected_convex_1_gen assms False inside_convex by blast
next
case 2
have coms: "compact s"
using assms apply (simp add: s compact_eq_bounded_closed)
by (meson bounded_inside bounded_subset compact_imp_bounded)
then have bst: "bounded (s \<union> t)"
by (simp add: compact_imp_bounded t)
then obtain r where "0 < r" and r: "s \<union> t \<subseteq> ball 0 r"
using bounded_subset_ballD by blast
have outst: "outside s \<inter> outside t \<noteq> {}"
proof -
have "- ball 0 r \<subseteq> outside s"
apply (rule outside_subset_convex)
using r by auto
moreover have "- ball 0 r \<subseteq> outside t"
apply (rule outside_subset_convex)
using r by auto
ultimately show ?thesis
by (metis Compl_subset_Compl_iff Int_subset_iff bounded_ball inf.orderE outside_bounded_nonempty outside_no_overlap)
qed
have "s \<inter> t = {}" using assms
by (metis disjoint_iff_not_equal inside_no_overlap subsetCE)
moreover have "outside s \<inter> inside t \<noteq> {}"
by (meson False assms(4) compact_eq_bounded_closed coms open_inside outside_compact_in_open t)
ultimately have "inside s \<inter> t = {}"
using inside_outside_intersect_connected [OF \<open>connected t\<close>, of s]
by (metis "2" compact_eq_bounded_closed coms connected_outside inf.commute inside_outside_intersect_connected outst)
then show ?thesis
using inside_inside [OF \<open>s \<subseteq> inside t\<close>] by blast
qed
qed
lemma connected_with_inside:
fixes s :: "'a :: real_normed_vector set"
assumes s: "closed s" and cons: "connected s"
shows "connected(s \<union> inside s)"
proof (cases "s \<union> inside s = UNIV")
case True with assms show ?thesis by auto
next
case False
then obtain b where b: "b \<notin> s" "b \<notin> inside s" by blast
have *: "\<exists>y t. y \<in> s \<and> connected t \<and> a \<in> t \<and> y \<in> t \<and> t \<subseteq> (s \<union> inside s)" if "a \<in> (s \<union> inside s)" for a
using that proof
assume "a \<in> s" then show ?thesis
apply (rule_tac x=a in exI)
apply (rule_tac x="{a}" in exI, simp)
done
next
assume a: "a \<in> inside s"
show ?thesis
apply (rule exists_path_subpath_to_frontier [OF path_linepath [of a b], of "inside s"])
using a apply (simp add: closure_def)
apply (simp add: b)
apply (rule_tac x="pathfinish h" in exI)
apply (rule_tac x="path_image h" in exI)
apply (simp add: pathfinish_in_path_image connected_path_image, auto)
using frontier_inside_subset s apply fastforce
by (metis (no_types, lifting) frontier_inside_subset insertE insert_Diff interior_eq open_inside pathfinish_in_path_image s subsetCE)
qed
show ?thesis
apply (simp add: connected_iff_connected_component)
apply (simp add: connected_component_def)
apply (clarify dest!: *)
apply (rename_tac u u' t t')
apply (rule_tac x="(s \<union> t \<union> t')" in exI)
apply (auto simp: intro!: connected_Un cons)
done
qed
text\<open>The proof is virtually the same as that above.\<close>
lemma connected_with_outside:
fixes s :: "'a :: real_normed_vector set"
assumes s: "closed s" and cons: "connected s"
shows "connected(s \<union> outside s)"
proof (cases "s \<union> outside s = UNIV")
case True with assms show ?thesis by auto
next
case False
then obtain b where b: "b \<notin> s" "b \<notin> outside s" by blast
have *: "\<exists>y t. y \<in> s \<and> connected t \<and> a \<in> t \<and> y \<in> t \<and> t \<subseteq> (s \<union> outside s)" if "a \<in> (s \<union> outside s)" for a
using that proof
assume "a \<in> s" then show ?thesis
apply (rule_tac x=a in exI)
apply (rule_tac x="{a}" in exI, simp)
done
next
assume a: "a \<in> outside s"
show ?thesis
apply (rule exists_path_subpath_to_frontier [OF path_linepath [of a b], of "outside s"])
using a apply (simp add: closure_def)
apply (simp add: b)
apply (rule_tac x="pathfinish h" in exI)
apply (rule_tac x="path_image h" in exI)
apply (simp add: pathfinish_in_path_image connected_path_image, auto)
using frontier_outside_subset s apply fastforce
by (metis (no_types, lifting) frontier_outside_subset insertE insert_Diff interior_eq open_outside pathfinish_in_path_image s subsetCE)
qed
show ?thesis
apply (simp add: connected_iff_connected_component)
apply (simp add: connected_component_def)
apply (clarify dest!: *)
apply (rename_tac u u' t t')
apply (rule_tac x="(s \<union> t \<union> t')" in exI)
apply (auto simp: intro!: connected_Un cons)
done
qed
lemma inside_inside_eq_empty [simp]:
fixes s :: "'a :: {real_normed_vector, perfect_space} set"
assumes s: "closed s" and cons: "connected s"
shows "inside (inside s) = {}"
by (metis (no_types) unbounded_outside connected_with_outside [OF assms] bounded_Un
inside_complement_unbounded_connected_empty unbounded_outside union_with_outside)
lemma inside_in_components:
"inside s \<in> components (- s) \<longleftrightarrow> connected(inside s) \<and> inside s \<noteq> {}"
apply (simp add: in_components_maximal)
apply (auto intro: inside_same_component connected_componentI)
apply (metis IntI empty_iff inside_no_overlap)
done
text\<open>The proof is virtually the same as that above.\<close>
lemma outside_in_components:
"outside s \<in> components (- s) \<longleftrightarrow> connected(outside s) \<and> outside s \<noteq> {}"
apply (simp add: in_components_maximal)
apply (auto intro: outside_same_component connected_componentI)
apply (metis IntI empty_iff outside_no_overlap)
done
lemma bounded_unique_outside:
fixes s :: "'a :: euclidean_space set"
shows "\<lbrakk>bounded s; DIM('a) \<ge> 2\<rbrakk> \<Longrightarrow> (c \<in> components (- s) \<and> \<not> bounded c \<longleftrightarrow> c = outside s)"
apply (rule iffI)
apply (metis cobounded_unique_unbounded_components connected_outside double_compl outside_bounded_nonempty outside_in_components unbounded_outside)
by (simp add: connected_outside outside_bounded_nonempty outside_in_components unbounded_outside)
subsection\<open>Condition for an open map's image to contain a ball\<close>
proposition ball_subset_open_map_image:
fixes f :: "'a::heine_borel \<Rightarrow> 'b :: {real_normed_vector,heine_borel}"
assumes contf: "continuous_on (closure S) f"
and oint: "open (f ` interior S)"
and le_no: "\<And>z. z \<in> frontier S \<Longrightarrow> r \<le> norm(f z - f a)"
and "bounded S" "a \<in> S" "0 < r"
shows "ball (f a) r \<subseteq> f ` S"
proof (cases "f ` S = UNIV")
case True then show ?thesis by simp
next
case False
obtain w where w: "w \<in> frontier (f ` S)"
and dw_le: "\<And>y. y \<in> frontier (f ` S) \<Longrightarrow> norm (f a - w) \<le> norm (f a - y)"
apply (rule distance_attains_inf [of "frontier(f ` S)" "f a"])
using \<open>a \<in> S\<close> by (auto simp: frontier_eq_empty dist_norm False)
then obtain \<xi> where \<xi>: "\<And>n. \<xi> n \<in> f ` S" and tendsw: "\<xi> \<longlonglongrightarrow> w"
by (metis Diff_iff frontier_def closure_sequential)
then have "\<And>n. \<exists>x \<in> S. \<xi> n = f x" by force
then obtain z where zs: "\<And>n. z n \<in> S" and fz: "\<And>n. \<xi> n = f (z n)"
by metis
then obtain y K where y: "y \<in> closure S" and "strict_mono (K :: nat \<Rightarrow> nat)"
and Klim: "(z \<circ> K) \<longlonglongrightarrow> y"
using \<open>bounded S\<close>
apply (simp add: compact_closure [symmetric] compact_def)
apply (drule_tac x=z in spec)
using closure_subset apply force
done
then have ftendsw: "((\<lambda>n. f (z n)) \<circ> K) \<longlonglongrightarrow> w"
by (metis LIMSEQ_subseq_LIMSEQ fun.map_cong0 fz tendsw)
have zKs: "\<And>n. (z \<circ> K) n \<in> S" by (simp add: zs)
have fz: "f \<circ> z = \<xi>" "(\<lambda>n. f (z n)) = \<xi>"
using fz by auto
then have "(\<xi> \<circ> K) \<longlonglongrightarrow> f y"
by (metis (no_types) Klim zKs y contf comp_assoc continuous_on_closure_sequentially)
with fz have wy: "w = f y" using fz LIMSEQ_unique ftendsw by auto
have rle: "r \<le> norm (f y - f a)"
apply (rule le_no)
using w wy oint
by (force simp: imageI image_mono interiorI interior_subset frontier_def y)
have **: "(b \<inter> (- S) \<noteq> {} \<and> b - (- S) \<noteq> {} \<Longrightarrow> b \<inter> f \<noteq> {})
\<Longrightarrow> (b \<inter> S \<noteq> {}) \<Longrightarrow> b \<inter> f = {} \<Longrightarrow>
b \<subseteq> S" for b f and S :: "'b set"
by blast
show ?thesis
apply (rule **) (*such a horrible mess*)
apply (rule connected_Int_frontier [where t = "f`S", OF connected_ball])
using \<open>a \<in> S\<close> \<open>0 < r\<close>
apply (auto simp: disjoint_iff_not_equal dist_norm)
by (metis dw_le norm_minus_commute not_less order_trans rle wy)
qed
subsubsection\<open>Special characterizations of classes of functions into and out of R.\<close>
lemma Hausdorff_space_euclidean [simp]: "Hausdorff_space (euclidean :: 'a::metric_space topology)"
proof -
have "\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> disjnt U V"
if "x \<noteq> y"
for x y :: 'a
proof (intro exI conjI)
let ?r = "dist x y / 2"
have [simp]: "?r > 0"
by (simp add: that)
show "open (ball x ?r)" "open (ball y ?r)" "x \<in> (ball x ?r)" "y \<in> (ball y ?r)"
by (auto simp add: that)
show "disjnt (ball x ?r) (ball y ?r)"
unfolding disjnt_def by (simp add: disjoint_ballI)
qed
then show ?thesis
by (simp add: Hausdorff_space_def)
qed
proposition embedding_map_into_euclideanreal:
assumes "path_connected_space X"
shows "embedding_map X euclideanreal f \<longleftrightarrow>
continuous_map X euclideanreal f \<and> inj_on f (topspace X)"
proof safe
show "continuous_map X euclideanreal f"
if "embedding_map X euclideanreal f"
using continuous_map_in_subtopology homeomorphic_imp_continuous_map that
unfolding embedding_map_def by blast
show "inj_on f (topspace X)"
if "embedding_map X euclideanreal f"
using that homeomorphic_imp_injective_map
unfolding embedding_map_def by blast
show "embedding_map X euclideanreal f"
if cont: "continuous_map X euclideanreal f" and inj: "inj_on f (topspace X)"
proof -
obtain g where gf: "\<And>x. x \<in> topspace X \<Longrightarrow> g (f x) = x"
using inv_into_f_f [OF inj] by auto
show ?thesis
unfolding embedding_map_def homeomorphic_map_maps homeomorphic_maps_def
proof (intro exI conjI)
show "continuous_map X (top_of_set (f ` topspace X)) f"
by (simp add: cont continuous_map_in_subtopology)
let ?S = "f ` topspace X"
have eq: "{x \<in> ?S. g x \<in> U} = f ` U" if "openin X U" for U
using openin_subset [OF that] by (auto simp: gf)
have 1: "g ` ?S \<subseteq> topspace X"
using eq by blast
have "openin (top_of_set ?S) {x \<in> ?S. g x \<in> T}"
if "openin X T" for T
proof -
have "T \<subseteq> topspace X"
by (simp add: openin_subset that)
have RR: "\<forall>x \<in> ?S \<inter> g -` T. \<exists>d>0. \<forall>x' \<in> ?S \<inter> ball x d. g x' \<in> T"
proof (clarsimp simp add: gf)
have pcS: "path_connectedin euclidean ?S"
using assms cont path_connectedin_continuous_map_image path_connectedin_topspace by blast
show "\<exists>d>0. \<forall>x'\<in>f ` topspace X \<inter> ball (f x) d. g x' \<in> T"
if "x \<in> T" for x
proof -
have x: "x \<in> topspace X"
using \<open>T \<subseteq> topspace X\<close> \<open>x \<in> T\<close> by blast
obtain u v d where "0 < d" "u \<in> topspace X" "v \<in> topspace X"
and sub_fuv: "?S \<inter> {f x - d .. f x + d} \<subseteq> {f u..f v}"
proof (cases "\<exists>u \<in> topspace X. f u < f x")
case True
then obtain u where u: "u \<in> topspace X" "f u < f x" ..
show ?thesis
proof (cases "\<exists>v \<in> topspace X. f x < f v")
case True
then obtain v where v: "v \<in> topspace X" "f x < f v" ..
show ?thesis
proof
let ?d = "min (f x - f u) (f v - f x)"
show "0 < ?d"
by (simp add: \<open>f u < f x\<close> \<open>f x < f v\<close>)
show "f ` topspace X \<inter> {f x - ?d..f x + ?d} \<subseteq> {f u..f v}"
by fastforce
qed (auto simp: u v)
next
case False
show ?thesis
proof
let ?d = "f x - f u"
show "0 < ?d"
by (simp add: u)
show "f ` topspace X \<inter> {f x - ?d..f x + ?d} \<subseteq> {f u..f x}"
using x u False by auto
qed (auto simp: x u)
qed
next
case False
note no_u = False
show ?thesis
proof (cases "\<exists>v \<in> topspace X. f x < f v")
case True
then obtain v where v: "v \<in> topspace X" "f x < f v" ..
show ?thesis
proof
let ?d = "f v - f x"
show "0 < ?d"
by (simp add: v)
show "f ` topspace X \<inter> {f x - ?d..f x + ?d} \<subseteq> {f x..f v}"
using False by auto
qed (auto simp: x v)
next
case False
show ?thesis
proof
show "f ` topspace X \<inter> {f x - 1..f x + 1} \<subseteq> {f x..f x}"
using False no_u by fastforce
qed (auto simp: x)
qed
qed
then obtain h where "pathin X h" "h 0 = u" "h 1 = v"
using assms unfolding path_connected_space_def by blast
obtain C where "compactin X C" "connectedin X C" "u \<in> C" "v \<in> C"
proof
show "compactin X (h ` {0..1})"
using that by (simp add: \<open>pathin X h\<close> compactin_path_image)
show "connectedin X (h ` {0..1})"
using \<open>pathin X h\<close> connectedin_path_image by blast
qed (use \<open>h 0 = u\<close> \<open>h 1 = v\<close> in auto)
have "continuous_map (subtopology euclideanreal (?S \<inter> {f x - d .. f x + d})) (subtopology X C) g"
proof (rule continuous_inverse_map)
show "compact_space (subtopology X C)"
using \<open>compactin X C\<close> compactin_subspace by blast
show "continuous_map (subtopology X C) euclideanreal f"
by (simp add: cont continuous_map_from_subtopology)
have "{f u .. f v} \<subseteq> f ` topspace (subtopology X C)"
proof (rule connected_contains_Icc)
show "connected (f ` topspace (subtopology X C))"
using connectedin_continuous_map_image [OF cont]
by (simp add: \<open>compactin X C\<close> \<open>connectedin X C\<close> compactin_subset_topspace inf_absorb2)
show "f u \<in> f ` topspace (subtopology X C)"
by (simp add: \<open>u \<in> C\<close> \<open>u \<in> topspace X\<close>)
show "f v \<in> f ` topspace (subtopology X C)"
by (simp add: \<open>v \<in> C\<close> \<open>v \<in> topspace X\<close>)
qed
then show "f ` topspace X \<inter> {f x - d..f x + d} \<subseteq> f ` topspace (subtopology X C)"
using sub_fuv by blast
qed (auto simp: gf)
then have contg: "continuous_map (subtopology euclideanreal (?S \<inter> {f x - d .. f x + d})) X g"
using continuous_map_in_subtopology by blast
have "\<exists>e>0. \<forall>x \<in> ?S \<inter> {f x - d .. f x + d} \<inter> ball (f x) e. g x \<in> T"
using openin_continuous_map_preimage [OF contg \<open>openin X T\<close>] x \<open>x \<in> T\<close> \<open>0 < d\<close>
unfolding openin_euclidean_subtopology_iff
by (force simp: gf dist_commute)
then obtain e where "e > 0 \<and> (\<forall>x\<in>f ` topspace X \<inter> {f x - d..f x + d} \<inter> ball (f x) e. g x \<in> T)"
by metis
with \<open>0 < d\<close> have "min d e > 0" "\<forall>u. u \<in> topspace X \<longrightarrow> \<bar>f x - f u\<bar> < min d e \<longrightarrow> u \<in> T"
using dist_real_def gf by force+
then show ?thesis
by (metis (full_types) Int_iff dist_real_def image_iff mem_ball gf)
qed
qed
then obtain d where d: "\<And>r. r \<in> ?S \<inter> g -` T \<Longrightarrow>
d r > 0 \<and> (\<forall>x \<in> ?S \<inter> ball r (d r). g x \<in> T)"
by metis
show ?thesis
unfolding openin_subtopology
proof (intro exI conjI)
show "{x \<in> ?S. g x \<in> T} = (\<Union>r \<in> ?S \<inter> g -` T. ball r (d r)) \<inter> f ` topspace X"
using d by (auto simp: gf)
qed auto
qed
then show "continuous_map (top_of_set ?S) X g"
by (simp add: continuous_map_def gf)
qed (auto simp: gf)
qed
qed
subsubsection \<open>An injective function into R is a homeomorphism and so an open map.\<close>
lemma injective_into_1d_eq_homeomorphism:
fixes f :: "'a::topological_space \<Rightarrow> real"
assumes f: "continuous_on S f" and S: "path_connected S"
shows "inj_on f S \<longleftrightarrow> (\<exists>g. homeomorphism S (f ` S) f g)"
proof
show "\<exists>g. homeomorphism S (f ` S) f g"
if "inj_on f S"
proof -
have "embedding_map (top_of_set S) euclideanreal f"
using that embedding_map_into_euclideanreal [of "top_of_set S" f] assms by auto
then show ?thesis
by (simp add: embedding_map_def) (metis all_closedin_homeomorphic_image f homeomorphism_injective_closed_map that)
qed
qed (metis homeomorphism_def inj_onI)
lemma injective_into_1d_imp_open_map:
fixes f :: "'a::topological_space \<Rightarrow> real"
assumes "continuous_on S f" "path_connected S" "inj_on f S" "openin (subtopology euclidean S) T"
shows "openin (subtopology euclidean (f ` S)) (f ` T)"
using assms homeomorphism_imp_open_map injective_into_1d_eq_homeomorphism by blast
lemma homeomorphism_into_1d:
fixes f :: "'a::topological_space \<Rightarrow> real"
assumes "path_connected S" "continuous_on S f" "f ` S = T" "inj_on f S"
shows "\<exists>g. homeomorphism S T f g"
using assms injective_into_1d_eq_homeomorphism by blast
subsection\<^marker>\<open>tag unimportant\<close> \<open>Rectangular paths\<close>
definition\<^marker>\<open>tag unimportant\<close> rectpath where
"rectpath a1 a3 = (let a2 = Complex (Re a3) (Im a1); a4 = Complex (Re a1) (Im a3)
in linepath a1 a2 +++ linepath a2 a3 +++ linepath a3 a4 +++ linepath a4 a1)"
lemma path_rectpath [simp, intro]: "path (rectpath a b)"
by (simp add: Let_def rectpath_def)
lemma pathstart_rectpath [simp]: "pathstart (rectpath a1 a3) = a1"
by (simp add: rectpath_def Let_def)
lemma pathfinish_rectpath [simp]: "pathfinish (rectpath a1 a3) = a1"
by (simp add: rectpath_def Let_def)
lemma simple_path_rectpath [simp, intro]:
assumes "Re a1 \<noteq> Re a3" "Im a1 \<noteq> Im a3"
shows "simple_path (rectpath a1 a3)"
unfolding rectpath_def Let_def using assms
by (intro simple_path_join_loop arc_join arc_linepath)
(auto simp: complex_eq_iff path_image_join closed_segment_same_Re closed_segment_same_Im)
lemma path_image_rectpath:
assumes "Re a1 \<le> Re a3" "Im a1 \<le> Im a3"
shows "path_image (rectpath a1 a3) =
{z. Re z \<in> {Re a1, Re a3} \<and> Im z \<in> {Im a1..Im a3}} \<union>
{z. Im z \<in> {Im a1, Im a3} \<and> Re z \<in> {Re a1..Re a3}}" (is "?lhs = ?rhs")
proof -
define a2 a4 where "a2 = Complex (Re a3) (Im a1)" and "a4 = Complex (Re a1) (Im a3)"
have "?lhs = closed_segment a1 a2 \<union> closed_segment a2 a3 \<union>
closed_segment a4 a3 \<union> closed_segment a1 a4"
by (simp_all add: rectpath_def Let_def path_image_join closed_segment_commute
a2_def a4_def Un_assoc)
also have "\<dots> = ?rhs" using assms
by (auto simp: rectpath_def Let_def path_image_join a2_def a4_def
closed_segment_same_Re closed_segment_same_Im closed_segment_eq_real_ivl)
finally show ?thesis .
qed
lemma path_image_rectpath_subset_cbox:
assumes "Re a \<le> Re b" "Im a \<le> Im b"
shows "path_image (rectpath a b) \<subseteq> cbox a b"
using assms by (auto simp: path_image_rectpath in_cbox_complex_iff)
lemma path_image_rectpath_inter_box:
assumes "Re a \<le> Re b" "Im a \<le> Im b"
shows "path_image (rectpath a b) \<inter> box a b = {}"
using assms by (auto simp: path_image_rectpath in_box_complex_iff)
lemma path_image_rectpath_cbox_minus_box:
assumes "Re a \<le> Re b" "Im a \<le> Im b"
shows "path_image (rectpath a b) = cbox a b - box a b"
using assms by (auto simp: path_image_rectpath in_cbox_complex_iff
in_box_complex_iff)
end
diff --git a/src/HOL/Analysis/Polytope.thy b/src/HOL/Analysis/Polytope.thy
--- a/src/HOL/Analysis/Polytope.thy
+++ b/src/HOL/Analysis/Polytope.thy
@@ -1,4005 +1,4005 @@
section \<open>Faces, Extreme Points, Polytopes, Polyhedra etc\<close>
text\<open>Ported from HOL Light by L C Paulson\<close>
theory Polytope
imports Cartesian_Euclidean_Space Path_Connected
begin
subsection \<open>Faces of a (usually convex) set\<close>
definition\<^marker>\<open>tag important\<close> face_of :: "['a::real_vector set, 'a set] \<Rightarrow> bool" (infixr "(face'_of)" 50)
where
"T face_of S \<longleftrightarrow>
T \<subseteq> S \<and> convex T \<and>
(\<forall>a \<in> S. \<forall>b \<in> S. \<forall>x \<in> T. x \<in> open_segment a b \<longrightarrow> a \<in> T \<and> b \<in> T)"
lemma face_ofD: "\<lbrakk>T face_of S; x \<in> open_segment a b; a \<in> S; b \<in> S; x \<in> T\<rbrakk> \<Longrightarrow> a \<in> T \<and> b \<in> T"
unfolding face_of_def by blast
lemma face_of_translation_eq [simp]:
"((+) a ` T face_of (+) a ` S) \<longleftrightarrow> T face_of S"
proof -
have *: "\<And>a T S. T face_of S \<Longrightarrow> ((+) a ` T face_of (+) a ` S)"
apply (simp add: face_of_def Ball_def, clarify)
by (meson imageI open_segment_translation_eq)
show ?thesis
apply (rule iffI)
apply (force simp: image_comp o_def dest: * [where a = "-a"])
apply (blast intro: *)
done
qed
lemma face_of_linear_image:
assumes "linear f" "inj f"
shows "(f ` c face_of f ` S) \<longleftrightarrow> c face_of S"
by (simp add: face_of_def inj_image_subset_iff inj_image_mem_iff open_segment_linear_image assms)
lemma face_of_refl: "convex S \<Longrightarrow> S face_of S"
by (auto simp: face_of_def)
lemma face_of_refl_eq: "S face_of S \<longleftrightarrow> convex S"
by (auto simp: face_of_def)
lemma empty_face_of [iff]: "{} face_of S"
by (simp add: face_of_def)
lemma face_of_empty [simp]: "S face_of {} \<longleftrightarrow> S = {}"
by (meson empty_face_of face_of_def subset_empty)
lemma face_of_trans [trans]: "\<lbrakk>S face_of T; T face_of u\<rbrakk> \<Longrightarrow> S face_of u"
unfolding face_of_def by (safe; blast)
lemma face_of_face: "T face_of S \<Longrightarrow> (f face_of T \<longleftrightarrow> f face_of S \<and> f \<subseteq> T)"
unfolding face_of_def by (safe; blast)
lemma face_of_subset: "\<lbrakk>F face_of S; F \<subseteq> T; T \<subseteq> S\<rbrakk> \<Longrightarrow> F face_of T"
unfolding face_of_def by (safe; blast)
lemma face_of_slice: "\<lbrakk>F face_of S; convex T\<rbrakk> \<Longrightarrow> (F \<inter> T) face_of (S \<inter> T)"
unfolding face_of_def by (blast intro: convex_Int)
lemma face_of_Int: "\<lbrakk>t1 face_of S; t2 face_of S\<rbrakk> \<Longrightarrow> (t1 \<inter> t2) face_of S"
unfolding face_of_def by (blast intro: convex_Int)
lemma face_of_Inter: "\<lbrakk>A \<noteq> {}; \<And>T. T \<in> A \<Longrightarrow> T face_of S\<rbrakk> \<Longrightarrow> (\<Inter> A) face_of S"
unfolding face_of_def by (blast intro: convex_Inter)
lemma face_of_Int_Int: "\<lbrakk>F face_of T; F' face_of t'\<rbrakk> \<Longrightarrow> (F \<inter> F') face_of (T \<inter> t')"
unfolding face_of_def by (blast intro: convex_Int)
lemma face_of_imp_subset: "T face_of S \<Longrightarrow> T \<subseteq> S"
unfolding face_of_def by blast
proposition face_of_imp_eq_affine_Int:
fixes S :: "'a::euclidean_space set"
assumes S: "convex S" and T: "T face_of S"
shows "T = (affine hull T) \<inter> S"
proof -
have "convex T" using T by (simp add: face_of_def)
have *: False if x: "x \<in> affine hull T" and "x \<in> S" "x \<notin> T" and y: "y \<in> rel_interior T" for x y
proof -
obtain e where "e>0" and e: "cball y e \<inter> affine hull T \<subseteq> T"
using y by (auto simp: rel_interior_cball)
have "y \<noteq> x" "y \<in> S" "y \<in> T"
using face_of_imp_subset rel_interior_subset T that by blast+
then have zne: "\<And>u. \<lbrakk>u \<in> {0<..<1}; (1 - u) *\<^sub>R y + u *\<^sub>R x \<in> T\<rbrakk> \<Longrightarrow> False"
using \<open>x \<in> S\<close> \<open>x \<notin> T\<close> \<open>T face_of S\<close> unfolding face_of_def
apply clarify
apply (drule_tac x=x in bspec, assumption)
apply (drule_tac x=y in bspec, assumption)
apply (subst (asm) open_segment_commute)
apply (force simp: open_segment_image_interval image_def)
done
have in01: "min (1/2) (e / norm (x - y)) \<in> {0<..<1}"
using \<open>y \<noteq> x\<close> \<open>e > 0\<close> by simp
show ?thesis
apply (rule zne [OF in01])
apply (rule e [THEN subsetD])
apply (rule IntI)
using \<open>y \<noteq> x\<close> \<open>e > 0\<close>
apply (simp add: cball_def dist_norm algebra_simps)
apply (simp add: Real_Vector_Spaces.scaleR_diff_right [symmetric] norm_minus_commute min_mult_distrib_right)
apply (rule mem_affine [OF affine_affine_hull _ x])
using \<open>y \<in> T\<close> apply (auto simp: hull_inc)
done
qed
show ?thesis
apply (rule subset_antisym)
using assms apply (simp add: hull_subset face_of_imp_subset)
apply (cases "T={}", simp)
apply (force simp: rel_interior_eq_empty [symmetric] \<open>convex T\<close> intro: *)
done
qed
lemma face_of_imp_closed:
fixes S :: "'a::euclidean_space set"
assumes "convex S" "closed S" "T face_of S" shows "closed T"
by (metis affine_affine_hull affine_closed closed_Int face_of_imp_eq_affine_Int assms)
lemma face_of_Int_supporting_hyperplane_le_strong:
assumes "convex(S \<inter> {x. a \<bullet> x = b})" and aleb: "\<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<le> b"
shows "(S \<inter> {x. a \<bullet> x = b}) face_of S"
proof -
have *: "a \<bullet> u = a \<bullet> x" if "x \<in> open_segment u v" "u \<in> S" "v \<in> S" and b: "b = a \<bullet> x"
for u v x
proof (rule antisym)
show "a \<bullet> u \<le> a \<bullet> x"
using aleb \<open>u \<in> S\<close> \<open>b = a \<bullet> x\<close> by blast
next
obtain \<xi> where "b = a \<bullet> ((1 - \<xi>) *\<^sub>R u + \<xi> *\<^sub>R v)" "0 < \<xi>" "\<xi> < 1"
using \<open>b = a \<bullet> x\<close> \<open>x \<in> open_segment u v\<close> in_segment
by (auto simp: open_segment_image_interval split: if_split_asm)
then have "b + \<xi> * (a \<bullet> u) \<le> a \<bullet> u + \<xi> * b"
using aleb [OF \<open>v \<in> S\<close>] by (simp add: algebra_simps)
then have "(1 - \<xi>) * b \<le> (1 - \<xi>) * (a \<bullet> u)"
by (simp add: algebra_simps)
then have "b \<le> a \<bullet> u"
using \<open>\<xi> < 1\<close> by auto
with b show "a \<bullet> x \<le> a \<bullet> u" by simp
qed
show ?thesis
apply (simp add: face_of_def assms)
using "*" open_segment_commute by blast
qed
lemma face_of_Int_supporting_hyperplane_ge_strong:
"\<lbrakk>convex(S \<inter> {x. a \<bullet> x = b}); \<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<ge> b\<rbrakk>
\<Longrightarrow> (S \<inter> {x. a \<bullet> x = b}) face_of S"
using face_of_Int_supporting_hyperplane_le_strong [of S "-a" "-b"] by simp
lemma face_of_Int_supporting_hyperplane_le:
"\<lbrakk>convex S; \<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<le> b\<rbrakk> \<Longrightarrow> (S \<inter> {x. a \<bullet> x = b}) face_of S"
by (simp add: convex_Int convex_hyperplane face_of_Int_supporting_hyperplane_le_strong)
lemma face_of_Int_supporting_hyperplane_ge:
"\<lbrakk>convex S; \<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<ge> b\<rbrakk> \<Longrightarrow> (S \<inter> {x. a \<bullet> x = b}) face_of S"
by (simp add: convex_Int convex_hyperplane face_of_Int_supporting_hyperplane_ge_strong)
lemma face_of_imp_convex: "T face_of S \<Longrightarrow> convex T"
using face_of_def by blast
lemma face_of_imp_compact:
fixes S :: "'a::euclidean_space set"
shows "\<lbrakk>convex S; compact S; T face_of S\<rbrakk> \<Longrightarrow> compact T"
by (meson bounded_subset compact_eq_bounded_closed face_of_imp_closed face_of_imp_subset)
lemma face_of_Int_subface:
"\<lbrakk>A \<inter> B face_of A; A \<inter> B face_of B; C face_of A; D face_of B\<rbrakk>
\<Longrightarrow> (C \<inter> D) face_of C \<and> (C \<inter> D) face_of D"
by (meson face_of_Int_Int face_of_face inf_le1 inf_le2)
lemma subset_of_face_of:
fixes S :: "'a::real_normed_vector set"
assumes "T face_of S" "u \<subseteq> S" "T \<inter> (rel_interior u) \<noteq> {}"
shows "u \<subseteq> T"
proof
fix c
assume "c \<in> u"
obtain b where "b \<in> T" "b \<in> rel_interior u" using assms by auto
then obtain e where "e>0" "b \<in> u" and e: "cball b e \<inter> affine hull u \<subseteq> u"
by (auto simp: rel_interior_cball)
show "c \<in> T"
proof (cases "b=c")
case True with \<open>b \<in> T\<close> show ?thesis by blast
next
case False
define d where "d = b + (e / norm(b - c)) *\<^sub>R (b - c)"
have "d \<in> cball b e \<inter> affine hull u"
using \<open>e > 0\<close> \<open>b \<in> u\<close> \<open>c \<in> u\<close>
by (simp add: d_def dist_norm hull_inc mem_affine_3_minus False)
with e have "d \<in> u" by blast
have nbc: "norm (b - c) + e > 0" using \<open>e > 0\<close>
by (metis add.commute le_less_trans less_add_same_cancel2 norm_ge_zero)
then have [simp]: "d \<noteq> c" using False scaleR_cancel_left [of "1 + (e / norm (b - c))" b c]
by (simp add: algebra_simps d_def) (simp add: field_split_simps)
have [simp]: "((e - e * e / (e + norm (b - c))) / norm (b - c)) = (e / (e + norm (b - c)))"
using False nbc
by (simp add: divide_simps) (simp add: algebra_simps)
have "b \<in> open_segment d c"
apply (simp add: open_segment_image_interval)
apply (simp add: d_def algebra_simps image_def)
apply (rule_tac x="e / (e + norm (b - c))" in bexI)
using False nbc \<open>0 < e\<close>
apply (auto simp: algebra_simps)
done
then have "d \<in> T \<and> c \<in> T"
apply (rule face_ofD [OF \<open>T face_of S\<close>])
using \<open>d \<in> u\<close> \<open>c \<in> u\<close> \<open>u \<subseteq> S\<close> \<open>b \<in> T\<close> apply auto
done
then show ?thesis ..
qed
qed
lemma face_of_eq:
fixes S :: "'a::real_normed_vector set"
assumes "T face_of S" "u face_of S" "(rel_interior T) \<inter> (rel_interior u) \<noteq> {}"
shows "T = u"
apply (rule subset_antisym)
apply (metis assms disjoint_iff_not_equal face_of_imp_subset rel_interior_subset subsetCE subset_of_face_of)
by (metis assms disjoint_iff_not_equal face_of_imp_subset rel_interior_subset subset_iff subset_of_face_of)
lemma face_of_disjoint_rel_interior:
fixes S :: "'a::real_normed_vector set"
assumes "T face_of S" "T \<noteq> S"
shows "T \<inter> rel_interior S = {}"
by (meson assms subset_of_face_of face_of_imp_subset order_refl subset_antisym)
lemma face_of_disjoint_interior:
fixes S :: "'a::real_normed_vector set"
assumes "T face_of S" "T \<noteq> S"
shows "T \<inter> interior S = {}"
proof -
have "T \<inter> interior S \<subseteq> rel_interior S"
by (meson inf_sup_ord(2) interior_subset_rel_interior order.trans)
thus ?thesis
by (metis (no_types) Int_greatest assms face_of_disjoint_rel_interior inf_sup_ord(1) subset_empty)
qed
lemma face_of_subset_rel_boundary:
fixes S :: "'a::real_normed_vector set"
assumes "T face_of S" "T \<noteq> S"
shows "T \<subseteq> (S - rel_interior S)"
by (meson DiffI assms disjoint_iff_not_equal face_of_disjoint_rel_interior face_of_imp_subset rev_subsetD subsetI)
lemma face_of_subset_rel_frontier:
fixes S :: "'a::real_normed_vector set"
assumes "T face_of S" "T \<noteq> S"
shows "T \<subseteq> rel_frontier S"
using assms closure_subset face_of_disjoint_rel_interior face_of_imp_subset rel_frontier_def by fastforce
lemma face_of_aff_dim_lt:
fixes S :: "'a::euclidean_space set"
assumes "convex S" "T face_of S" "T \<noteq> S"
shows "aff_dim T < aff_dim S"
proof -
have "aff_dim T \<le> aff_dim S"
by (simp add: face_of_imp_subset aff_dim_subset assms)
moreover have "aff_dim T \<noteq> aff_dim S"
proof (cases "T = {}")
case True then show ?thesis
by (metis aff_dim_empty \<open>T \<noteq> S\<close>)
next case False then show ?thesis
by (metis Set.set_insert assms convex_rel_frontier_aff_dim dual_order.irrefl face_of_imp_convex face_of_subset_rel_frontier insert_not_empty subsetI)
qed
ultimately show ?thesis
by simp
qed
lemma subset_of_face_of_affine_hull:
fixes S :: "'a::euclidean_space set"
assumes T: "T face_of S" and "convex S" "U \<subseteq> S" and dis: "\<not> disjnt (affine hull T) (rel_interior U)"
shows "U \<subseteq> T"
apply (rule subset_of_face_of [OF T \<open>U \<subseteq> S\<close>])
using face_of_imp_eq_affine_Int [OF \<open>convex S\<close> T]
using rel_interior_subset [of U] dis
using \<open>U \<subseteq> S\<close> disjnt_def by fastforce
lemma affine_hull_face_of_disjoint_rel_interior:
fixes S :: "'a::euclidean_space set"
assumes "convex S" "F face_of S" "F \<noteq> S"
shows "affine hull F \<inter> rel_interior S = {}"
by (metis assms disjnt_def face_of_imp_subset order_refl subset_antisym subset_of_face_of_affine_hull)
lemma affine_diff_divide:
assumes "affine S" "k \<noteq> 0" "k \<noteq> 1" and xy: "x \<in> S" "y /\<^sub>R (1 - k) \<in> S"
shows "(x - y) /\<^sub>R k \<in> S"
proof -
have "inverse(k) *\<^sub>R (x - y) = (1 - inverse k) *\<^sub>R inverse(1 - k) *\<^sub>R y + inverse(k) *\<^sub>R x"
using assms
by (simp add: algebra_simps) (simp add: scaleR_left_distrib [symmetric] field_split_simps)
then show ?thesis
using \<open>affine S\<close> xy by (auto simp: affine_alt)
qed
proposition face_of_convex_hulls:
assumes S: "finite S" "T \<subseteq> S" and disj: "affine hull T \<inter> convex hull (S - T) = {}"
shows "(convex hull T) face_of (convex hull S)"
proof -
have fin: "finite T" "finite (S - T)" using assms
by (auto simp: finite_subset)
have *: "x \<in> convex hull T"
if x: "x \<in> convex hull S" and y: "y \<in> convex hull S" and w: "w \<in> convex hull T" "w \<in> open_segment x y"
for x y w
proof -
have waff: "w \<in> affine hull T"
using convex_hull_subset_affine_hull w by blast
obtain a b where a: "\<And>i. i \<in> S \<Longrightarrow> 0 \<le> a i" and asum: "sum a S = 1" and aeqx: "(\<Sum>i\<in>S. a i *\<^sub>R i) = x"
and b: "\<And>i. i \<in> S \<Longrightarrow> 0 \<le> b i" and bsum: "sum b S = 1" and beqy: "(\<Sum>i\<in>S. b i *\<^sub>R i) = y"
using x y by (auto simp: assms convex_hull_finite)
obtain u where "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> convex hull T" "x \<noteq> y" and weq: "w = (1 - u) *\<^sub>R x + u *\<^sub>R y"
and u01: "0 < u" "u < 1"
using w by (auto simp: open_segment_image_interval split: if_split_asm)
define c where "c i = (1 - u) * a i + u * b i" for i
have cge0: "\<And>i. i \<in> S \<Longrightarrow> 0 \<le> c i"
using a b u01 by (simp add: c_def)
have sumc1: "sum c S = 1"
by (simp add: c_def sum.distrib sum_distrib_left [symmetric] asum bsum)
have sumci_xy: "(\<Sum>i\<in>S. c i *\<^sub>R i) = (1 - u) *\<^sub>R x + u *\<^sub>R y"
apply (simp add: c_def sum.distrib scaleR_left_distrib)
by (simp only: scaleR_scaleR [symmetric] Real_Vector_Spaces.scaleR_right.sum [symmetric] aeqx beqy)
show ?thesis
proof (cases "sum c (S - T) = 0")
case True
have ci0: "\<And>i. i \<in> (S - T) \<Longrightarrow> c i = 0"
using True cge0 fin(2) sum_nonneg_eq_0_iff by auto
have a0: "a i = 0" if "i \<in> (S - T)" for i
using ci0 [OF that] u01 a [of i] b [of i] that
by (simp add: c_def Groups.ordered_comm_monoid_add_class.add_nonneg_eq_0_iff)
have [simp]: "sum a T = 1"
using assms by (metis sum.mono_neutral_cong_right a0 asum)
show ?thesis
apply (simp add: convex_hull_finite \<open>finite T\<close>)
apply (rule_tac x=a in exI)
using a0 assms
apply (auto simp: cge0 a aeqx [symmetric] sum.mono_neutral_right)
done
next
case False
define k where "k = sum c (S - T)"
have "k > 0" using False
unfolding k_def by (metis DiffD1 antisym_conv cge0 sum_nonneg not_less)
have weq_sumsum: "w = sum (\<lambda>x. c x *\<^sub>R x) T + sum (\<lambda>x. c x *\<^sub>R x) (S - T)"
by (metis (no_types) add.commute S(1) S(2) sum.subset_diff sumci_xy weq)
show ?thesis
proof (cases "k = 1")
case True
then have "sum c T = 0"
by (simp add: S k_def sum_diff sumc1)
then have [simp]: "sum c (S - T) = 1"
by (simp add: S sum_diff sumc1)
have ci0: "\<And>i. i \<in> T \<Longrightarrow> c i = 0"
by (meson \<open>finite T\<close> \<open>sum c T = 0\<close> \<open>T \<subseteq> S\<close> cge0 sum_nonneg_eq_0_iff subsetCE)
then have [simp]: "(\<Sum>i\<in>S-T. c i *\<^sub>R i) = w"
by (simp add: weq_sumsum)
have "w \<in> convex hull (S - T)"
apply (simp add: convex_hull_finite fin)
apply (rule_tac x=c in exI)
apply (auto simp: cge0 weq True k_def)
done
then show ?thesis
using disj waff by blast
next
case False
then have sumcf: "sum c T = 1 - k"
by (simp add: S k_def sum_diff sumc1)
have "(\<Sum>i\<in>T. c i *\<^sub>R i) /\<^sub>R (1 - k) \<in> convex hull T"
apply (simp add: convex_hull_finite fin)
apply (rule_tac x="\<lambda>i. inverse (1-k) * c i" in exI)
apply auto
apply (metis sumcf cge0 inverse_nonnegative_iff_nonnegative mult_nonneg_nonneg S(2) sum_nonneg subsetCE)
apply (metis False mult.commute right_inverse right_minus_eq sum_distrib_left sumcf)
by (metis (mono_tags, lifting) scaleR_right.sum scaleR_scaleR sum.cong)
with \<open>0 < k\<close> have "inverse(k) *\<^sub>R (w - sum (\<lambda>i. c i *\<^sub>R i) T) \<in> affine hull T"
by (simp add: affine_diff_divide [OF affine_affine_hull] False waff convex_hull_subset_affine_hull [THEN subsetD])
moreover have "inverse(k) *\<^sub>R (w - sum (\<lambda>x. c x *\<^sub>R x) T) \<in> convex hull (S - T)"
apply (simp add: weq_sumsum convex_hull_finite fin)
apply (rule_tac x="\<lambda>i. inverse k * c i" in exI)
using \<open>k > 0\<close> cge0
apply (auto simp: scaleR_right.sum sum_distrib_left [symmetric] k_def [symmetric])
done
ultimately show ?thesis
using disj by blast
qed
qed
qed
have [simp]: "convex hull T \<subseteq> convex hull S"
by (simp add: \<open>T \<subseteq> S\<close> hull_mono)
show ?thesis
using open_segment_commute by (auto simp: face_of_def intro: *)
qed
proposition face_of_convex_hull_insert:
"\<lbrakk>finite S; a \<notin> affine hull S; T face_of convex hull S\<rbrakk> \<Longrightarrow> T face_of convex hull insert a S"
apply (rule face_of_trans, blast)
apply (rule face_of_convex_hulls; force simp: insert_Diff_if)
done
proposition face_of_affine_trivial:
assumes "affine S" "T face_of S"
shows "T = {} \<or> T = S"
proof (rule ccontr, clarsimp)
assume "T \<noteq> {}" "T \<noteq> S"
then obtain a where "a \<in> T" by auto
then have "a \<in> S"
using \<open>T face_of S\<close> face_of_imp_subset by blast
have "S \<subseteq> T"
proof
fix b assume "b \<in> S"
show "b \<in> T"
proof (cases "a = b")
case True with \<open>a \<in> T\<close> show ?thesis by auto
next
case False
then have "a \<in> open_segment (2 *\<^sub>R a - b) b"
apply (auto simp: open_segment_def closed_segment_def)
apply (rule_tac x="1/2" in exI)
apply (simp add: algebra_simps)
by (simp add: scaleR_2)
moreover have "2 *\<^sub>R a - b \<in> S"
by (rule mem_affine [OF \<open>affine S\<close> \<open>a \<in> S\<close> \<open>b \<in> S\<close>, of 2 "-1", simplified])
moreover note \<open>b \<in> S\<close> \<open>a \<in> T\<close>
ultimately show ?thesis
by (rule face_ofD [OF \<open>T face_of S\<close>, THEN conjunct2])
qed
qed
then show False
using \<open>T \<noteq> S\<close> \<open>T face_of S\<close> face_of_imp_subset by blast
qed
lemma face_of_affine_eq:
"affine S \<Longrightarrow> (T face_of S \<longleftrightarrow> T = {} \<or> T = S)"
using affine_imp_convex face_of_affine_trivial face_of_refl by auto
proposition Inter_faces_finite_altbound:
fixes T :: "'a::euclidean_space set set"
assumes cfaI: "\<And>c. c \<in> T \<Longrightarrow> c face_of S"
shows "\<exists>F'. finite F' \<and> F' \<subseteq> T \<and> card F' \<le> DIM('a) + 2 \<and> \<Inter>F' = \<Inter>T"
proof (cases "\<forall>F'. finite F' \<and> F' \<subseteq> T \<and> card F' \<le> DIM('a) + 2 \<longrightarrow> (\<exists>c. c \<in> T \<and> c \<inter> (\<Inter>F') \<subset> (\<Inter>F'))")
case True
then obtain c where c:
"\<And>F'. \<lbrakk>finite F'; F' \<subseteq> T; card F' \<le> DIM('a) + 2\<rbrakk> \<Longrightarrow> c F' \<in> T \<and> c F' \<inter> (\<Inter>F') \<subset> (\<Inter>F')"
by metis
define d where "d = rec_nat {c{}} (\<lambda>n r. insert (c r) r)"
have [simp]: "d 0 = {c {}}"
by (simp add: d_def)
have dSuc [simp]: "\<And>n. d (Suc n) = insert (c (d n)) (d n)"
by (simp add: d_def)
have dn_notempty: "d n \<noteq> {}" for n
by (induction n) auto
have dn_le_Suc: "d n \<subseteq> T \<and> finite(d n) \<and> card(d n) \<le> Suc n" if "n \<le> DIM('a) + 2" for n
using that
proof (induction n)
case 0
then show ?case by (simp add: c)
next
case (Suc n)
then show ?case by (auto simp: c card_insert_if)
qed
have aff_dim_le: "aff_dim(\<Inter>(d n)) \<le> DIM('a) - int n" if "n \<le> DIM('a) + 2" for n
using that
proof (induction n)
case 0
then show ?case
by (simp add: aff_dim_le_DIM)
next
case (Suc n)
have fs: "\<Inter>(d (Suc n)) face_of S"
by (meson Suc.prems cfaI dn_le_Suc dn_notempty face_of_Inter subsetCE)
have condn: "convex (\<Inter>(d n))"
using Suc.prems nat_le_linear not_less_eq_eq
by (blast intro: face_of_imp_convex cfaI convex_Inter dest: dn_le_Suc)
have fdn: "\<Inter>(d (Suc n)) face_of \<Inter>(d n)"
by (metis (no_types, lifting) Inter_anti_mono Suc.prems dSuc cfaI dn_le_Suc dn_notempty face_of_Inter face_of_imp_subset face_of_subset subset_iff subset_insertI)
have ne: "\<Inter>(d (Suc n)) \<noteq> \<Inter>(d n)"
by (metis (no_types, lifting) Suc.prems Suc_leD c complete_lattice_class.Inf_insert dSuc dn_le_Suc less_irrefl order.trans)
have *: "\<And>m::int. \<And>d. \<And>d'::int. d < d' \<and> d' \<le> m - n \<Longrightarrow> d \<le> m - of_nat(n+1)"
by arith
have "aff_dim (\<Inter>(d (Suc n))) < aff_dim (\<Inter>(d n))"
by (rule face_of_aff_dim_lt [OF condn fdn ne])
moreover have "aff_dim (\<Inter>(d n)) \<le> int (DIM('a)) - int n"
using Suc by auto
ultimately
have "aff_dim (\<Inter>(d (Suc n))) \<le> int (DIM('a)) - (n+1)" by arith
then show ?case by linarith
qed
have "aff_dim (\<Inter>(d (DIM('a) + 2))) \<le> -2"
using aff_dim_le [OF order_refl] by simp
with aff_dim_geq [of "\<Inter>(d (DIM('a) + 2))"] show ?thesis
using order.trans by fastforce
next
case False
then show ?thesis
apply simp
apply (erule ex_forward)
by blast
qed
lemma faces_of_translation:
"{F. F face_of image (\<lambda>x. a + x) S} = image (image (\<lambda>x. a + x)) {F. F face_of S}"
apply (rule subset_antisym, clarify)
apply (auto simp: image_iff)
apply (metis face_of_imp_subset face_of_translation_eq subset_imageE)
done
proposition face_of_Times:
assumes "F face_of S" and "F' face_of S'"
shows "(F \<times> F') face_of (S \<times> S')"
proof -
have "F \<times> F' \<subseteq> S \<times> S'"
using assms [unfolded face_of_def] by blast
moreover
have "convex (F \<times> F')"
using assms [unfolded face_of_def] by (blast intro: convex_Times)
moreover
have "a \<in> F \<and> a' \<in> F' \<and> b \<in> F \<and> b' \<in> F'"
if "a \<in> S" "b \<in> S" "a' \<in> S'" "b' \<in> S'" "x \<in> F \<times> F'" "x \<in> open_segment (a,a') (b,b')"
for a b a' b' x
proof (cases "b=a \<or> b'=a'")
case True with that show ?thesis
using assms
by (force simp: in_segment dest: face_ofD)
next
case False with assms [unfolded face_of_def] that show ?thesis
by (blast dest!: open_segment_PairD)
qed
ultimately show ?thesis
unfolding face_of_def by blast
qed
corollary face_of_Times_decomp:
fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
shows "c face_of (S \<times> S') \<longleftrightarrow> (\<exists>F F'. F face_of S \<and> F' face_of S' \<and> c = F \<times> F')"
(is "?lhs = ?rhs")
proof
assume c: ?lhs
show ?rhs
proof (cases "c = {}")
case True then show ?thesis by auto
next
case False
have 1: "fst ` c \<subseteq> S" "snd ` c \<subseteq> S'"
using c face_of_imp_subset by fastforce+
have "convex c"
using c by (metis face_of_imp_convex)
have conv: "convex (fst ` c)" "convex (snd ` c)"
by (simp_all add: \<open>convex c\<close> convex_linear_image linear_fst linear_snd)
have fstab: "a \<in> fst ` c \<and> b \<in> fst ` c"
if "a \<in> S" "b \<in> S" "x \<in> open_segment a b" "(x,x') \<in> c" for a b x x'
proof -
have *: "(x,x') \<in> open_segment (a,x') (b,x')"
using that by (auto simp: in_segment)
show ?thesis
using face_ofD [OF c *] that face_of_imp_subset [OF c] by force
qed
have fst: "fst ` c face_of S"
by (force simp: face_of_def 1 conv fstab)
have sndab: "a' \<in> snd ` c \<and> b' \<in> snd ` c"
if "a' \<in> S'" "b' \<in> S'" "x' \<in> open_segment a' b'" "(x,x') \<in> c" for a' b' x x'
proof -
have *: "(x,x') \<in> open_segment (x,a') (x,b')"
using that by (auto simp: in_segment)
show ?thesis
using face_ofD [OF c *] that face_of_imp_subset [OF c] by force
qed
have snd: "snd ` c face_of S'"
by (force simp: face_of_def 1 conv sndab)
have cc: "rel_interior c \<subseteq> rel_interior (fst ` c) \<times> rel_interior (snd ` c)"
by (force simp: face_of_Times rel_interior_Times conv fst snd \<open>convex c\<close> linear_fst linear_snd rel_interior_convex_linear_image [symmetric])
have "c = fst ` c \<times> snd ` c"
apply (rule face_of_eq [OF c])
apply (simp_all add: face_of_Times rel_interior_Times conv fst snd)
using False rel_interior_eq_empty \<open>convex c\<close> cc
apply blast
done
with fst snd show ?thesis by metis
qed
next
assume ?rhs with face_of_Times show ?lhs by auto
qed
lemma face_of_Times_eq:
fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
shows "(F \<times> F') face_of (S \<times> S') \<longleftrightarrow>
F = {} \<or> F' = {} \<or> F face_of S \<and> F' face_of S'"
by (auto simp: face_of_Times_decomp times_eq_iff)
lemma hyperplane_face_of_halfspace_le: "{x. a \<bullet> x = b} face_of {x. a \<bullet> x \<le> b}"
proof -
have "{x. a \<bullet> x \<le> b} \<inter> {x. a \<bullet> x = b} = {x. a \<bullet> x = b}"
by auto
with face_of_Int_supporting_hyperplane_le [OF convex_halfspace_le [of a b], of a b]
show ?thesis by auto
qed
lemma hyperplane_face_of_halfspace_ge: "{x. a \<bullet> x = b} face_of {x. a \<bullet> x \<ge> b}"
proof -
have "{x. a \<bullet> x \<ge> b} \<inter> {x. a \<bullet> x = b} = {x. a \<bullet> x = b}"
by auto
with face_of_Int_supporting_hyperplane_ge [OF convex_halfspace_ge [of b a], of b a]
show ?thesis by auto
qed
lemma face_of_halfspace_le:
fixes a :: "'n::euclidean_space"
shows "F face_of {x. a \<bullet> x \<le> b} \<longleftrightarrow>
F = {} \<or> F = {x. a \<bullet> x = b} \<or> F = {x. a \<bullet> x \<le> b}"
(is "?lhs = ?rhs")
proof (cases "a = 0")
case True then show ?thesis
using face_of_affine_eq affine_UNIV by auto
next
case False
then have ine: "interior {x. a \<bullet> x \<le> b} \<noteq> {}"
using halfspace_eq_empty_lt interior_halfspace_le by blast
show ?thesis
proof
assume L: ?lhs
have "F \<noteq> {x. a \<bullet> x \<le> b} \<Longrightarrow> F face_of {x. a \<bullet> x = b}"
using False
apply (simp add: frontier_halfspace_le [symmetric] rel_frontier_nonempty_interior [OF ine, symmetric])
apply (rule face_of_subset [OF L])
apply (simp add: face_of_subset_rel_frontier [OF L])
apply (force simp: rel_frontier_def closed_halfspace_le)
done
with L show ?rhs
using affine_hyperplane face_of_affine_eq by blast
next
assume ?rhs
then show ?lhs
by (metis convex_halfspace_le empty_face_of face_of_refl hyperplane_face_of_halfspace_le)
qed
qed
lemma face_of_halfspace_ge:
fixes a :: "'n::euclidean_space"
shows "F face_of {x. a \<bullet> x \<ge> b} \<longleftrightarrow>
F = {} \<or> F = {x. a \<bullet> x = b} \<or> F = {x. a \<bullet> x \<ge> b}"
using face_of_halfspace_le [of F "-a" "-b"] by simp
subsection\<open>Exposed faces\<close>
text\<open>That is, faces that are intersection with supporting hyperplane\<close>
definition\<^marker>\<open>tag important\<close> exposed_face_of :: "['a::euclidean_space set, 'a set] \<Rightarrow> bool"
(infixr "(exposed'_face'_of)" 50)
where "T exposed_face_of S \<longleftrightarrow>
T face_of S \<and> (\<exists>a b. S \<subseteq> {x. a \<bullet> x \<le> b} \<and> T = S \<inter> {x. a \<bullet> x = b})"
lemma empty_exposed_face_of [iff]: "{} exposed_face_of S"
apply (simp add: exposed_face_of_def)
apply (rule_tac x=0 in exI)
apply (rule_tac x=1 in exI, force)
done
lemma exposed_face_of_refl_eq [simp]: "S exposed_face_of S \<longleftrightarrow> convex S"
apply (simp add: exposed_face_of_def face_of_refl_eq, auto)
apply (rule_tac x=0 in exI)+
apply force
done
lemma exposed_face_of_refl: "convex S \<Longrightarrow> S exposed_face_of S"
by simp
lemma exposed_face_of:
"T exposed_face_of S \<longleftrightarrow>
T face_of S \<and>
(T = {} \<or> T = S \<or>
(\<exists>a b. a \<noteq> 0 \<and> S \<subseteq> {x. a \<bullet> x \<le> b} \<and> T = S \<inter> {x. a \<bullet> x = b}))"
proof (cases "T = {}")
case True then show ?thesis
by simp
next
case False
show ?thesis
proof (cases "T = S")
case True then show ?thesis
by (simp add: face_of_refl_eq)
next
case False
with \<open>T \<noteq> {}\<close> show ?thesis
apply (auto simp: exposed_face_of_def)
apply (metis inner_zero_left)
done
qed
qed
lemma exposed_face_of_Int_supporting_hyperplane_le:
"\<lbrakk>convex S; \<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<le> b\<rbrakk> \<Longrightarrow> (S \<inter> {x. a \<bullet> x = b}) exposed_face_of S"
by (force simp: exposed_face_of_def face_of_Int_supporting_hyperplane_le)
lemma exposed_face_of_Int_supporting_hyperplane_ge:
"\<lbrakk>convex S; \<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<ge> b\<rbrakk> \<Longrightarrow> (S \<inter> {x. a \<bullet> x = b}) exposed_face_of S"
using exposed_face_of_Int_supporting_hyperplane_le [of S "-a" "-b"] by simp
proposition exposed_face_of_Int:
assumes "T exposed_face_of S"
and "u exposed_face_of S"
shows "(T \<inter> u) exposed_face_of S"
proof -
obtain a b where T: "S \<inter> {x. a \<bullet> x = b} face_of S"
and S: "S \<subseteq> {x. a \<bullet> x \<le> b}"
and teq: "T = S \<inter> {x. a \<bullet> x = b}"
using assms by (auto simp: exposed_face_of_def)
obtain a' b' where u: "S \<inter> {x. a' \<bullet> x = b'} face_of S"
and s': "S \<subseteq> {x. a' \<bullet> x \<le> b'}"
and ueq: "u = S \<inter> {x. a' \<bullet> x = b'}"
using assms by (auto simp: exposed_face_of_def)
have tu: "T \<inter> u face_of S"
using T teq u ueq by (simp add: face_of_Int)
have ss: "S \<subseteq> {x. (a + a') \<bullet> x \<le> b + b'}"
using S s' by (force simp: inner_left_distrib)
show ?thesis
apply (simp add: exposed_face_of_def tu)
apply (rule_tac x="a+a'" in exI)
apply (rule_tac x="b+b'" in exI)
using S s'
apply (fastforce simp: ss inner_left_distrib teq ueq)
done
qed
proposition exposed_face_of_Inter:
fixes P :: "'a::euclidean_space set set"
assumes "P \<noteq> {}"
and "\<And>T. T \<in> P \<Longrightarrow> T exposed_face_of S"
shows "\<Inter>P exposed_face_of S"
proof -
obtain Q where "finite Q" and QsubP: "Q \<subseteq> P" "card Q \<le> DIM('a) + 2" and IntQ: "\<Inter>Q = \<Inter>P"
using Inter_faces_finite_altbound [of P S] assms [unfolded exposed_face_of]
by force
show ?thesis
proof (cases "Q = {}")
case True then show ?thesis
by (metis IntQ Inter_UNIV_conv(2) assms(1) assms(2) ex_in_conv)
next
case False
have "Q \<subseteq> {T. T exposed_face_of S}"
using QsubP assms by blast
moreover have "Q \<subseteq> {T. T exposed_face_of S} \<Longrightarrow> \<Inter>Q exposed_face_of S"
using \<open>finite Q\<close> False
apply (induction Q rule: finite_induct)
using exposed_face_of_Int apply fastforce+
done
ultimately show ?thesis
by (simp add: IntQ)
qed
qed
proposition exposed_face_of_sums:
assumes "convex S" and "convex T"
and "F exposed_face_of {x + y | x y. x \<in> S \<and> y \<in> T}"
(is "F exposed_face_of ?ST")
obtains k l
where "k exposed_face_of S" "l exposed_face_of T"
"F = {x + y | x y. x \<in> k \<and> y \<in> l}"
proof (cases "F = {}")
case True then show ?thesis
using that by blast
next
case False
show ?thesis
proof (cases "F = ?ST")
case True then show ?thesis
using assms exposed_face_of_refl_eq that by blast
next
case False
obtain p where "p \<in> F" using \<open>F \<noteq> {}\<close> by blast
moreover
obtain u z where T: "?ST \<inter> {x. u \<bullet> x = z} face_of ?ST"
and S: "?ST \<subseteq> {x. u \<bullet> x \<le> z}"
and feq: "F = ?ST \<inter> {x. u \<bullet> x = z}"
using assms by (auto simp: exposed_face_of_def)
ultimately obtain a0 b0
where p: "p = a0 + b0" and "a0 \<in> S" "b0 \<in> T" and z: "u \<bullet> p = z"
by auto
have lez: "u \<bullet> (x + y) \<le> z" if "x \<in> S" "y \<in> T" for x y
using S that by auto
have sef: "S \<inter> {x. u \<bullet> x = u \<bullet> a0} exposed_face_of S"
apply (rule exposed_face_of_Int_supporting_hyperplane_le [OF \<open>convex S\<close>])
apply (metis p z add_le_cancel_right inner_right_distrib lez [OF _ \<open>b0 \<in> T\<close>])
done
have tef: "T \<inter> {x. u \<bullet> x = u \<bullet> b0} exposed_face_of T"
apply (rule exposed_face_of_Int_supporting_hyperplane_le [OF \<open>convex T\<close>])
apply (metis p z add.commute add_le_cancel_right inner_right_distrib lez [OF \<open>a0 \<in> S\<close>])
done
have "{x + y |x y. x \<in> S \<and> u \<bullet> x = u \<bullet> a0 \<and> y \<in> T \<and> u \<bullet> y = u \<bullet> b0} \<subseteq> F"
by (auto simp: feq) (metis inner_right_distrib p z)
moreover have "F \<subseteq> {x + y |x y. x \<in> S \<and> u \<bullet> x = u \<bullet> a0 \<and> y \<in> T \<and> u \<bullet> y = u \<bullet> b0}"
apply (auto simp: feq)
apply (rename_tac x y)
apply (rule_tac x=x in exI)
apply (rule_tac x=y in exI, simp)
using z p \<open>a0 \<in> S\<close> \<open>b0 \<in> T\<close>
apply clarify
apply (simp add: inner_right_distrib)
apply (metis add_le_cancel_right antisym lez [unfolded inner_right_distrib] add.commute)
done
ultimately have "F = {x + y |x y. x \<in> S \<inter> {x. u \<bullet> x = u \<bullet> a0} \<and> y \<in> T \<inter> {x. u \<bullet> x = u \<bullet> b0}}"
by blast
then show ?thesis
by (rule that [OF sef tef])
qed
qed
proposition exposed_face_of_parallel:
"T exposed_face_of S \<longleftrightarrow>
T face_of S \<and>
(\<exists>a b. S \<subseteq> {x. a \<bullet> x \<le> b} \<and> T = S \<inter> {x. a \<bullet> x = b} \<and>
(T \<noteq> {} \<longrightarrow> T \<noteq> S \<longrightarrow> a \<noteq> 0) \<and>
(T \<noteq> S \<longrightarrow> (\<forall>w \<in> affine hull S. (w + a) \<in> affine hull S)))"
(is "?lhs = ?rhs")
proof
assume ?lhs then show ?rhs
proof (clarsimp simp: exposed_face_of_def)
fix a b
assume faceS: "S \<inter> {x. a \<bullet> x = b} face_of S" and Ssub: "S \<subseteq> {x. a \<bullet> x \<le> b}"
show "\<exists>c d. S \<subseteq> {x. c \<bullet> x \<le> d} \<and>
S \<inter> {x. a \<bullet> x = b} = S \<inter> {x. c \<bullet> x = d} \<and>
(S \<inter> {x. a \<bullet> x = b} \<noteq> {} \<longrightarrow> S \<inter> {x. a \<bullet> x = b} \<noteq> S \<longrightarrow> c \<noteq> 0) \<and>
(S \<inter> {x. a \<bullet> x = b} \<noteq> S \<longrightarrow> (\<forall>w \<in> affine hull S. w + c \<in> affine hull S))"
proof (cases "affine hull S \<inter> {x. -a \<bullet> x \<le> -b} = {} \<or> affine hull S \<subseteq> {x. - a \<bullet> x \<le> - b}")
case True
then show ?thesis
proof
assume "affine hull S \<inter> {x. - a \<bullet> x \<le> - b} = {}"
then show ?thesis
apply (rule_tac x="0" in exI)
apply (rule_tac x="1" in exI)
using hull_subset by fastforce
next
assume "affine hull S \<subseteq> {x. - a \<bullet> x \<le> - b}"
then show ?thesis
apply (rule_tac x="0" in exI)
apply (rule_tac x="0" in exI)
using Ssub hull_subset by fastforce
qed
next
case False
then obtain a' b' where "a' \<noteq> 0"
and le: "affine hull S \<inter> {x. a' \<bullet> x \<le> b'} = affine hull S \<inter> {x. - a \<bullet> x \<le> - b}"
and eq: "affine hull S \<inter> {x. a' \<bullet> x = b'} = affine hull S \<inter> {x. - a \<bullet> x = - b}"
and mem: "\<And>w. w \<in> affine hull S \<Longrightarrow> w + a' \<in> affine hull S"
using affine_parallel_slice affine_affine_hull by metis
show ?thesis
proof (intro conjI impI allI ballI exI)
have *: "S \<subseteq> - (affine hull S \<inter> {x. P x}) \<union> affine hull S \<inter> {x. Q x} \<Longrightarrow> S \<subseteq> {x. \<not> P x \<or> Q x}"
for P Q
using hull_subset by fastforce
have "S \<subseteq> {x. \<not> (a' \<bullet> x \<le> b') \<or> a' \<bullet> x = b'}"
apply (rule *)
apply (simp only: le eq)
using Ssub by auto
then show "S \<subseteq> {x. - a' \<bullet> x \<le> - b'}"
by auto
show "S \<inter> {x. a \<bullet> x = b} = S \<inter> {x. - a' \<bullet> x = - b'}"
using eq hull_subset [of S affine] by force
show "\<lbrakk>S \<inter> {x. a \<bullet> x = b} \<noteq> {}; S \<inter> {x. a \<bullet> x = b} \<noteq> S\<rbrakk> \<Longrightarrow> - a' \<noteq> 0"
using \<open>a' \<noteq> 0\<close> by auto
show "w + - a' \<in> affine hull S"
if "S \<inter> {x. a \<bullet> x = b} \<noteq> S" "w \<in> affine hull S" for w
proof -
have "w + 1 *\<^sub>R (w - (w + a')) \<in> affine hull S"
using affine_affine_hull mem mem_affine_3_minus that(2) by blast
then show ?thesis by simp
qed
qed
qed
qed
next
assume ?rhs then show ?lhs
unfolding exposed_face_of_def by blast
qed
subsection\<open>Extreme points of a set: its singleton faces\<close>
definition\<^marker>\<open>tag important\<close> extreme_point_of :: "['a::real_vector, 'a set] \<Rightarrow> bool"
(infixr "(extreme'_point'_of)" 50)
where "x extreme_point_of S \<longleftrightarrow>
x \<in> S \<and> (\<forall>a \<in> S. \<forall>b \<in> S. x \<notin> open_segment a b)"
lemma extreme_point_of_stillconvex:
"convex S \<Longrightarrow> (x extreme_point_of S \<longleftrightarrow> x \<in> S \<and> convex(S - {x}))"
by (fastforce simp add: convex_contains_segment extreme_point_of_def open_segment_def)
lemma face_of_singleton:
"{x} face_of S \<longleftrightarrow> x extreme_point_of S"
by (fastforce simp add: extreme_point_of_def face_of_def)
lemma extreme_point_not_in_REL_INTERIOR:
fixes S :: "'a::real_normed_vector set"
shows "\<lbrakk>x extreme_point_of S; S \<noteq> {x}\<rbrakk> \<Longrightarrow> x \<notin> rel_interior S"
apply (simp add: face_of_singleton [symmetric])
apply (blast dest: face_of_disjoint_rel_interior)
done
lemma extreme_point_not_in_interior:
fixes S :: "'a::{real_normed_vector, perfect_space} set"
shows "x extreme_point_of S \<Longrightarrow> x \<notin> interior S"
apply (case_tac "S = {x}")
apply (simp add: empty_interior_finite)
by (meson contra_subsetD extreme_point_not_in_REL_INTERIOR interior_subset_rel_interior)
lemma extreme_point_of_face:
"F face_of S \<Longrightarrow> v extreme_point_of F \<longleftrightarrow> v extreme_point_of S \<and> v \<in> F"
by (meson empty_subsetI face_of_face face_of_singleton insert_subset)
lemma extreme_point_of_convex_hull:
"x extreme_point_of (convex hull S) \<Longrightarrow> x \<in> S"
apply (simp add: extreme_point_of_stillconvex)
using hull_minimal [of S "(convex hull S) - {x}" convex]
using hull_subset [of S convex]
apply blast
done
proposition extreme_points_of_convex_hull:
"{x. x extreme_point_of (convex hull S)} \<subseteq> S"
using extreme_point_of_convex_hull by auto
lemma extreme_point_of_empty [simp]: "\<not> (x extreme_point_of {})"
by (simp add: extreme_point_of_def)
lemma extreme_point_of_singleton [iff]: "x extreme_point_of {a} \<longleftrightarrow> x = a"
using extreme_point_of_stillconvex by auto
lemma extreme_point_of_translation_eq:
"(a + x) extreme_point_of (image (\<lambda>x. a + x) S) \<longleftrightarrow> x extreme_point_of S"
by (auto simp: extreme_point_of_def)
lemma extreme_points_of_translation:
"{x. x extreme_point_of (image (\<lambda>x. a + x) S)} =
(\<lambda>x. a + x) ` {x. x extreme_point_of S}"
using extreme_point_of_translation_eq
by auto (metis (no_types, lifting) image_iff mem_Collect_eq minus_add_cancel)
lemma extreme_points_of_translation_subtract:
"{x. x extreme_point_of (image (\<lambda>x. x - a) S)} =
(\<lambda>x. x - a) ` {x. x extreme_point_of S}"
using extreme_points_of_translation [of "- a" S]
by simp
lemma extreme_point_of_Int:
"\<lbrakk>x extreme_point_of S; x extreme_point_of T\<rbrakk> \<Longrightarrow> x extreme_point_of (S \<inter> T)"
by (simp add: extreme_point_of_def)
lemma extreme_point_of_Int_supporting_hyperplane_le:
"\<lbrakk>S \<inter> {x. a \<bullet> x = b} = {c}; \<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<le> b\<rbrakk> \<Longrightarrow> c extreme_point_of S"
apply (simp add: face_of_singleton [symmetric])
by (metis face_of_Int_supporting_hyperplane_le_strong convex_singleton)
lemma extreme_point_of_Int_supporting_hyperplane_ge:
"\<lbrakk>S \<inter> {x. a \<bullet> x = b} = {c}; \<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<ge> b\<rbrakk> \<Longrightarrow> c extreme_point_of S"
apply (simp add: face_of_singleton [symmetric])
by (metis face_of_Int_supporting_hyperplane_ge_strong convex_singleton)
lemma exposed_point_of_Int_supporting_hyperplane_le:
"\<lbrakk>S \<inter> {x. a \<bullet> x = b} = {c}; \<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<le> b\<rbrakk> \<Longrightarrow> {c} exposed_face_of S"
apply (simp add: exposed_face_of_def face_of_singleton)
apply (force simp: extreme_point_of_Int_supporting_hyperplane_le)
done
lemma exposed_point_of_Int_supporting_hyperplane_ge:
"\<lbrakk>S \<inter> {x. a \<bullet> x = b} = {c}; \<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<ge> b\<rbrakk> \<Longrightarrow> {c} exposed_face_of S"
using exposed_point_of_Int_supporting_hyperplane_le [of S "-a" "-b" c]
by simp
lemma extreme_point_of_convex_hull_insert:
"\<lbrakk>finite S; a \<notin> convex hull S\<rbrakk> \<Longrightarrow> a extreme_point_of (convex hull (insert a S))"
apply (case_tac "a \<in> S")
apply (simp add: hull_inc)
using face_of_convex_hulls [of "insert a S" "{a}"]
apply (auto simp: face_of_singleton hull_same)
done
subsection\<open>Facets\<close>
definition\<^marker>\<open>tag important\<close> facet_of :: "['a::euclidean_space set, 'a set] \<Rightarrow> bool"
(infixr "(facet'_of)" 50)
where "F facet_of S \<longleftrightarrow> F face_of S \<and> F \<noteq> {} \<and> aff_dim F = aff_dim S - 1"
lemma facet_of_empty [simp]: "\<not> S facet_of {}"
by (simp add: facet_of_def)
lemma facet_of_irrefl [simp]: "\<not> S facet_of S "
by (simp add: facet_of_def)
lemma facet_of_imp_face_of: "F facet_of S \<Longrightarrow> F face_of S"
by (simp add: facet_of_def)
lemma facet_of_imp_subset: "F facet_of S \<Longrightarrow> F \<subseteq> S"
by (simp add: face_of_imp_subset facet_of_def)
lemma hyperplane_facet_of_halfspace_le:
"a \<noteq> 0 \<Longrightarrow> {x. a \<bullet> x = b} facet_of {x. a \<bullet> x \<le> b}"
unfolding facet_of_def hyperplane_eq_empty
by (auto simp: hyperplane_face_of_halfspace_ge hyperplane_face_of_halfspace_le
- DIM_positive Suc_leI of_nat_diff aff_dim_halfspace_le)
+ Suc_leI of_nat_diff aff_dim_halfspace_le)
lemma hyperplane_facet_of_halfspace_ge:
"a \<noteq> 0 \<Longrightarrow> {x. a \<bullet> x = b} facet_of {x. a \<bullet> x \<ge> b}"
unfolding facet_of_def hyperplane_eq_empty
by (auto simp: hyperplane_face_of_halfspace_le hyperplane_face_of_halfspace_ge
- DIM_positive Suc_leI of_nat_diff aff_dim_halfspace_ge)
+ Suc_leI of_nat_diff aff_dim_halfspace_ge)
lemma facet_of_halfspace_le:
"F facet_of {x. a \<bullet> x \<le> b} \<longleftrightarrow> a \<noteq> 0 \<and> F = {x. a \<bullet> x = b}"
(is "?lhs = ?rhs")
proof
assume c: ?lhs
with c facet_of_irrefl show ?rhs
by (force simp: aff_dim_halfspace_le facet_of_def face_of_halfspace_le cong: conj_cong split: if_split_asm)
next
assume ?rhs then show ?lhs
by (simp add: hyperplane_facet_of_halfspace_le)
qed
lemma facet_of_halfspace_ge:
"F facet_of {x. a \<bullet> x \<ge> b} \<longleftrightarrow> a \<noteq> 0 \<and> F = {x. a \<bullet> x = b}"
using facet_of_halfspace_le [of F "-a" "-b"] by simp
subsection \<open>Edges: faces of affine dimension 1\<close> (*FIXME too small subsection, rearrange? *)
definition\<^marker>\<open>tag important\<close> edge_of :: "['a::euclidean_space set, 'a set] \<Rightarrow> bool" (infixr "(edge'_of)" 50)
where "e edge_of S \<longleftrightarrow> e face_of S \<and> aff_dim e = 1"
lemma edge_of_imp_subset:
"S edge_of T \<Longrightarrow> S \<subseteq> T"
by (simp add: edge_of_def face_of_imp_subset)
subsection\<open>Existence of extreme points\<close>
proposition different_norm_3_collinear_points:
fixes a :: "'a::euclidean_space"
assumes "x \<in> open_segment a b" "norm(a) = norm(b)" "norm(x) = norm(b)"
shows False
proof -
obtain u where "norm ((1 - u) *\<^sub>R a + u *\<^sub>R b) = norm b"
and "a \<noteq> b"
and u01: "0 < u" "u < 1"
using assms by (auto simp: open_segment_image_interval if_splits)
then have "(1 - u) *\<^sub>R a \<bullet> (1 - u) *\<^sub>R a + ((1 - u) * 2) *\<^sub>R a \<bullet> u *\<^sub>R b =
(1 - u * u) *\<^sub>R (a \<bullet> a)"
using assms by (simp add: norm_eq algebra_simps inner_commute)
then have "(1 - u) *\<^sub>R ((1 - u) *\<^sub>R a \<bullet> a + (2 * u) *\<^sub>R a \<bullet> b) =
(1 - u) *\<^sub>R ((1 + u) *\<^sub>R (a \<bullet> a))"
by (simp add: algebra_simps)
then have "(1 - u) *\<^sub>R (a \<bullet> a) + (2 * u) *\<^sub>R (a \<bullet> b) = (1 + u) *\<^sub>R (a \<bullet> a)"
using u01 by auto
then have "a \<bullet> b = a \<bullet> a"
using u01 by (simp add: algebra_simps)
then have "a = b"
using \<open>norm(a) = norm(b)\<close> norm_eq vector_eq by fastforce
then show ?thesis
using \<open>a \<noteq> b\<close> by force
qed
proposition extreme_point_exists_convex:
fixes S :: "'a::euclidean_space set"
assumes "compact S" "convex S" "S \<noteq> {}"
obtains x where "x extreme_point_of S"
proof -
obtain x where "x \<in> S" and xsup: "\<And>y. y \<in> S \<Longrightarrow> norm y \<le> norm x"
using distance_attains_sup [of S 0] assms by auto
have False if "a \<in> S" "b \<in> S" and x: "x \<in> open_segment a b" for a b
proof -
have noax: "norm a \<le> norm x" and nobx: "norm b \<le> norm x" using xsup that by auto
have "a \<noteq> b"
using empty_iff open_segment_idem x by auto
have *: "(1 - u) * na + u * nb < norm x" if "na < norm x" "nb \<le> norm x" "0 < u" "u < 1" for na nb u
proof -
have "(1 - u) * na + u * nb < (1 - u) * norm x + u * nb"
by (simp add: that)
also have "... \<le> (1 - u) * norm x + u * norm x"
by (simp add: that)
finally have "(1 - u) * na + u * nb < (1 - u) * norm x + u * norm x" .
then show ?thesis
using scaleR_collapse [symmetric, of "norm x" u] by auto
qed
have "norm x < norm x" if "norm a < norm x"
using x
apply (clarsimp simp only: open_segment_image_interval \<open>a \<noteq> b\<close> if_False)
apply (rule norm_triangle_lt)
apply (simp add: norm_mult)
using * [of "norm a" "norm b"] nobx that
apply blast
done
moreover have "norm x < norm x" if "norm b < norm x"
using x
apply (clarsimp simp only: open_segment_image_interval \<open>a \<noteq> b\<close> if_False)
apply (rule norm_triangle_lt)
apply (simp add: norm_mult)
using * [of "norm b" "norm a" "1-u" for u] noax that
apply (simp add: add.commute)
done
ultimately have "\<not> (norm a < norm x) \<and> \<not> (norm b < norm x)"
by auto
then show ?thesis
using different_norm_3_collinear_points noax nobx that(3) by fastforce
qed
then show ?thesis
apply (rule_tac x=x in that)
apply (force simp: extreme_point_of_def \<open>x \<in> S\<close>)
done
qed
subsection\<open>Krein-Milman, the weaker form\<close>
proposition Krein_Milman:
fixes S :: "'a::euclidean_space set"
assumes "compact S" "convex S"
shows "S = closure(convex hull {x. x extreme_point_of S})"
proof (cases "S = {}")
case True then show ?thesis by simp
next
case False
have "closed S"
by (simp add: \<open>compact S\<close> compact_imp_closed)
have "closure (convex hull {x. x extreme_point_of S}) \<subseteq> S"
apply (rule closure_minimal [OF hull_minimal \<open>closed S\<close>])
using assms
apply (auto simp: extreme_point_of_def)
done
moreover have "u \<in> closure (convex hull {x. x extreme_point_of S})"
if "u \<in> S" for u
proof (rule ccontr)
assume unot: "u \<notin> closure(convex hull {x. x extreme_point_of S})"
then obtain a b where "a \<bullet> u < b"
and ab: "\<And>x. x \<in> closure(convex hull {x. x extreme_point_of S}) \<Longrightarrow> b < a \<bullet> x"
using separating_hyperplane_closed_point [of "closure(convex hull {x. x extreme_point_of S})"]
by blast
have "continuous_on S ((\<bullet>) a)"
by (rule continuous_intros)+
then obtain m where "m \<in> S" and m: "\<And>y. y \<in> S \<Longrightarrow> a \<bullet> m \<le> a \<bullet> y"
using continuous_attains_inf [of S "\<lambda>x. a \<bullet> x"] \<open>compact S\<close> \<open>u \<in> S\<close>
by auto
define T where "T = S \<inter> {x. a \<bullet> x = a \<bullet> m}"
have "m \<in> T"
by (simp add: T_def \<open>m \<in> S\<close>)
moreover have "compact T"
by (simp add: T_def compact_Int_closed [OF \<open>compact S\<close> closed_hyperplane])
moreover have "convex T"
by (simp add: T_def convex_Int [OF \<open>convex S\<close> convex_hyperplane])
ultimately obtain v where v: "v extreme_point_of T"
using extreme_point_exists_convex [of T] by auto
then have "{v} face_of T"
by (simp add: face_of_singleton)
also have "T face_of S"
by (simp add: T_def m face_of_Int_supporting_hyperplane_ge [OF \<open>convex S\<close>])
finally have "v extreme_point_of S"
by (simp add: face_of_singleton)
then have "b < a \<bullet> v"
using closure_subset by (simp add: closure_hull hull_inc ab)
then show False
using \<open>a \<bullet> u < b\<close> \<open>{v} face_of T\<close> face_of_imp_subset m T_def that by fastforce
qed
ultimately show ?thesis
by blast
qed
text\<open>Now the sharper form.\<close>
lemma Krein_Milman_Minkowski_aux:
fixes S :: "'a::euclidean_space set"
assumes n: "dim S = n" and S: "compact S" "convex S" "0 \<in> S"
shows "0 \<in> convex hull {x. x extreme_point_of S}"
using n S
proof (induction n arbitrary: S rule: less_induct)
case (less n S) show ?case
proof (cases "0 \<in> rel_interior S")
case True with Krein_Milman show ?thesis
by (metis subsetD convex_convex_hull convex_rel_interior_closure less.prems(2) less.prems(3) rel_interior_subset)
next
case False
have "rel_interior S \<noteq> {}"
by (simp add: rel_interior_convex_nonempty_aux less)
then obtain c where c: "c \<in> rel_interior S" by blast
obtain a where "a \<noteq> 0"
and le_ay: "\<And>y. y \<in> S \<Longrightarrow> a \<bullet> 0 \<le> a \<bullet> y"
and less_ay: "\<And>y. y \<in> rel_interior S \<Longrightarrow> a \<bullet> 0 < a \<bullet> y"
by (blast intro: supporting_hyperplane_rel_boundary intro!: less False)
have face: "S \<inter> {x. a \<bullet> x = 0} face_of S"
apply (rule face_of_Int_supporting_hyperplane_ge [OF \<open>convex S\<close>])
using le_ay by auto
then have co: "compact (S \<inter> {x. a \<bullet> x = 0})" "convex (S \<inter> {x. a \<bullet> x = 0})"
using less.prems by (blast intro: face_of_imp_compact face_of_imp_convex)+
have "a \<bullet> y = 0" if "y \<in> span (S \<inter> {x. a \<bullet> x = 0})" for y
proof -
have "y \<in> span {x. a \<bullet> x = 0}"
by (metis inf.cobounded2 span_mono subsetCE that)
then show ?thesis
by (blast intro: span_induct [OF _ subspace_hyperplane])
qed
then have "dim (S \<inter> {x. a \<bullet> x = 0}) < n"
by (metis (no_types) less_ay c subsetD dim_eq_span inf.strict_order_iff
inf_le1 \<open>dim S = n\<close> not_le rel_interior_subset span_0 span_base)
then have "0 \<in> convex hull {x. x extreme_point_of (S \<inter> {x. a \<bullet> x = 0})}"
by (rule less.IH) (auto simp: co less.prems)
then show ?thesis
by (metis (mono_tags, lifting) Collect_mono_iff \<open>S \<inter> {x. a \<bullet> x = 0} face_of S\<close> extreme_point_of_face hull_mono subset_iff)
qed
qed
theorem Krein_Milman_Minkowski:
fixes S :: "'a::euclidean_space set"
assumes "compact S" "convex S"
shows "S = convex hull {x. x extreme_point_of S}"
proof
show "S \<subseteq> convex hull {x. x extreme_point_of S}"
proof
fix a assume [simp]: "a \<in> S"
have 1: "compact ((+) (- a) ` S)"
by (simp add: \<open>compact S\<close> compact_translation_subtract cong: image_cong_simp)
have 2: "convex ((+) (- a) ` S)"
by (simp add: \<open>convex S\<close> compact_translation_subtract)
show a_invex: "a \<in> convex hull {x. x extreme_point_of S}"
using Krein_Milman_Minkowski_aux [OF refl 1 2]
convex_hull_translation [of "-a"]
by (auto simp: extreme_points_of_translation_subtract translation_assoc cong: image_cong_simp)
qed
next
show "convex hull {x. x extreme_point_of S} \<subseteq> S"
proof -
have "{a. a extreme_point_of S} \<subseteq> S"
using extreme_point_of_def by blast
then show ?thesis
by (simp add: \<open>convex S\<close> hull_minimal)
qed
qed
subsection\<open>Applying it to convex hulls of explicitly indicated finite sets\<close>
corollary Krein_Milman_polytope:
fixes S :: "'a::euclidean_space set"
shows
"finite S
\<Longrightarrow> convex hull S =
convex hull {x. x extreme_point_of (convex hull S)}"
by (simp add: Krein_Milman_Minkowski finite_imp_compact_convex_hull)
lemma extreme_points_of_convex_hull_eq:
fixes S :: "'a::euclidean_space set"
shows
"\<lbrakk>compact S; \<And>T. T \<subset> S \<Longrightarrow> convex hull T \<noteq> convex hull S\<rbrakk>
\<Longrightarrow> {x. x extreme_point_of (convex hull S)} = S"
by (metis (full_types) Krein_Milman_Minkowski compact_convex_hull convex_convex_hull extreme_points_of_convex_hull psubsetI)
lemma extreme_point_of_convex_hull_eq:
fixes S :: "'a::euclidean_space set"
shows
"\<lbrakk>compact S; \<And>T. T \<subset> S \<Longrightarrow> convex hull T \<noteq> convex hull S\<rbrakk>
\<Longrightarrow> (x extreme_point_of (convex hull S) \<longleftrightarrow> x \<in> S)"
using extreme_points_of_convex_hull_eq by auto
lemma extreme_point_of_convex_hull_convex_independent:
fixes S :: "'a::euclidean_space set"
assumes "compact S" and S: "\<And>a. a \<in> S \<Longrightarrow> a \<notin> convex hull (S - {a})"
shows "(x extreme_point_of (convex hull S) \<longleftrightarrow> x \<in> S)"
proof -
have "convex hull T \<noteq> convex hull S" if "T \<subset> S" for T
proof -
obtain a where "T \<subseteq> S" "a \<in> S" "a \<notin> T" using \<open>T \<subset> S\<close> by blast
then show ?thesis
by (metis (full_types) Diff_eq_empty_iff Diff_insert0 S hull_mono hull_subset insert_Diff_single subsetCE)
qed
then show ?thesis
by (rule extreme_point_of_convex_hull_eq [OF \<open>compact S\<close>])
qed
lemma extreme_point_of_convex_hull_affine_independent:
fixes S :: "'a::euclidean_space set"
shows
"\<not> affine_dependent S
\<Longrightarrow> (x extreme_point_of (convex hull S) \<longleftrightarrow> x \<in> S)"
by (metis aff_independent_finite affine_dependent_def affine_hull_convex_hull extreme_point_of_convex_hull_convex_independent finite_imp_compact hull_inc)
text\<open>Elementary proofs exist, not requiring Euclidean spaces and all this development\<close>
lemma extreme_point_of_convex_hull_2:
fixes x :: "'a::euclidean_space"
shows "x extreme_point_of (convex hull {a,b}) \<longleftrightarrow> x = a \<or> x = b"
proof -
have "x extreme_point_of (convex hull {a,b}) \<longleftrightarrow> x \<in> {a,b}"
by (intro extreme_point_of_convex_hull_affine_independent affine_independent_2)
then show ?thesis
by simp
qed
lemma extreme_point_of_segment:
fixes x :: "'a::euclidean_space"
shows
"x extreme_point_of closed_segment a b \<longleftrightarrow> x = a \<or> x = b"
by (simp add: extreme_point_of_convex_hull_2 segment_convex_hull)
lemma face_of_convex_hull_subset:
fixes S :: "'a::euclidean_space set"
assumes "compact S" and T: "T face_of (convex hull S)"
obtains s' where "s' \<subseteq> S" "T = convex hull s'"
apply (rule_tac s' = "{x. x extreme_point_of T}" in that)
using T extreme_point_of_convex_hull extreme_point_of_face apply blast
by (metis (no_types) Krein_Milman_Minkowski assms compact_convex_hull convex_convex_hull face_of_imp_compact face_of_imp_convex)
lemma face_of_convex_hull_aux:
assumes eq: "x *\<^sub>R p = u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c"
and x: "u + v + w = x" "x \<noteq> 0" and S: "affine S" "a \<in> S" "b \<in> S" "c \<in> S"
shows "p \<in> S"
proof -
have "p = (u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c) /\<^sub>R x"
by (metis \<open>x \<noteq> 0\<close> eq mult.commute right_inverse scaleR_one scaleR_scaleR)
moreover have "affine hull {a,b,c} \<subseteq> S"
by (simp add: S hull_minimal)
moreover have "(u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c) /\<^sub>R x \<in> affine hull {a,b,c}"
apply (simp add: affine_hull_3)
apply (rule_tac x="u/x" in exI)
apply (rule_tac x="v/x" in exI)
apply (rule_tac x="w/x" in exI)
- using x apply (auto simp: algebra_simps field_split_simps)
+ using x apply (auto simp: field_split_simps)
done
ultimately show ?thesis by force
qed
proposition face_of_convex_hull_insert_eq:
fixes a :: "'a :: euclidean_space"
assumes "finite S" and a: "a \<notin> affine hull S"
shows "(F face_of (convex hull (insert a S)) \<longleftrightarrow>
F face_of (convex hull S) \<or>
(\<exists>F'. F' face_of (convex hull S) \<and> F = convex hull (insert a F')))"
(is "F face_of ?CAS \<longleftrightarrow> _")
proof safe
assume F: "F face_of ?CAS"
and *: "\<nexists>F'. F' face_of convex hull S \<and> F = convex hull insert a F'"
obtain T where T: "T \<subseteq> insert a S" and FeqT: "F = convex hull T"
by (metis F \<open>finite S\<close> compact_insert finite_imp_compact face_of_convex_hull_subset)
show "F face_of convex hull S"
proof (cases "a \<in> T")
case True
have "F = convex hull insert a (convex hull T \<inter> convex hull S)"
proof
have "T \<subseteq> insert a (convex hull T \<inter> convex hull S)"
using T hull_subset by fastforce
then show "F \<subseteq> convex hull insert a (convex hull T \<inter> convex hull S)"
by (simp add: FeqT hull_mono)
show "convex hull insert a (convex hull T \<inter> convex hull S) \<subseteq> F"
apply (rule hull_minimal)
using True by (auto simp: \<open>F = convex hull T\<close> hull_inc)
qed
moreover have "convex hull T \<inter> convex hull S face_of convex hull S"
by (metis F FeqT convex_convex_hull face_of_slice hull_mono inf.absorb_iff2 subset_insertI)
ultimately show ?thesis
using * by force
next
case False
then show ?thesis
by (metis FeqT F T face_of_subset hull_mono subset_insert subset_insertI)
qed
next
assume "F face_of convex hull S"
show "F face_of ?CAS"
by (simp add: \<open>F face_of convex hull S\<close> a face_of_convex_hull_insert \<open>finite S\<close>)
next
fix F
assume F: "F face_of convex hull S"
show "convex hull insert a F face_of ?CAS"
proof (cases "S = {}")
case True
then show ?thesis
using F face_of_affine_eq by auto
next
case False
have anotc: "a \<notin> convex hull S"
by (metis (no_types) a affine_hull_convex_hull hull_inc)
show ?thesis
proof (cases "F = {}")
case True show ?thesis
using anotc by (simp add: \<open>F = {}\<close> \<open>finite S\<close> extreme_point_of_convex_hull_insert face_of_singleton)
next
case False
have "convex hull insert a F \<subseteq> ?CAS"
by (simp add: F a \<open>finite S\<close> convex_hull_subset face_of_convex_hull_insert face_of_imp_subset hull_inc)
moreover
have "(\<exists>y v. (1 - ub) *\<^sub>R a + ub *\<^sub>R b = (1 - v) *\<^sub>R a + v *\<^sub>R y \<and>
0 \<le> v \<and> v \<le> 1 \<and> y \<in> F) \<and>
(\<exists>x u. (1 - uc) *\<^sub>R a + uc *\<^sub>R c = (1 - u) *\<^sub>R a + u *\<^sub>R x \<and>
0 \<le> u \<and> u \<le> 1 \<and> x \<in> F)"
if *: "(1 - ux) *\<^sub>R a + ux *\<^sub>R x
\<in> open_segment ((1 - ub) *\<^sub>R a + ub *\<^sub>R b) ((1 - uc) *\<^sub>R a + uc *\<^sub>R c)"
and "0 \<le> ub" "ub \<le> 1" "0 \<le> uc" "uc \<le> 1" "0 \<le> ux" "ux \<le> 1"
and b: "b \<in> convex hull S" and c: "c \<in> convex hull S" and "x \<in> F"
for b c ub uc ux x
proof -
obtain v where ne: "(1 - ub) *\<^sub>R a + ub *\<^sub>R b \<noteq> (1 - uc) *\<^sub>R a + uc *\<^sub>R c"
and eq: "(1 - ux) *\<^sub>R a + ux *\<^sub>R x =
(1 - v) *\<^sub>R ((1 - ub) *\<^sub>R a + ub *\<^sub>R b) + v *\<^sub>R ((1 - uc) *\<^sub>R a + uc *\<^sub>R c)"
and "0 < v" "v < 1"
using * by (auto simp: in_segment)
then have 0: "((1 - ux) - ((1 - v) * (1 - ub) + v * (1 - uc))) *\<^sub>R a +
(ux *\<^sub>R x - (((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c)) = 0"
by (auto simp: algebra_simps)
then have "((1 - ux) - ((1 - v) * (1 - ub) + v * (1 - uc))) *\<^sub>R a =
((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c + (-ux) *\<^sub>R x"
by (auto simp: algebra_simps)
then have "a \<in> affine hull S" if "1 - ux - ((1 - v) * (1 - ub) + v * (1 - uc)) \<noteq> 0"
apply (rule face_of_convex_hull_aux)
using b c that apply (auto simp: algebra_simps)
using F convex_hull_subset_affine_hull face_of_imp_subset \<open>x \<in> F\<close> apply blast+
done
then have "1 - ux - ((1 - v) * (1 - ub) + v * (1 - uc)) = 0"
using a by blast
with 0 have equx: "(1 - v) * ub + v * uc = ux"
and uxx: "ux *\<^sub>R x = (((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c)"
by auto (auto simp: algebra_simps)
show ?thesis
proof (cases "uc = 0")
case True
then show ?thesis
using equx 0 \<open>0 \<le> ub\<close> \<open>ub \<le> 1\<close> \<open>v < 1\<close> \<open>x \<in> F\<close>
apply (auto simp: algebra_simps)
apply (rule_tac x=x in exI, simp)
apply (rule_tac x=ub in exI, auto)
apply (metis add.left_neutral diff_eq_eq less_irrefl mult.commute mult_cancel_right1 real_vector.scale_cancel_left real_vector.scale_left_diff_distrib)
using \<open>x \<in> F\<close> \<open>uc \<le> 1\<close> apply blast
done
next
case False
show ?thesis
proof (cases "ub = 0")
case True
then show ?thesis
using equx 0 \<open>0 \<le> uc\<close> \<open>uc \<le> 1\<close> \<open>0 < v\<close> \<open>x \<in> F\<close> \<open>uc \<noteq> 0\<close> by (force simp: algebra_simps)
next
case False
then have "0 < ub" "0 < uc"
using \<open>uc \<noteq> 0\<close> \<open>0 \<le> ub\<close> \<open>0 \<le> uc\<close> by auto
then have "ux \<noteq> 0"
by (metis \<open>0 < v\<close> \<open>v < 1\<close> diff_ge_0_iff_ge dual_order.strict_implies_order equx leD le_add_same_cancel2 zero_le_mult_iff zero_less_mult_iff)
have "b \<in> F \<and> c \<in> F"
proof (cases "b = c")
case True
then show ?thesis
by (metis \<open>ux \<noteq> 0\<close> equx real_vector.scale_cancel_left scaleR_add_left uxx \<open>x \<in> F\<close>)
next
case False
have "x = (((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c) /\<^sub>R ux"
by (metis \<open>ux \<noteq> 0\<close> uxx mult.commute right_inverse scaleR_one scaleR_scaleR)
also have "... = (1 - v * uc / ux) *\<^sub>R b + (v * uc / ux) *\<^sub>R c"
- using \<open>ux \<noteq> 0\<close> equx apply (auto simp: algebra_simps field_split_simps)
+ using \<open>ux \<noteq> 0\<close> equx apply (auto simp: field_split_simps)
by (metis add.commute add_diff_eq add_divide_distrib diff_add_cancel scaleR_add_left)
finally have "x = (1 - v * uc / ux) *\<^sub>R b + (v * uc / ux) *\<^sub>R c" .
then have "x \<in> open_segment b c"
apply (simp add: in_segment \<open>b \<noteq> c\<close>)
apply (rule_tac x="(v * uc) / ux" in exI)
using \<open>0 \<le> ux\<close> \<open>ux \<noteq> 0\<close> \<open>0 < uc\<close> \<open>0 < v\<close> \<open>0 < ub\<close> \<open>v < 1\<close> equx
- apply (force simp: algebra_simps field_split_simps)
+ apply (force simp: field_split_simps)
done
then show ?thesis
by (rule face_ofD [OF F _ b c \<open>x \<in> F\<close>])
qed
with \<open>0 \<le> ub\<close> \<open>ub \<le> 1\<close> \<open>0 \<le> uc\<close> \<open>uc \<le> 1\<close> show ?thesis by blast
qed
qed
qed
moreover have "convex hull F = F"
by (meson F convex_hull_eq face_of_imp_convex)
ultimately show ?thesis
unfolding face_of_def by (fastforce simp: convex_hull_insert_alt \<open>S \<noteq> {}\<close> \<open>F \<noteq> {}\<close>)
qed
qed
qed
lemma face_of_convex_hull_insert2:
fixes a :: "'a :: euclidean_space"
assumes S: "finite S" and a: "a \<notin> affine hull S" and F: "F face_of convex hull S"
shows "convex hull (insert a F) face_of convex hull (insert a S)"
by (metis F face_of_convex_hull_insert_eq [OF S a])
proposition face_of_convex_hull_affine_independent:
fixes S :: "'a::euclidean_space set"
assumes "\<not> affine_dependent S"
shows "(T face_of (convex hull S) \<longleftrightarrow> (\<exists>c. c \<subseteq> S \<and> T = convex hull c))"
(is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
by (meson \<open>T face_of convex hull S\<close> aff_independent_finite assms face_of_convex_hull_subset finite_imp_compact)
next
assume ?rhs
then obtain c where "c \<subseteq> S" and T: "T = convex hull c"
by blast
have "affine hull c \<inter> affine hull (S - c) = {}"
apply (rule disjoint_affine_hull [OF assms \<open>c \<subseteq> S\<close>], auto)
done
then have "affine hull c \<inter> convex hull (S - c) = {}"
using convex_hull_subset_affine_hull by fastforce
then show ?lhs
by (metis face_of_convex_hulls \<open>c \<subseteq> S\<close> aff_independent_finite assms T)
qed
lemma facet_of_convex_hull_affine_independent:
fixes S :: "'a::euclidean_space set"
assumes "\<not> affine_dependent S"
shows "T facet_of (convex hull S) \<longleftrightarrow>
T \<noteq> {} \<and> (\<exists>u. u \<in> S \<and> T = convex hull (S - {u}))"
(is "?lhs = ?rhs")
proof
assume ?lhs
then have "T face_of (convex hull S)" "T \<noteq> {}"
and afft: "aff_dim T = aff_dim (convex hull S) - 1"
by (auto simp: facet_of_def)
then obtain c where "c \<subseteq> S" and c: "T = convex hull c"
by (auto simp: face_of_convex_hull_affine_independent [OF assms])
then have affs: "aff_dim S = aff_dim c + 1"
by (metis aff_dim_convex_hull afft eq_diff_eq)
have "\<not> affine_dependent c"
using \<open>c \<subseteq> S\<close> affine_dependent_subset assms by blast
with affs have "card (S - c) = 1"
apply (simp add: aff_dim_affine_independent [symmetric] aff_dim_convex_hull)
by (metis aff_dim_affine_independent aff_independent_finite One_nat_def \<open>c \<subseteq> S\<close> add.commute
add_diff_cancel_right' assms card_Diff_subset card_mono of_nat_1 of_nat_diff of_nat_eq_iff)
then obtain u where u: "u \<in> S - c"
by (metis DiffI \<open>c \<subseteq> S\<close> aff_independent_finite assms cancel_comm_monoid_add_class.diff_cancel
card_Diff_subset subsetI subset_antisym zero_neq_one)
then have u: "S = insert u c"
by (metis Diff_subset \<open>c \<subseteq> S\<close> \<open>card (S - c) = 1\<close> card_1_singletonE double_diff insert_Diff insert_subset singletonD)
have "T = convex hull (c - {u})"
by (metis Diff_empty Diff_insert0 \<open>T facet_of convex hull S\<close> c facet_of_irrefl insert_absorb u)
with \<open>T \<noteq> {}\<close> show ?rhs
using c u by auto
next
assume ?rhs
then obtain u where "T \<noteq> {}" "u \<in> S" and u: "T = convex hull (S - {u})"
by (force simp: facet_of_def)
then have "\<not> S \<subseteq> {u}"
using \<open>T \<noteq> {}\<close> u by auto
have [simp]: "aff_dim (convex hull (S - {u})) = aff_dim (convex hull S) - 1"
using assms \<open>u \<in> S\<close>
apply (simp add: aff_dim_convex_hull affine_dependent_def)
apply (drule bspec, assumption)
by (metis add_diff_cancel_right' aff_dim_insert insert_Diff [of u S])
show ?lhs
apply (subst u)
apply (simp add: \<open>\<not> S \<subseteq> {u}\<close> facet_of_def face_of_convex_hull_affine_independent [OF assms], blast)
done
qed
lemma facet_of_convex_hull_affine_independent_alt:
fixes S :: "'a::euclidean_space set"
shows
"\<not>affine_dependent S
\<Longrightarrow> (T facet_of (convex hull S) \<longleftrightarrow>
2 \<le> card S \<and> (\<exists>u. u \<in> S \<and> T = convex hull (S - {u})))"
apply (simp add: facet_of_convex_hull_affine_independent)
apply (auto simp: Set.subset_singleton_iff)
apply (metis Diff_cancel Int_empty_right Int_insert_right_if1 aff_independent_finite card_eq_0_iff card_insert_if card_mono card_subset_eq convex_hull_eq_empty eq_iff equals0D finite_insert finite_subset inf.absorb_iff2 insert_absorb insert_not_empty not_less_eq_eq numeral_2_eq_2)
done
lemma segment_face_of:
assumes "(closed_segment a b) face_of S"
shows "a extreme_point_of S" "b extreme_point_of S"
proof -
have as: "{a} face_of S"
by (metis (no_types) assms convex_hull_singleton empty_iff extreme_point_of_convex_hull_insert face_of_face face_of_singleton finite.emptyI finite.insertI insert_absorb insert_iff segment_convex_hull)
moreover have "{b} face_of S"
proof -
have "b \<in> convex hull {a} \<or> b extreme_point_of convex hull {b, a}"
by (meson extreme_point_of_convex_hull_insert finite.emptyI finite.insertI)
moreover have "closed_segment a b = convex hull {b, a}"
using closed_segment_commute segment_convex_hull by blast
ultimately show ?thesis
by (metis as assms face_of_face convex_hull_singleton empty_iff face_of_singleton insertE)
qed
ultimately show "a extreme_point_of S" "b extreme_point_of S"
using face_of_singleton by blast+
qed
proposition Krein_Milman_frontier:
fixes S :: "'a::euclidean_space set"
assumes "convex S" "compact S"
shows "S = convex hull (frontier S)"
(is "?lhs = ?rhs")
proof
have "?lhs \<subseteq> convex hull {x. x extreme_point_of S}"
using Krein_Milman_Minkowski assms by blast
also have "... \<subseteq> ?rhs"
apply (rule hull_mono)
apply (auto simp: frontier_def extreme_point_not_in_interior)
using closure_subset apply (force simp: extreme_point_of_def)
done
finally show "?lhs \<subseteq> ?rhs" .
next
have "?rhs \<subseteq> convex hull S"
by (metis Diff_subset \<open>compact S\<close> closure_closed compact_eq_bounded_closed frontier_def hull_mono)
also have "... \<subseteq> ?lhs"
by (simp add: \<open>convex S\<close> hull_same)
finally show "?rhs \<subseteq> ?lhs" .
qed
subsection\<open>Polytopes\<close>
definition\<^marker>\<open>tag important\<close> polytope where
"polytope S \<equiv> \<exists>v. finite v \<and> S = convex hull v"
lemma polytope_translation_eq: "polytope (image (\<lambda>x. a + x) S) \<longleftrightarrow> polytope S"
apply (simp add: polytope_def, safe)
apply (metis convex_hull_translation finite_imageI translation_galois)
by (metis convex_hull_translation finite_imageI)
lemma polytope_linear_image: "\<lbrakk>linear f; polytope p\<rbrakk> \<Longrightarrow> polytope(image f p)"
unfolding polytope_def using convex_hull_linear_image by blast
lemma polytope_empty: "polytope {}"
using convex_hull_empty polytope_def by blast
lemma polytope_convex_hull: "finite S \<Longrightarrow> polytope(convex hull S)"
using polytope_def by auto
lemma polytope_Times: "\<lbrakk>polytope S; polytope T\<rbrakk> \<Longrightarrow> polytope(S \<times> T)"
unfolding polytope_def
by (metis finite_cartesian_product convex_hull_Times)
lemma face_of_polytope_polytope:
fixes S :: "'a::euclidean_space set"
shows "\<lbrakk>polytope S; F face_of S\<rbrakk> \<Longrightarrow> polytope F"
unfolding polytope_def
by (meson face_of_convex_hull_subset finite_imp_compact finite_subset)
lemma finite_polytope_faces:
fixes S :: "'a::euclidean_space set"
assumes "polytope S"
shows "finite {F. F face_of S}"
proof -
obtain v where "finite v" "S = convex hull v"
using assms polytope_def by auto
have "finite ((hull) convex ` {T. T \<subseteq> v})"
by (simp add: \<open>finite v\<close>)
moreover have "{F. F face_of S} \<subseteq> ((hull) convex ` {T. T \<subseteq> v})"
by (metis (no_types, lifting) \<open>finite v\<close> \<open>S = convex hull v\<close> face_of_convex_hull_subset finite_imp_compact image_eqI mem_Collect_eq subsetI)
ultimately show ?thesis
by (blast intro: finite_subset)
qed
lemma finite_polytope_facets:
assumes "polytope S"
shows "finite {T. T facet_of S}"
by (simp add: assms facet_of_def finite_polytope_faces)
lemma polytope_scaling:
assumes "polytope S" shows "polytope (image (\<lambda>x. c *\<^sub>R x) S)"
by (simp add: assms polytope_linear_image)
lemma polytope_imp_compact:
fixes S :: "'a::real_normed_vector set"
shows "polytope S \<Longrightarrow> compact S"
by (metis finite_imp_compact_convex_hull polytope_def)
lemma polytope_imp_convex: "polytope S \<Longrightarrow> convex S"
by (metis convex_convex_hull polytope_def)
lemma polytope_imp_closed:
fixes S :: "'a::real_normed_vector set"
shows "polytope S \<Longrightarrow> closed S"
by (simp add: compact_imp_closed polytope_imp_compact)
lemma polytope_imp_bounded:
fixes S :: "'a::real_normed_vector set"
shows "polytope S \<Longrightarrow> bounded S"
by (simp add: compact_imp_bounded polytope_imp_compact)
lemma polytope_interval: "polytope(cbox a b)"
unfolding polytope_def by (meson closed_interval_as_convex_hull)
lemma polytope_sing: "polytope {a}"
using polytope_def by force
lemma face_of_polytope_insert:
"\<lbrakk>polytope S; a \<notin> affine hull S; F face_of S\<rbrakk> \<Longrightarrow> F face_of convex hull (insert a S)"
by (metis (no_types, lifting) affine_hull_convex_hull face_of_convex_hull_insert hull_insert polytope_def)
proposition face_of_polytope_insert2:
fixes a :: "'a :: euclidean_space"
assumes "polytope S" "a \<notin> affine hull S" "F face_of S"
shows "convex hull (insert a F) face_of convex hull (insert a S)"
proof -
obtain V where "finite V" "S = convex hull V"
using assms by (auto simp: polytope_def)
then have "convex hull (insert a F) face_of convex hull (insert a V)"
using affine_hull_convex_hull assms face_of_convex_hull_insert2 by blast
then show ?thesis
by (metis \<open>S = convex hull V\<close> hull_insert)
qed
subsection\<open>Polyhedra\<close>
definition\<^marker>\<open>tag important\<close> polyhedron where
"polyhedron S \<equiv>
\<exists>F. finite F \<and>
S = \<Inter> F \<and>
(\<forall>h \<in> F. \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x \<le> b})"
lemma polyhedron_Int [intro,simp]:
"\<lbrakk>polyhedron S; polyhedron T\<rbrakk> \<Longrightarrow> polyhedron (S \<inter> T)"
apply (simp add: polyhedron_def, clarify)
apply (rename_tac F G)
apply (rule_tac x="F \<union> G" in exI, auto)
done
lemma polyhedron_UNIV [iff]: "polyhedron UNIV"
unfolding polyhedron_def
by (rule_tac x="{}" in exI) auto
lemma polyhedron_Inter [intro,simp]:
"\<lbrakk>finite F; \<And>S. S \<in> F \<Longrightarrow> polyhedron S\<rbrakk> \<Longrightarrow> polyhedron(\<Inter>F)"
by (induction F rule: finite_induct) auto
lemma polyhedron_empty [iff]: "polyhedron ({} :: 'a :: euclidean_space set)"
proof -
have "\<exists>a. a \<noteq> 0 \<and>
(\<exists>b. {x. (SOME i. i \<in> Basis) \<bullet> x \<le> - 1} = {x. a \<bullet> x \<le> b})"
by (rule_tac x="(SOME i. i \<in> Basis)" in exI) (force simp: SOME_Basis nonzero_Basis)
moreover have "\<exists>a b. a \<noteq> 0 \<and>
{x. - (SOME i. i \<in> Basis) \<bullet> x \<le> - 1} = {x. a \<bullet> x \<le> b}"
apply (rule_tac x="-(SOME i. i \<in> Basis)" in exI)
apply (rule_tac x="-1" in exI)
apply (simp add: SOME_Basis nonzero_Basis)
done
ultimately show ?thesis
unfolding polyhedron_def
apply (rule_tac x="{{x. (SOME i. i \<in> Basis) \<bullet> x \<le> -1},
{x. -(SOME i. i \<in> Basis) \<bullet> x \<le> -1}}" in exI)
apply force
done
qed
lemma polyhedron_halfspace_le:
fixes a :: "'a :: euclidean_space"
shows "polyhedron {x. a \<bullet> x \<le> b}"
proof (cases "a = 0")
case True then show ?thesis by auto
next
case False
then show ?thesis
unfolding polyhedron_def
by (rule_tac x="{{x. a \<bullet> x \<le> b}}" in exI) auto
qed
lemma polyhedron_halfspace_ge:
fixes a :: "'a :: euclidean_space"
shows "polyhedron {x. a \<bullet> x \<ge> b}"
using polyhedron_halfspace_le [of "-a" "-b"] by simp
lemma polyhedron_hyperplane:
fixes a :: "'a :: euclidean_space"
shows "polyhedron {x. a \<bullet> x = b}"
proof -
have "{x. a \<bullet> x = b} = {x. a \<bullet> x \<le> b} \<inter> {x. a \<bullet> x \<ge> b}"
by force
then show ?thesis
by (simp add: polyhedron_halfspace_ge polyhedron_halfspace_le)
qed
lemma affine_imp_polyhedron:
fixes S :: "'a :: euclidean_space set"
shows "affine S \<Longrightarrow> polyhedron S"
by (metis affine_hull_eq polyhedron_Inter polyhedron_hyperplane affine_hull_finite_intersection_hyperplanes [of S])
lemma polyhedron_imp_closed:
fixes S :: "'a :: euclidean_space set"
shows "polyhedron S \<Longrightarrow> closed S"
apply (simp add: polyhedron_def)
using closed_halfspace_le by fastforce
lemma polyhedron_imp_convex:
fixes S :: "'a :: euclidean_space set"
shows "polyhedron S \<Longrightarrow> convex S"
apply (simp add: polyhedron_def)
using convex_Inter convex_halfspace_le by fastforce
lemma polyhedron_affine_hull:
fixes S :: "'a :: euclidean_space set"
shows "polyhedron(affine hull S)"
by (simp add: affine_imp_polyhedron)
subsection\<open>Canonical polyhedron representation making facial structure explicit\<close>
proposition polyhedron_Int_affine:
fixes S :: "'a :: euclidean_space set"
shows "polyhedron S \<longleftrightarrow>
(\<exists>F. finite F \<and> S = (affine hull S) \<inter> \<Inter>F \<and>
(\<forall>h \<in> F. \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x \<le> b}))"
(is "?lhs = ?rhs")
proof
assume ?lhs then show ?rhs
apply (simp add: polyhedron_def)
apply (erule ex_forward)
using hull_subset apply force
done
next
assume ?rhs then show ?lhs
apply clarify
apply (erule ssubst)
apply (force intro: polyhedron_affine_hull polyhedron_halfspace_le)
done
qed
proposition rel_interior_polyhedron_explicit:
assumes "finite F"
and seq: "S = affine hull S \<inter> \<Inter>F"
and faceq: "\<And>h. h \<in> F \<Longrightarrow> a h \<noteq> 0 \<and> h = {x. a h \<bullet> x \<le> b h}"
and psub: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<subset> affine hull S \<inter> \<Inter>F'"
shows "rel_interior S = {x \<in> S. \<forall>h \<in> F. a h \<bullet> x < b h}"
proof -
have rels: "\<And>x. x \<in> rel_interior S \<Longrightarrow> x \<in> S"
by (meson IntE mem_rel_interior)
moreover have "a i \<bullet> x < b i" if x: "x \<in> rel_interior S" and "i \<in> F" for x i
proof -
have fif: "F - {i} \<subset> F"
using \<open>i \<in> F\<close> Diff_insert_absorb Diff_subset set_insert psubsetI by blast
then have "S \<subset> affine hull S \<inter> \<Inter>(F - {i})"
by (rule psub)
then obtain z where ssub: "S \<subseteq> \<Inter>(F - {i})" and zint: "z \<in> \<Inter>(F - {i})"
and "z \<notin> S" and zaff: "z \<in> affine hull S"
by auto
have "z \<noteq> x"
using \<open>z \<notin> S\<close> rels x by blast
have "z \<notin> affine hull S \<inter> \<Inter>F"
using \<open>z \<notin> S\<close> seq by auto
then have aiz: "a i \<bullet> z > b i"
using faceq zint zaff by fastforce
obtain e where "e > 0" "x \<in> S" and e: "ball x e \<inter> affine hull S \<subseteq> S"
using x by (auto simp: mem_rel_interior_ball)
then have ins: "\<And>y. \<lbrakk>norm (x - y) < e; y \<in> affine hull S\<rbrakk> \<Longrightarrow> y \<in> S"
by (metis IntI subsetD dist_norm mem_ball)
define \<xi> where "\<xi> = min (1/2) (e / 2 / norm(z - x))"
have "norm (\<xi> *\<^sub>R x - \<xi> *\<^sub>R z) = norm (\<xi> *\<^sub>R (x - z))"
by (simp add: \<xi>_def algebra_simps norm_mult)
also have "... = \<xi> * norm (x - z)"
using \<open>e > 0\<close> by (simp add: \<xi>_def)
also have "... < e"
using \<open>z \<noteq> x\<close> \<open>e > 0\<close> by (simp add: \<xi>_def min_def field_split_simps norm_minus_commute)
finally have lte: "norm (\<xi> *\<^sub>R x - \<xi> *\<^sub>R z) < e" .
have \<xi>_aff: "\<xi> *\<^sub>R z + (1 - \<xi>) *\<^sub>R x \<in> affine hull S"
by (metis \<open>x \<in> S\<close> add.commute affine_affine_hull diff_add_cancel hull_inc mem_affine zaff)
have "\<xi> *\<^sub>R z + (1 - \<xi>) *\<^sub>R x \<in> S"
apply (rule ins [OF _ \<xi>_aff])
apply (simp add: algebra_simps lte)
done
then obtain l where l: "0 < l" "l < 1" and ls: "(l *\<^sub>R z + (1 - l) *\<^sub>R x) \<in> S"
apply (rule_tac l = \<xi> in that)
using \<open>e > 0\<close> \<open>z \<noteq> x\<close> apply (auto simp: \<xi>_def)
done
then have i: "l *\<^sub>R z + (1 - l) *\<^sub>R x \<in> i"
using seq \<open>i \<in> F\<close> by auto
have "b i * l + (a i \<bullet> x) * (1 - l) < a i \<bullet> (l *\<^sub>R z + (1 - l) *\<^sub>R x)"
using l by (simp add: algebra_simps aiz)
also have "\<dots> \<le> b i" using i l
using faceq mem_Collect_eq \<open>i \<in> F\<close> by blast
finally have "(a i \<bullet> x) * (1 - l) < b i * (1 - l)"
by (simp add: algebra_simps)
with l show ?thesis
by simp
qed
moreover have "x \<in> rel_interior S"
if "x \<in> S" and less: "\<And>h. h \<in> F \<Longrightarrow> a h \<bullet> x < b h" for x
proof -
have 1: "\<And>h. h \<in> F \<Longrightarrow> x \<in> interior h"
by (metis interior_halfspace_le mem_Collect_eq less faceq)
have 2: "\<And>y. \<lbrakk>\<forall>h\<in>F. y \<in> interior h; y \<in> affine hull S\<rbrakk> \<Longrightarrow> y \<in> S"
by (metis IntI Inter_iff contra_subsetD interior_subset seq)
show ?thesis
apply (simp add: rel_interior \<open>x \<in> S\<close>)
apply (rule_tac x="\<Inter>h\<in>F. interior h" in exI)
apply (auto simp: \<open>finite F\<close> open_INT 1 2)
done
qed
ultimately show ?thesis by blast
qed
lemma polyhedron_Int_affine_parallel:
fixes S :: "'a :: euclidean_space set"
shows "polyhedron S \<longleftrightarrow>
(\<exists>F. finite F \<and>
S = (affine hull S) \<inter> (\<Inter>F) \<and>
(\<forall>h \<in> F. \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x \<le> b} \<and>
(\<forall>x \<in> affine hull S. (x + a) \<in> affine hull S)))"
(is "?lhs = ?rhs")
proof
assume ?lhs
then obtain F where "finite F" and seq: "S = (affine hull S) \<inter> \<Inter>F"
and faces: "\<And>h. h \<in> F \<Longrightarrow> \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x \<le> b}"
by (fastforce simp add: polyhedron_Int_affine)
then obtain a b where ab: "\<And>h. h \<in> F \<Longrightarrow> a h \<noteq> 0 \<and> h = {x. a h \<bullet> x \<le> b h}"
by metis
show ?rhs
proof -
have "\<exists>a' b'. a' \<noteq> 0 \<and>
affine hull S \<inter> {x. a' \<bullet> x \<le> b'} = affine hull S \<inter> h \<and>
(\<forall>w \<in> affine hull S. (w + a') \<in> affine hull S)"
if "h \<in> F" "\<not>(affine hull S \<subseteq> h)" for h
proof -
have "a h \<noteq> 0" and "h = {x. a h \<bullet> x \<le> b h}" "h \<inter> \<Inter>F = \<Inter>F"
using \<open>h \<in> F\<close> ab by auto
then have "(affine hull S) \<inter> {x. a h \<bullet> x \<le> b h} \<noteq> {}"
by (metis (no_types) affine_hull_eq_empty inf.absorb_iff2 inf_assoc inf_bot_right inf_commute seq that(2))
moreover have "\<not> (affine hull S \<subseteq> {x. a h \<bullet> x \<le> b h})"
using \<open>h = {x. a h \<bullet> x \<le> b h}\<close> that(2) by blast
ultimately show ?thesis
using affine_parallel_slice [of "affine hull S"]
by (metis \<open>h = {x. a h \<bullet> x \<le> b h}\<close> affine_affine_hull)
qed
then obtain a b
where ab: "\<And>h. \<lbrakk>h \<in> F; \<not> (affine hull S \<subseteq> h)\<rbrakk>
\<Longrightarrow> a h \<noteq> 0 \<and>
affine hull S \<inter> {x. a h \<bullet> x \<le> b h} = affine hull S \<inter> h \<and>
(\<forall>w \<in> affine hull S. (w + a h) \<in> affine hull S)"
by metis
have seq2: "S = affine hull S \<inter> (\<Inter>h\<in>{h \<in> F. \<not> affine hull S \<subseteq> h}. {x. a h \<bullet> x \<le> b h})"
by (subst seq) (auto simp: ab INT_extend_simps)
show ?thesis
apply (rule_tac x="(\<lambda>h. {x. a h \<bullet> x \<le> b h}) ` {h. h \<in> F \<and> \<not>(affine hull S \<subseteq> h)}" in exI)
apply (intro conjI seq2)
using \<open>finite F\<close> apply force
using ab apply blast
done
qed
next
assume ?rhs then show ?lhs
apply (simp add: polyhedron_Int_affine)
by metis
qed
proposition polyhedron_Int_affine_parallel_minimal:
fixes S :: "'a :: euclidean_space set"
shows "polyhedron S \<longleftrightarrow>
(\<exists>F. finite F \<and>
S = (affine hull S) \<inter> (\<Inter>F) \<and>
(\<forall>h \<in> F. \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x \<le> b} \<and>
(\<forall>x \<in> affine hull S. (x + a) \<in> affine hull S)) \<and>
(\<forall>F'. F' \<subset> F \<longrightarrow> S \<subset> (affine hull S) \<inter> (\<Inter>F')))"
(is "?lhs = ?rhs")
proof
assume ?lhs
then obtain f0
where f0: "finite f0"
"S = (affine hull S) \<inter> (\<Inter>f0)"
(is "?P f0")
"\<forall>h \<in> f0. \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x \<le> b} \<and>
(\<forall>x \<in> affine hull S. (x + a) \<in> affine hull S)"
(is "?Q f0")
by (force simp: polyhedron_Int_affine_parallel)
define n where "n = (LEAST n. \<exists>F. card F = n \<and> finite F \<and> ?P F \<and> ?Q F)"
have nf: "\<exists>F. card F = n \<and> finite F \<and> ?P F \<and> ?Q F"
apply (simp add: n_def)
apply (rule LeastI [where k = "card f0"])
using f0 apply auto
done
then obtain F where F: "card F = n" "finite F" and seq: "?P F" and aff: "?Q F"
by blast
then have "\<not> (finite g \<and> ?P g \<and> ?Q g)" if "card g < n" for g
using that by (auto simp: n_def dest!: not_less_Least)
then have *: "\<not> (?P g \<and> ?Q g)" if "g \<subset> F" for g
using that \<open>finite F\<close> psubset_card_mono \<open>card F = n\<close>
by (metis finite_Int inf.strict_order_iff)
have 1: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<subseteq> affine hull S \<inter> \<Inter>F'"
by (subst seq) blast
have 2: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<noteq> affine hull S \<inter> \<Inter>F'"
apply (frule *)
by (metis aff subsetCE subset_iff_psubset_eq)
show ?rhs
by (metis \<open>finite F\<close> seq aff psubsetI 1 2)
next
assume ?rhs then show ?lhs
by (auto simp: polyhedron_Int_affine_parallel)
qed
lemma polyhedron_Int_affine_minimal:
fixes S :: "'a :: euclidean_space set"
shows "polyhedron S \<longleftrightarrow>
(\<exists>F. finite F \<and> S = (affine hull S) \<inter> \<Inter>F \<and>
(\<forall>h \<in> F. \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x \<le> b}) \<and>
(\<forall>F'. F' \<subset> F \<longrightarrow> S \<subset> (affine hull S) \<inter> \<Inter>F'))"
apply (rule iffI)
apply (force simp: polyhedron_Int_affine_parallel_minimal elim!: ex_forward)
apply (auto simp: polyhedron_Int_affine elim!: ex_forward)
done
proposition facet_of_polyhedron_explicit:
assumes "finite F"
and seq: "S = affine hull S \<inter> \<Inter>F"
and faceq: "\<And>h. h \<in> F \<Longrightarrow> a h \<noteq> 0 \<and> h = {x. a h \<bullet> x \<le> b h}"
and psub: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<subset> affine hull S \<inter> \<Inter>F'"
shows "c facet_of S \<longleftrightarrow> (\<exists>h. h \<in> F \<and> c = S \<inter> {x. a h \<bullet> x = b h})"
proof (cases "S = {}")
case True with psub show ?thesis by force
next
case False
have "polyhedron S"
apply (simp add: polyhedron_Int_affine)
apply (rule_tac x=F in exI)
using assms apply force
done
then have "convex S"
by (rule polyhedron_imp_convex)
with False rel_interior_eq_empty have "rel_interior S \<noteq> {}" by blast
then obtain x where "x \<in> rel_interior S" by auto
then obtain T where "open T" "x \<in> T" "x \<in> S" "T \<inter> affine hull S \<subseteq> S"
by (force simp: mem_rel_interior)
then have xaff: "x \<in> affine hull S" and xint: "x \<in> \<Inter>F"
using seq hull_inc by auto
have "rel_interior S = {x \<in> S. \<forall>h\<in>F. a h \<bullet> x < b h}"
by (rule rel_interior_polyhedron_explicit [OF \<open>finite F\<close> seq faceq psub])
with \<open>x \<in> rel_interior S\<close>
have [simp]: "\<And>h. h\<in>F \<Longrightarrow> a h \<bullet> x < b h" by blast
have *: "(S \<inter> {x. a h \<bullet> x = b h}) facet_of S" if "h \<in> F" for h
proof -
have "S \<subset> affine hull S \<inter> \<Inter>(F - {h})"
using psub that by (metis Diff_disjoint Diff_subset insert_disjoint(2) psubsetI)
then obtain z where zaff: "z \<in> affine hull S" and zint: "z \<in> \<Inter>(F - {h})" and "z \<notin> S"
by force
then have "z \<noteq> x" "z \<notin> h" using seq \<open>x \<in> S\<close> by auto
have "x \<in> h" using that xint by auto
then have able: "a h \<bullet> x \<le> b h"
using faceq that by blast
also have "... < a h \<bullet> z" using \<open>z \<notin> h\<close> faceq [OF that] xint by auto
finally have xltz: "a h \<bullet> x < a h \<bullet> z" .
define l where "l = (b h - a h \<bullet> x) / (a h \<bullet> z - a h \<bullet> x)"
define w where "w = (1 - l) *\<^sub>R x + l *\<^sub>R z"
have "0 < l" "l < 1"
using able xltz \<open>b h < a h \<bullet> z\<close> \<open>h \<in> F\<close>
by (auto simp: l_def field_split_simps)
have awlt: "a i \<bullet> w < b i" if "i \<in> F" "i \<noteq> h" for i
proof -
have "(1 - l) * (a i \<bullet> x) < (1 - l) * b i"
by (simp add: \<open>l < 1\<close> \<open>i \<in> F\<close>)
moreover have "l * (a i \<bullet> z) \<le> l * b i"
apply (rule mult_left_mono)
apply (metis Diff_insert_absorb Inter_iff Set.set_insert \<open>h \<in> F\<close> faceq insertE mem_Collect_eq that zint)
using \<open>0 < l\<close>
apply simp
done
ultimately show ?thesis by (simp add: w_def algebra_simps)
qed
have weq: "a h \<bullet> w = b h"
using xltz unfolding w_def l_def
by (simp add: algebra_simps) (simp add: field_simps)
have "w \<in> affine hull S"
by (simp add: w_def mem_affine xaff zaff)
moreover have "w \<in> \<Inter>F"
using \<open>a h \<bullet> w = b h\<close> awlt faceq less_eq_real_def by blast
ultimately have "w \<in> S"
using seq by blast
with weq have "S \<inter> {x. a h \<bullet> x = b h} \<noteq> {}" by blast
moreover have "S \<inter> {x. a h \<bullet> x = b h} face_of S"
apply (rule face_of_Int_supporting_hyperplane_le)
apply (rule \<open>convex S\<close>)
apply (subst (asm) seq)
using faceq that apply fastforce
done
moreover have "affine hull (S \<inter> {x. a h \<bullet> x = b h}) =
(affine hull S) \<inter> {x. a h \<bullet> x = b h}"
proof
show "affine hull (S \<inter> {x. a h \<bullet> x = b h}) \<subseteq> affine hull S \<inter> {x. a h \<bullet> x = b h}"
apply (intro Int_greatest hull_mono Int_lower1)
apply (metis affine_hull_eq affine_hyperplane hull_mono inf_le2)
done
next
show "affine hull S \<inter> {x. a h \<bullet> x = b h} \<subseteq> affine hull (S \<inter> {x. a h \<bullet> x = b h})"
proof
fix y
assume yaff: "y \<in> affine hull S \<inter> {y. a h \<bullet> y = b h}"
obtain T where "0 < T"
and T: "\<And>j. \<lbrakk>j \<in> F; j \<noteq> h\<rbrakk> \<Longrightarrow> T * (a j \<bullet> y - a j \<bullet> w) \<le> b j - a j \<bullet> w"
proof (cases "F - {h} = {}")
case True then show ?thesis
by (rule_tac T=1 in that) auto
next
case False
then obtain h' where h': "h' \<in> F - {h}" by auto
let ?body = "(\<lambda>j. if 0 < a j \<bullet> y - a j \<bullet> w
then (b j - a j \<bullet> w) / (a j \<bullet> y - a j \<bullet> w)
else 1) ` (F - {h})"
define inff where "inff = Inf ?body"
from \<open>finite F\<close> have "finite ?body"
by blast
moreover from h' have "?body \<noteq> {}"
by blast
moreover have "j > 0" if "j \<in> ?body" for j
proof -
from that obtain x where "x \<in> F" and "x \<noteq> h" and *: "j =
(if 0 < a x \<bullet> y - a x \<bullet> w
then (b x - a x \<bullet> w) / (a x \<bullet> y - a x \<bullet> w) else 1)"
by blast
with awlt [of x] have "a x \<bullet> w < b x"
by simp
with * show ?thesis
by simp
qed
ultimately have "0 < inff"
by (simp_all add: finite_less_Inf_iff inff_def)
moreover have "inff * (a j \<bullet> y - a j \<bullet> w) \<le> b j - a j \<bullet> w"
if "j \<in> F" "j \<noteq> h" for j
proof (cases "a j \<bullet> w < a j \<bullet> y")
case True
then have "inff \<le> (b j - a j \<bullet> w) / (a j \<bullet> y - a j \<bullet> w)"
apply (simp add: inff_def)
apply (rule cInf_le_finite)
using \<open>finite F\<close> apply blast
apply (simp add: that split: if_split_asm)
done
then show ?thesis
using \<open>0 < inff\<close> awlt [OF that] mult_strict_left_mono
- by (fastforce simp add: algebra_simps field_split_simps split: if_split_asm)
+ by (fastforce simp add: field_split_simps split: if_split_asm)
next
case False
with \<open>0 < inff\<close> have "inff * (a j \<bullet> y - a j \<bullet> w) \<le> 0"
by (simp add: mult_le_0_iff)
also have "... < b j - a j \<bullet> w"
by (simp add: awlt that)
finally show ?thesis by simp
qed
ultimately show ?thesis
by (blast intro: that)
qed
define c where "c = (1 - T) *\<^sub>R w + T *\<^sub>R y"
have "(1 - T) *\<^sub>R w + T *\<^sub>R y \<in> j" if "j \<in> F" for j
proof (cases "j = h")
case True
have "(1 - T) *\<^sub>R w + T *\<^sub>R y \<in> {x. a h \<bullet> x \<le> b h}"
using weq yaff by (auto simp: algebra_simps)
with True faceq [OF that] show ?thesis by metis
next
case False
with T that have "(1 - T) *\<^sub>R w + T *\<^sub>R y \<in> {x. a j \<bullet> x \<le> b j}"
by (simp add: algebra_simps)
with faceq [OF that] show ?thesis by simp
qed
moreover have "(1 - T) *\<^sub>R w + T *\<^sub>R y \<in> affine hull S"
apply (rule affine_affine_hull [simplified affine_alt, rule_format])
apply (simp add: \<open>w \<in> affine hull S\<close>)
using yaff apply blast
done
ultimately have "c \<in> S"
using seq by (force simp: c_def)
moreover have "a h \<bullet> c = b h"
using yaff by (force simp: c_def algebra_simps weq)
ultimately have caff: "c \<in> affine hull (S \<inter> {y. a h \<bullet> y = b h})"
by (simp add: hull_inc)
have waff: "w \<in> affine hull (S \<inter> {y. a h \<bullet> y = b h})"
using \<open>w \<in> S\<close> weq by (blast intro: hull_inc)
have yeq: "y = (1 - inverse T) *\<^sub>R w + c /\<^sub>R T"
using \<open>0 < T\<close> by (simp add: c_def algebra_simps)
show "y \<in> affine hull (S \<inter> {y. a h \<bullet> y = b h})"
by (metis yeq affine_affine_hull [simplified affine_alt, rule_format, OF waff caff])
qed
qed
ultimately show ?thesis
apply (simp add: facet_of_def)
apply (subst aff_dim_affine_hull [symmetric])
using \<open>b h < a h \<bullet> z\<close> zaff
apply (force simp: aff_dim_affine_Int_hyperplane)
done
qed
show ?thesis
proof
show "\<exists>h. h \<in> F \<and> c = S \<inter> {x. a h \<bullet> x = b h} \<Longrightarrow> c facet_of S"
using * by blast
next
assume "c facet_of S"
then have "c face_of S" "convex c" "c \<noteq> {}" and affc: "aff_dim c = aff_dim S - 1"
by (auto simp: facet_of_def face_of_imp_convex)
then obtain x where x: "x \<in> rel_interior c"
by (force simp: rel_interior_eq_empty)
then have "x \<in> c"
by (meson subsetD rel_interior_subset)
then have "x \<in> S"
using \<open>c facet_of S\<close> facet_of_imp_subset by blast
have rels: "rel_interior S = {x \<in> S. \<forall>h\<in>F. a h \<bullet> x < b h}"
by (rule rel_interior_polyhedron_explicit [OF assms])
have "c \<noteq> S"
using \<open>c facet_of S\<close> facet_of_irrefl by blast
then have "x \<notin> rel_interior S"
by (metis IntI empty_iff \<open>x \<in> c\<close> \<open>c \<noteq> S\<close> \<open>c face_of S\<close> face_of_disjoint_rel_interior)
with rels \<open>x \<in> S\<close> obtain i where "i \<in> F" and i: "a i \<bullet> x \<ge> b i"
by force
have "x \<in> {u. a i \<bullet> u \<le> b i}"
by (metis IntD2 InterE \<open>i \<in> F\<close> \<open>x \<in> S\<close> faceq seq)
then have "a i \<bullet> x \<le> b i" by simp
then have "a i \<bullet> x = b i" using i by auto
have "c \<subseteq> S \<inter> {x. a i \<bullet> x = b i}"
apply (rule subset_of_face_of [of _ S])
apply (simp add: "*" \<open>i \<in> F\<close> facet_of_imp_face_of)
apply (simp add: \<open>c face_of S\<close> face_of_imp_subset)
using \<open>a i \<bullet> x = b i\<close> \<open>x \<in> S\<close> x by blast
then have cface: "c face_of (S \<inter> {x. a i \<bullet> x = b i})"
by (meson \<open>c face_of S\<close> face_of_subset inf_le1)
have con: "convex (S \<inter> {x. a i \<bullet> x = b i})"
by (simp add: \<open>convex S\<close> convex_Int convex_hyperplane)
show "\<exists>h. h \<in> F \<and> c = S \<inter> {x. a h \<bullet> x = b h}"
apply (rule_tac x=i in exI)
apply (simp add: \<open>i \<in> F\<close>)
by (metis (no_types) * \<open>i \<in> F\<close> affc facet_of_def less_irrefl face_of_aff_dim_lt [OF con cface])
qed
qed
lemma face_of_polyhedron_subset_explicit:
fixes S :: "'a :: euclidean_space set"
assumes "finite F"
and seq: "S = affine hull S \<inter> \<Inter>F"
and faceq: "\<And>h. h \<in> F \<Longrightarrow> a h \<noteq> 0 \<and> h = {x. a h \<bullet> x \<le> b h}"
and psub: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<subset> affine hull S \<inter> \<Inter>F'"
and c: "c face_of S" and "c \<noteq> {}" "c \<noteq> S"
obtains h where "h \<in> F" "c \<subseteq> S \<inter> {x. a h \<bullet> x = b h}"
proof -
have "c \<subseteq> S" using \<open>c face_of S\<close>
by (simp add: face_of_imp_subset)
have "polyhedron S"
apply (simp add: polyhedron_Int_affine)
by (metis \<open>finite F\<close> faceq seq)
then have "convex S"
by (simp add: polyhedron_imp_convex)
then have *: "(S \<inter> {x. a h \<bullet> x = b h}) face_of S" if "h \<in> F" for h
apply (rule face_of_Int_supporting_hyperplane_le)
using faceq seq that by fastforce
have "rel_interior c \<noteq> {}"
using c \<open>c \<noteq> {}\<close> face_of_imp_convex rel_interior_eq_empty by blast
then obtain x where "x \<in> rel_interior c" by auto
have rels: "rel_interior S = {x \<in> S. \<forall>h\<in>F. a h \<bullet> x < b h}"
by (rule rel_interior_polyhedron_explicit [OF \<open>finite F\<close> seq faceq psub])
then have xnot: "x \<notin> rel_interior S"
by (metis IntI \<open>x \<in> rel_interior c\<close> c \<open>c \<noteq> S\<close> contra_subsetD empty_iff face_of_disjoint_rel_interior rel_interior_subset)
then have "x \<in> S"
using \<open>c \<subseteq> S\<close> \<open>x \<in> rel_interior c\<close> rel_interior_subset by auto
then have xint: "x \<in> \<Inter>F"
using seq by blast
have "F \<noteq> {}" using assms
by (metis affine_Int affine_Inter affine_affine_hull ex_in_conv face_of_affine_trivial)
then obtain i where "i \<in> F" "\<not> (a i \<bullet> x < b i)"
using \<open>x \<in> S\<close> rels xnot by auto
with xint have "a i \<bullet> x = b i"
by (metis eq_iff mem_Collect_eq not_le Inter_iff faceq)
have face: "S \<inter> {x. a i \<bullet> x = b i} face_of S"
by (simp add: "*" \<open>i \<in> F\<close>)
show ?thesis
apply (rule_tac h = i in that)
apply (rule \<open>i \<in> F\<close>)
apply (rule subset_of_face_of [OF face \<open>c \<subseteq> S\<close>])
using \<open>a i \<bullet> x = b i\<close> \<open>x \<in> rel_interior c\<close> \<open>x \<in> S\<close> apply blast
done
qed
text\<open>Initial part of proof duplicates that above\<close>
proposition face_of_polyhedron_explicit:
fixes S :: "'a :: euclidean_space set"
assumes "finite F"
and seq: "S = affine hull S \<inter> \<Inter>F"
and faceq: "\<And>h. h \<in> F \<Longrightarrow> a h \<noteq> 0 \<and> h = {x. a h \<bullet> x \<le> b h}"
and psub: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<subset> affine hull S \<inter> \<Inter>F'"
and c: "c face_of S" and "c \<noteq> {}" "c \<noteq> S"
shows "c = \<Inter>{S \<inter> {x. a h \<bullet> x = b h} | h. h \<in> F \<and> c \<subseteq> S \<inter> {x. a h \<bullet> x = b h}}"
proof -
let ?ab = "\<lambda>h. {x. a h \<bullet> x = b h}"
have "c \<subseteq> S" using \<open>c face_of S\<close>
by (simp add: face_of_imp_subset)
have "polyhedron S"
apply (simp add: polyhedron_Int_affine)
by (metis \<open>finite F\<close> faceq seq)
then have "convex S"
by (simp add: polyhedron_imp_convex)
then have *: "(S \<inter> ?ab h) face_of S" if "h \<in> F" for h
apply (rule face_of_Int_supporting_hyperplane_le)
using faceq seq that by fastforce
have "rel_interior c \<noteq> {}"
using c \<open>c \<noteq> {}\<close> face_of_imp_convex rel_interior_eq_empty by blast
then obtain z where z: "z \<in> rel_interior c" by auto
have rels: "rel_interior S = {z \<in> S. \<forall>h\<in>F. a h \<bullet> z < b h}"
by (rule rel_interior_polyhedron_explicit [OF \<open>finite F\<close> seq faceq psub])
then have xnot: "z \<notin> rel_interior S"
by (metis IntI \<open>z \<in> rel_interior c\<close> c \<open>c \<noteq> S\<close> contra_subsetD empty_iff face_of_disjoint_rel_interior rel_interior_subset)
then have "z \<in> S"
using \<open>c \<subseteq> S\<close> \<open>z \<in> rel_interior c\<close> rel_interior_subset by auto
with seq have xint: "z \<in> \<Inter>F" by blast
have "open (\<Inter>h\<in>{h \<in> F. a h \<bullet> z < b h}. {w. a h \<bullet> w < b h})"
by (auto simp: \<open>finite F\<close> open_halfspace_lt open_INT)
then obtain e where "0 < e"
"ball z e \<subseteq> (\<Inter>h\<in>{h \<in> F. a h \<bullet> z < b h}. {w. a h \<bullet> w < b h})"
by (auto intro: openE [of _ z])
then have e: "\<And>h. \<lbrakk>h \<in> F; a h \<bullet> z < b h\<rbrakk> \<Longrightarrow> ball z e \<subseteq> {w. a h \<bullet> w < b h}"
by blast
have "c \<subseteq> (S \<inter> ?ab h) \<longleftrightarrow> z \<in> S \<inter> ?ab h" if "h \<in> F" for h
proof
show "z \<in> S \<inter> ?ab h \<Longrightarrow> c \<subseteq> S \<inter> ?ab h"
apply (rule subset_of_face_of [of _ S])
using that \<open>c \<subseteq> S\<close> \<open>z \<in> rel_interior c\<close>
using facet_of_polyhedron_explicit [OF \<open>finite F\<close> seq faceq psub]
unfolding facet_of_def
apply auto
done
next
show "c \<subseteq> S \<inter> ?ab h \<Longrightarrow> z \<in> S \<inter> ?ab h"
using \<open>z \<in> rel_interior c\<close> rel_interior_subset by force
qed
then have **: "{S \<inter> ?ab h | h. h \<in> F \<and> c \<subseteq> S \<and> c \<subseteq> ?ab h} =
{S \<inter> ?ab h |h. h \<in> F \<and> z \<in> S \<inter> ?ab h}"
by blast
have bsub: "ball z e \<inter> affine hull \<Inter>{S \<inter> ?ab h |h. h \<in> F \<and> a h \<bullet> z = b h}
\<subseteq> affine hull S \<inter> \<Inter>F \<inter> \<Inter>{?ab h |h. h \<in> F \<and> a h \<bullet> z = b h}"
if "i \<in> F" and i: "a i \<bullet> z = b i" for i
proof -
have sub: "ball z e \<inter> \<Inter>{?ab h |h. h \<in> F \<and> a h \<bullet> z = b h} \<subseteq> j"
if "j \<in> F" for j
proof -
have "a j \<bullet> z \<le> b j" using faceq that xint by auto
then consider "a j \<bullet> z < b j" | "a j \<bullet> z = b j" by linarith
then have "\<exists>G. G \<in> {?ab h |h. h \<in> F \<and> a h \<bullet> z = b h} \<and> ball z e \<inter> G \<subseteq> j"
proof cases
assume "a j \<bullet> z < b j"
then have "ball z e \<inter> {x. a i \<bullet> x = b i} \<subseteq> j"
using e [OF \<open>j \<in> F\<close>] faceq that
by (fastforce simp: ball_def)
then show ?thesis
by (rule_tac x="{x. a i \<bullet> x = b i}" in exI) (force simp: \<open>i \<in> F\<close> i)
next
assume eq: "a j \<bullet> z = b j"
with faceq that show ?thesis
by (rule_tac x="{x. a j \<bullet> x = b j}" in exI) (fastforce simp add: \<open>j \<in> F\<close>)
qed
then show ?thesis by blast
qed
have 1: "affine hull \<Inter>{S \<inter> ?ab h |h. h \<in> F \<and> a h \<bullet> z = b h} \<subseteq> affine hull S"
apply (rule hull_mono)
using that \<open>z \<in> S\<close> by auto
have 2: "affine hull \<Inter>{S \<inter> ?ab h |h. h \<in> F \<and> a h \<bullet> z = b h}
\<subseteq> \<Inter>{?ab h |h. h \<in> F \<and> a h \<bullet> z = b h}"
by (rule hull_minimal) (auto intro: affine_hyperplane)
have 3: "ball z e \<inter> \<Inter>{?ab h |h. h \<in> F \<and> a h \<bullet> z = b h} \<subseteq> \<Inter>F"
by (iprover intro: sub Inter_greatest)
have *: "\<lbrakk>A \<subseteq> (B :: 'a set); A \<subseteq> C; E \<inter> C \<subseteq> D\<rbrakk> \<Longrightarrow> E \<inter> A \<subseteq> (B \<inter> D) \<inter> C"
for A B C D E by blast
show ?thesis by (intro * 1 2 3)
qed
have "\<exists>h. h \<in> F \<and> c \<subseteq> ?ab h"
apply (rule face_of_polyhedron_subset_explicit [OF \<open>finite F\<close> seq faceq psub])
using assms by auto
then have fac: "\<Inter>{S \<inter> ?ab h |h. h \<in> F \<and> c \<subseteq> S \<inter> ?ab h} face_of S"
using * by (force simp: \<open>c \<subseteq> S\<close> intro: face_of_Inter)
have red:
"(\<And>a. P a \<Longrightarrow> T \<subseteq> S \<inter> \<Inter>{F x |x. P x}) \<Longrightarrow> T \<subseteq> \<Inter>{S \<inter> F x |x. P x}"
for P T F by blast
have "ball z e \<inter> affine hull \<Inter>{S \<inter> ?ab h |h. h \<in> F \<and> a h \<bullet> z = b h}
\<subseteq> \<Inter>{S \<inter> ?ab h |h. h \<in> F \<and> a h \<bullet> z = b h}"
apply (rule red)
apply (metis seq bsub)
done
with \<open>0 < e\<close> have zinrel: "z \<in> rel_interior
(\<Inter>{S \<inter> ?ab h |h. h \<in> F \<and> z \<in> S \<and> a h \<bullet> z = b h})"
by (auto simp: mem_rel_interior_ball \<open>z \<in> S\<close>)
show ?thesis
apply (rule face_of_eq [OF c fac])
using z zinrel apply (force simp: **)
done
qed
subsection\<open>More general corollaries from the explicit representation\<close>
corollary facet_of_polyhedron:
assumes "polyhedron S" and "c facet_of S"
obtains a b where "a \<noteq> 0" "S \<subseteq> {x. a \<bullet> x \<le> b}" "c = S \<inter> {x. a \<bullet> x = b}"
proof -
obtain F where "finite F" and seq: "S = affine hull S \<inter> \<Inter>F"
and faces: "\<And>h. h \<in> F \<Longrightarrow> \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x \<le> b}"
and min: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<subset> (affine hull S) \<inter> \<Inter>F'"
using assms by (simp add: polyhedron_Int_affine_minimal) meson
then obtain a b where ab: "\<And>h. h \<in> F \<Longrightarrow> a h \<noteq> 0 \<and> h = {x. a h \<bullet> x \<le> b h}"
by metis
obtain i where "i \<in> F" and c: "c = S \<inter> {x. a i \<bullet> x = b i}"
using facet_of_polyhedron_explicit [OF \<open>finite F\<close> seq ab min] assms
by force
moreover have ssub: "S \<subseteq> {x. a i \<bullet> x \<le> b i}"
apply (subst seq)
using \<open>i \<in> F\<close> ab by auto
ultimately show ?thesis
by (rule_tac a = "a i" and b = "b i" in that) (simp_all add: ab)
qed
corollary face_of_polyhedron:
assumes "polyhedron S" and "c face_of S" and "c \<noteq> {}" and "c \<noteq> S"
shows "c = \<Inter>{F. F facet_of S \<and> c \<subseteq> F}"
proof -
obtain F where "finite F" and seq: "S = affine hull S \<inter> \<Inter>F"
and faces: "\<And>h. h \<in> F \<Longrightarrow> \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x \<le> b}"
and min: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<subset> (affine hull S) \<inter> \<Inter>F'"
using assms by (simp add: polyhedron_Int_affine_minimal) meson
then obtain a b where ab: "\<And>h. h \<in> F \<Longrightarrow> a h \<noteq> 0 \<and> h = {x. a h \<bullet> x \<le> b h}"
by metis
show ?thesis
apply (subst face_of_polyhedron_explicit [OF \<open>finite F\<close> seq ab min])
apply (auto simp: assms facet_of_polyhedron_explicit [OF \<open>finite F\<close> seq ab min] cong: Collect_cong)
done
qed
lemma face_of_polyhedron_subset_facet:
assumes "polyhedron S" and "c face_of S" and "c \<noteq> {}" and "c \<noteq> S"
obtains F where "F facet_of S" "c \<subseteq> F"
using face_of_polyhedron assms
by (metis (no_types, lifting) Inf_greatest antisym_conv face_of_imp_subset mem_Collect_eq)
lemma exposed_face_of_polyhedron:
assumes "polyhedron S"
shows "F exposed_face_of S \<longleftrightarrow> F face_of S"
proof
show "F exposed_face_of S \<Longrightarrow> F face_of S"
by (simp add: exposed_face_of_def)
next
assume "F face_of S"
show "F exposed_face_of S"
proof (cases "F = {} \<or> F = S")
case True then show ?thesis
using \<open>F face_of S\<close> exposed_face_of by blast
next
case False
then have "{g. g facet_of S \<and> F \<subseteq> g} \<noteq> {}"
by (metis Collect_empty_eq_bot \<open>F face_of S\<close> assms empty_def face_of_polyhedron_subset_facet)
moreover have "\<And>T. \<lbrakk>T facet_of S; F \<subseteq> T\<rbrakk> \<Longrightarrow> T exposed_face_of S"
by (metis assms exposed_face_of facet_of_imp_face_of facet_of_polyhedron)
ultimately have "\<Inter>{fa.
fa facet_of S \<and> F \<subseteq> fa} exposed_face_of S"
by (metis (no_types, lifting) mem_Collect_eq exposed_face_of_Inter)
then show ?thesis
using False
apply (subst face_of_polyhedron [OF assms \<open>F face_of S\<close>], auto)
done
qed
qed
lemma face_of_polyhedron_polyhedron:
fixes S :: "'a :: euclidean_space set"
assumes "polyhedron S" "c face_of S" shows "polyhedron c"
by (metis assms face_of_imp_eq_affine_Int polyhedron_Int polyhedron_affine_hull polyhedron_imp_convex)
lemma finite_polyhedron_faces:
fixes S :: "'a :: euclidean_space set"
assumes "polyhedron S"
shows "finite {F. F face_of S}"
proof -
obtain F where "finite F" and seq: "S = affine hull S \<inter> \<Inter>F"
and faces: "\<And>h. h \<in> F \<Longrightarrow> \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x \<le> b}"
and min: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<subset> (affine hull S) \<inter> \<Inter>F'"
using assms by (simp add: polyhedron_Int_affine_minimal) meson
then obtain a b where ab: "\<And>h. h \<in> F \<Longrightarrow> a h \<noteq> 0 \<and> h = {x. a h \<bullet> x \<le> b h}"
by metis
have "finite {\<Inter>{S \<inter> {x. a h \<bullet> x = b h} |h. h \<in> F'}| F'. F' \<in> Pow F}"
by (simp add: \<open>finite F\<close>)
moreover have "{F. F face_of S} - {{}, S} \<subseteq> {\<Inter>{S \<inter> {x. a h \<bullet> x = b h} |h. h \<in> F'}| F'. F' \<in> Pow F}"
apply clarify
apply (rename_tac c)
apply (drule face_of_polyhedron_explicit [OF \<open>finite F\<close> seq ab min, simplified], simp_all)
apply (erule ssubst)
apply (rule_tac x="{h \<in> F. c \<subseteq> S \<inter> {x. a h \<bullet> x = b h}}" in exI, auto)
done
ultimately show ?thesis
by (meson finite.emptyI finite.insertI finite_Diff2 finite_subset)
qed
lemma finite_polyhedron_exposed_faces:
"polyhedron S \<Longrightarrow> finite {F. F exposed_face_of S}"
using exposed_face_of_polyhedron finite_polyhedron_faces by fastforce
lemma finite_polyhedron_extreme_points:
fixes S :: "'a :: euclidean_space set"
shows "polyhedron S \<Longrightarrow> finite {v. v extreme_point_of S}"
apply (simp add: face_of_singleton [symmetric])
apply (rule finite_subset [OF _ finite_vimageI [OF finite_polyhedron_faces]], auto)
done
lemma finite_polyhedron_facets:
fixes S :: "'a :: euclidean_space set"
shows "polyhedron S \<Longrightarrow> finite {F. F facet_of S}"
unfolding facet_of_def
by (blast intro: finite_subset [OF _ finite_polyhedron_faces])
proposition rel_interior_of_polyhedron:
fixes S :: "'a :: euclidean_space set"
assumes "polyhedron S"
shows "rel_interior S = S - \<Union>{F. F facet_of S}"
proof -
obtain F where "finite F" and seq: "S = affine hull S \<inter> \<Inter>F"
and faces: "\<And>h. h \<in> F \<Longrightarrow> \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x \<le> b}"
and min: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<subset> (affine hull S) \<inter> \<Inter>F'"
using assms by (simp add: polyhedron_Int_affine_minimal) meson
then obtain a b where ab: "\<And>h. h \<in> F \<Longrightarrow> a h \<noteq> 0 \<and> h = {x. a h \<bullet> x \<le> b h}"
by metis
have facet: "(c facet_of S) \<longleftrightarrow> (\<exists>h. h \<in> F \<and> c = S \<inter> {x. a h \<bullet> x = b h})" for c
by (rule facet_of_polyhedron_explicit [OF \<open>finite F\<close> seq ab min])
have rel: "rel_interior S = {x \<in> S. \<forall>h\<in>F. a h \<bullet> x < b h}"
by (rule rel_interior_polyhedron_explicit [OF \<open>finite F\<close> seq ab min])
have "a h \<bullet> x < b h" if "x \<in> S" "h \<in> F" and xnot: "x \<notin> \<Union>{F. F facet_of S}" for x h
proof -
have "x \<in> \<Inter>F" using seq that by force
with \<open>h \<in> F\<close> ab have "a h \<bullet> x \<le> b h" by auto
then consider "a h \<bullet> x < b h" | "a h \<bullet> x = b h" by linarith
then show ?thesis
proof cases
case 1 then show ?thesis .
next
case 2
have "Collect ((\<in>) x) \<notin> Collect ((\<in>) (\<Union>{A. A facet_of S}))"
using xnot by fastforce
then have "F \<notin> Collect ((\<in>) h)"
using 2 \<open>x \<in> S\<close> facet by blast
with \<open>h \<in> F\<close> have "\<Inter>F \<subseteq> S \<inter> {x. a h \<bullet> x = b h}" by blast
with 2 that \<open>x \<in> \<Inter>F\<close> show ?thesis
apply simp
apply (drule_tac x="\<Inter>F" in spec)
apply (simp add: facet)
apply (drule_tac x=h in spec)
using seq by auto
qed
qed
moreover have "\<exists>h\<in>F. a h \<bullet> x \<ge> b h" if "x \<in> \<Union>{F. F facet_of S}" for x
using that by (force simp: facet)
ultimately show ?thesis
by (force simp: rel)
qed
lemma rel_boundary_of_polyhedron:
fixes S :: "'a :: euclidean_space set"
assumes "polyhedron S"
shows "S - rel_interior S = \<Union> {F. F facet_of S}"
using facet_of_imp_subset by (fastforce simp add: rel_interior_of_polyhedron assms)
lemma rel_frontier_of_polyhedron:
fixes S :: "'a :: euclidean_space set"
assumes "polyhedron S"
shows "rel_frontier S = \<Union> {F. F facet_of S}"
by (simp add: assms rel_frontier_def polyhedron_imp_closed rel_boundary_of_polyhedron)
lemma rel_frontier_of_polyhedron_alt:
fixes S :: "'a :: euclidean_space set"
assumes "polyhedron S"
shows "rel_frontier S = \<Union> {F. F face_of S \<and> (F \<noteq> S)}"
apply (rule subset_antisym)
apply (force simp: rel_frontier_of_polyhedron facet_of_def assms)
using face_of_subset_rel_frontier by fastforce
text\<open>A characterization of polyhedra as having finitely many faces\<close>
proposition polyhedron_eq_finite_exposed_faces:
fixes S :: "'a :: euclidean_space set"
shows "polyhedron S \<longleftrightarrow> closed S \<and> convex S \<and> finite {F. F exposed_face_of S}"
(is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
by (auto simp: polyhedron_imp_closed polyhedron_imp_convex finite_polyhedron_exposed_faces)
next
assume ?rhs
then have "closed S" "convex S" and fin: "finite {F. F exposed_face_of S}" by auto
show ?lhs
proof (cases "S = {}")
case True then show ?thesis by auto
next
case False
define F where "F = {h. h exposed_face_of S \<and> h \<noteq> {} \<and> h \<noteq> S}"
have "finite F" by (simp add: fin F_def)
have hface: "h face_of S"
and "\<exists>a b. a \<noteq> 0 \<and> S \<subseteq> {x. a \<bullet> x \<le> b} \<and> h = S \<inter> {x. a \<bullet> x = b}"
if "h \<in> F" for h
using exposed_face_of F_def that by simp_all auto
then obtain a b where ab:
"\<And>h. h \<in> F \<Longrightarrow> a h \<noteq> 0 \<and> S \<subseteq> {x. a h \<bullet> x \<le> b h} \<and> h = S \<inter> {x. a h \<bullet> x = b h}"
by metis
have *: "False"
if paff: "p \<in> affine hull S" and "p \<notin> S"
and pint: "p \<in> \<Inter>{{x. a h \<bullet> x \<le> b h} |h. h \<in> F}" for p
proof -
have "rel_interior S \<noteq> {}"
by (simp add: \<open>S \<noteq> {}\<close> \<open>convex S\<close> rel_interior_eq_empty)
then obtain c where c: "c \<in> rel_interior S" by auto
with rel_interior_subset have "c \<in> S" by blast
have ccp: "closed_segment c p \<subseteq> affine hull S"
by (meson affine_affine_hull affine_imp_convex c closed_segment_subset hull_subset paff rel_interior_subset subsetCE)
obtain x where xcl: "x \<in> closed_segment c p" and "x \<in> S" and xnot: "x \<notin> rel_interior S"
using connected_openin [of "closed_segment c p"]
apply simp
apply (drule_tac x="closed_segment c p \<inter> rel_interior S" in spec)
apply (erule impE)
apply (force simp: openin_rel_interior openin_Int intro: openin_subtopology_Int_subset [OF _ ccp])
apply (drule_tac x="closed_segment c p \<inter> (- S)" in spec)
using rel_interior_subset \<open>closed S\<close> c \<open>p \<notin> S\<close> apply blast
done
then obtain \<mu> where "0 \<le> \<mu>" "\<mu> \<le> 1" and xeq: "x = (1 - \<mu>) *\<^sub>R c + \<mu> *\<^sub>R p"
by (auto simp: in_segment)
show False
proof (cases "\<mu>=0 \<or> \<mu>=1")
case True with xeq c xnot \<open>x \<in> S\<close> \<open>p \<notin> S\<close>
show False by auto
next
case False
then have xos: "x \<in> open_segment c p"
using \<open>x \<in> S\<close> c open_segment_def that(2) xcl xnot by auto
have xclo: "x \<in> closure S"
using \<open>x \<in> S\<close> closure_subset by blast
obtain d where "d \<noteq> 0"
and dle: "\<And>y. y \<in> closure S \<Longrightarrow> d \<bullet> x \<le> d \<bullet> y"
and dless: "\<And>y. y \<in> rel_interior S \<Longrightarrow> d \<bullet> x < d \<bullet> y"
by (metis supporting_hyperplane_relative_frontier [OF \<open>convex S\<close> xclo xnot])
have sex: "S \<inter> {y. d \<bullet> y = d \<bullet> x} exposed_face_of S"
by (simp add: \<open>closed S\<close> dle exposed_face_of_Int_supporting_hyperplane_ge [OF \<open>convex S\<close>])
have sne: "S \<inter> {y. d \<bullet> y = d \<bullet> x} \<noteq> {}"
using \<open>x \<in> S\<close> by blast
have sns: "S \<inter> {y. d \<bullet> y = d \<bullet> x} \<noteq> S"
by (metis (mono_tags) Int_Collect c subsetD dless not_le order_refl rel_interior_subset)
obtain h where "h \<in> F" "x \<in> h"
apply (rule_tac h="S \<inter> {y. d \<bullet> y = d \<bullet> x}" in that)
apply (simp_all add: F_def sex sne sns \<open>x \<in> S\<close>)
done
have abface: "{y. a h \<bullet> y = b h} face_of {y. a h \<bullet> y \<le> b h}"
using hyperplane_face_of_halfspace_le by blast
then have "c \<in> h"
using face_ofD [OF abface xos] \<open>c \<in> S\<close> \<open>h \<in> F\<close> ab pint \<open>x \<in> h\<close> by blast
with c have "h \<inter> rel_interior S \<noteq> {}" by blast
then show False
using \<open>h \<in> F\<close> F_def face_of_disjoint_rel_interior hface by auto
qed
qed
have "S \<subseteq> affine hull S \<inter> \<Inter>{{x. a h \<bullet> x \<le> b h} |h. h \<in> F}"
using ab by (auto simp: hull_subset)
moreover have "affine hull S \<inter> \<Inter>{{x. a h \<bullet> x \<le> b h} |h. h \<in> F} \<subseteq> S"
using * by blast
ultimately have "S = affine hull S \<inter> \<Inter> {{x. a h \<bullet> x \<le> b h} |h. h \<in> F}" ..
then show ?thesis
apply (rule ssubst)
apply (force intro: polyhedron_affine_hull polyhedron_halfspace_le simp: \<open>finite F\<close>)
done
qed
qed
corollary polyhedron_eq_finite_faces:
fixes S :: "'a :: euclidean_space set"
shows "polyhedron S \<longleftrightarrow> closed S \<and> convex S \<and> finite {F. F face_of S}"
(is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
by (simp add: finite_polyhedron_faces polyhedron_imp_closed polyhedron_imp_convex)
next
assume ?rhs
then show ?lhs
by (force simp: polyhedron_eq_finite_exposed_faces exposed_face_of intro: finite_subset)
qed
lemma polyhedron_linear_image_eq:
fixes h :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
assumes "linear h" "bij h"
shows "polyhedron (h ` S) \<longleftrightarrow> polyhedron S"
proof -
have *: "{f. P f} = (image h) ` {f. P (h ` f)}" for P
apply safe
apply (rule_tac x="inv h ` x" in image_eqI)
apply (auto simp: \<open>bij h\<close> bij_is_surj image_f_inv_f)
done
have "inj h" using bij_is_inj assms by blast
then have injim: "inj_on ((`) h) A" for A
by (simp add: inj_on_def inj_image_eq_iff)
show ?thesis
using \<open>linear h\<close> \<open>inj h\<close>
apply (simp add: polyhedron_eq_finite_faces closed_injective_linear_image_eq)
apply (simp add: * face_of_linear_image [of h _ S, symmetric] finite_image_iff injim)
done
qed
lemma polyhedron_negations:
fixes S :: "'a :: euclidean_space set"
shows "polyhedron S \<Longrightarrow> polyhedron(image uminus S)"
by (subst polyhedron_linear_image_eq)
(auto simp: bij_uminus intro!: linear_uminus)
subsection\<open>Relation between polytopes and polyhedra\<close>
proposition polytope_eq_bounded_polyhedron:
fixes S :: "'a :: euclidean_space set"
shows "polytope S \<longleftrightarrow> polyhedron S \<and> bounded S"
(is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
by (simp add: finite_polytope_faces polyhedron_eq_finite_faces
polytope_imp_closed polytope_imp_convex polytope_imp_bounded)
next
assume ?rhs then show ?lhs
unfolding polytope_def
apply (rule_tac x="{v. v extreme_point_of S}" in exI)
apply (simp add: finite_polyhedron_extreme_points Krein_Milman_Minkowski compact_eq_bounded_closed polyhedron_imp_closed polyhedron_imp_convex)
done
qed
lemma polytope_Int:
fixes S :: "'a :: euclidean_space set"
shows "\<lbrakk>polytope S; polytope T\<rbrakk> \<Longrightarrow> polytope(S \<inter> T)"
by (simp add: polytope_eq_bounded_polyhedron bounded_Int)
lemma polytope_Int_polyhedron:
fixes S :: "'a :: euclidean_space set"
shows "\<lbrakk>polytope S; polyhedron T\<rbrakk> \<Longrightarrow> polytope(S \<inter> T)"
by (simp add: bounded_Int polytope_eq_bounded_polyhedron)
lemma polyhedron_Int_polytope:
fixes S :: "'a :: euclidean_space set"
shows "\<lbrakk>polyhedron S; polytope T\<rbrakk> \<Longrightarrow> polytope(S \<inter> T)"
by (simp add: bounded_Int polytope_eq_bounded_polyhedron)
lemma polytope_imp_polyhedron:
fixes S :: "'a :: euclidean_space set"
shows "polytope S \<Longrightarrow> polyhedron S"
by (simp add: polytope_eq_bounded_polyhedron)
lemma polytope_facet_exists:
fixes p :: "'a :: euclidean_space set"
assumes "polytope p" "0 < aff_dim p"
obtains F where "F facet_of p"
proof (cases "p = {}")
case True with assms show ?thesis by auto
next
case False
then obtain v where "v extreme_point_of p"
using extreme_point_exists_convex
by (blast intro: \<open>polytope p\<close> polytope_imp_compact polytope_imp_convex)
then
show ?thesis
by (metis face_of_polyhedron_subset_facet polytope_imp_polyhedron aff_dim_sing
all_not_in_conv assms face_of_singleton less_irrefl singletonI that)
qed
lemma polyhedron_interval [iff]: "polyhedron(cbox a b)"
by (metis polytope_imp_polyhedron polytope_interval)
lemma polyhedron_convex_hull:
fixes S :: "'a :: euclidean_space set"
shows "finite S \<Longrightarrow> polyhedron(convex hull S)"
by (simp add: polytope_convex_hull polytope_imp_polyhedron)
subsection\<open>Relative and absolute frontier of a polytope\<close>
lemma rel_boundary_of_convex_hull:
fixes S :: "'a::euclidean_space set"
assumes "\<not> affine_dependent S"
shows "(convex hull S) - rel_interior(convex hull S) = (\<Union>a\<in>S. convex hull (S - {a}))"
proof -
have "finite S" by (metis assms aff_independent_finite)
then consider "card S = 0" | "card S = 1" | "2 \<le> card S" by arith
then show ?thesis
proof cases
case 1 then have "S = {}" by (simp add: \<open>finite S\<close>)
then show ?thesis by simp
next
case 2 show ?thesis
by (auto intro: card_1_singletonE [OF \<open>card S = 1\<close>])
next
case 3
with assms show ?thesis
by (auto simp: polyhedron_convex_hull rel_boundary_of_polyhedron facet_of_convex_hull_affine_independent_alt \<open>finite S\<close>)
qed
qed
proposition frontier_of_convex_hull:
fixes S :: "'a::euclidean_space set"
assumes "card S = Suc (DIM('a))"
shows "frontier(convex hull S) = \<Union> {convex hull (S - {a}) | a. a \<in> S}"
proof (cases "affine_dependent S")
case True
have [iff]: "finite S"
using assms using card_infinite by force
then have ccs: "closed (convex hull S)"
by (simp add: compact_imp_closed finite_imp_compact_convex_hull)
{ fix x T
assume "finite T" "T \<subseteq> S" "int (card T) \<le> aff_dim S + 1" "x \<in> convex hull T"
then have "S \<noteq> T"
using True \<open>finite S\<close> aff_dim_le_card affine_independent_iff_card by fastforce
then obtain a where "a \<in> S" "a \<notin> T"
using \<open>T \<subseteq> S\<close> by blast
then have "x \<in> (\<Union>a\<in>S. convex hull (S - {a}))"
using True affine_independent_iff_card [of S]
apply simp
apply (metis (no_types, hide_lams) Diff_eq_empty_iff Diff_insert0 \<open>a \<notin> T\<close> \<open>T \<subseteq> S\<close> \<open>x \<in> convex hull T\<close> hull_mono insert_Diff_single subsetCE)
done
} note * = this
have 1: "convex hull S \<subseteq> (\<Union> a\<in>S. convex hull (S - {a}))"
apply (subst caratheodory_aff_dim)
apply (blast intro: *)
done
have 2: "\<Union>((\<lambda>a. convex hull (S - {a})) ` S) \<subseteq> convex hull S"
by (rule Union_least) (metis (no_types, lifting) Diff_subset hull_mono imageE)
show ?thesis using True
apply (simp add: segment_convex_hull frontier_def)
using interior_convex_hull_eq_empty [OF assms]
apply (simp add: closure_closed [OF ccs])
apply (rule subset_antisym)
using 1 apply blast
using 2 apply blast
done
next
case False
then have "frontier (convex hull S) = (convex hull S) - rel_interior(convex hull S)"
apply (simp add: rel_boundary_of_convex_hull [symmetric] frontier_def)
by (metis aff_independent_finite assms closure_convex_hull finite_imp_compact_convex_hull hull_hull interior_convex_hull_eq_empty rel_interior_nonempty_interior)
also have "... = \<Union>{convex hull (S - {a}) |a. a \<in> S}"
proof -
have "convex hull S - rel_interior (convex hull S) = rel_frontier (convex hull S)"
by (simp add: False aff_independent_finite polyhedron_convex_hull rel_boundary_of_polyhedron rel_frontier_of_polyhedron)
then show ?thesis
by (simp add: False rel_frontier_convex_hull_cases)
qed
finally show ?thesis .
qed
subsection\<open>Special case of a triangle\<close>
proposition frontier_of_triangle:
fixes a :: "'a::euclidean_space"
assumes "DIM('a) = 2"
shows "frontier(convex hull {a,b,c}) = closed_segment a b \<union> closed_segment b c \<union> closed_segment c a"
(is "?lhs = ?rhs")
proof (cases "b = a \<or> c = a \<or> c = b")
case True then show ?thesis
by (auto simp: assms segment_convex_hull frontier_def empty_interior_convex_hull insert_commute card_insert_le_m1 hull_inc insert_absorb)
next
case False then have [simp]: "card {a, b, c} = Suc (DIM('a))"
by (simp add: card_insert Set.insert_Diff_if assms)
show ?thesis
proof
show "?lhs \<subseteq> ?rhs"
using False
by (force simp: segment_convex_hull frontier_of_convex_hull insert_Diff_if insert_commute split: if_split_asm)
show "?rhs \<subseteq> ?lhs"
using False
apply (simp add: frontier_of_convex_hull segment_convex_hull)
apply (intro conjI subsetI)
apply (rule_tac X="convex hull {a,b}" in UnionI; force simp: Set.insert_Diff_if)
apply (rule_tac X="convex hull {b,c}" in UnionI; force)
apply (rule_tac X="convex hull {a,c}" in UnionI; force simp: insert_commute Set.insert_Diff_if)
done
qed
qed
corollary inside_of_triangle:
fixes a :: "'a::euclidean_space"
assumes "DIM('a) = 2"
shows "inside (closed_segment a b \<union> closed_segment b c \<union> closed_segment c a) = interior(convex hull {a,b,c})"
by (metis assms frontier_of_triangle bounded_empty bounded_insert convex_convex_hull inside_frontier_eq_interior bounded_convex_hull)
corollary interior_of_triangle:
fixes a :: "'a::euclidean_space"
assumes "DIM('a) = 2"
shows "interior(convex hull {a,b,c}) =
convex hull {a,b,c} - (closed_segment a b \<union> closed_segment b c \<union> closed_segment c a)"
using interior_subset
by (force simp: frontier_of_triangle [OF assms, symmetric] frontier_def Diff_Diff_Int)
subsection\<open>Subdividing a cell complex\<close>
lemma subdivide_interval:
fixes x::real
assumes "a < \<bar>x - y\<bar>" "0 < a"
obtains n where "n \<in> \<int>" "x < n * a \<and> n * a < y \<or> y < n * a \<and> n * a < x"
proof -
consider "a + x < y" | "a + y < x"
using assms by linarith
then show ?thesis
proof cases
case 1
let ?n = "of_int (floor (x/a)) + 1"
have x: "x < ?n * a"
by (meson \<open>0 < a\<close> divide_less_eq floor_eq_iff)
have "?n * a \<le> a + x"
apply (simp add: algebra_simps)
by (metis \<open>0 < a\<close> floor_correct less_irrefl nonzero_mult_div_cancel_left real_mult_le_cancel_iff2 times_divide_eq_right)
also have "... < y"
by (rule 1)
finally have "?n * a < y" .
with x show ?thesis
using Ints_1 Ints_add Ints_of_int that by blast
next
case 2
let ?n = "of_int (floor (y/a)) + 1"
have y: "y < ?n * a"
by (meson \<open>0 < a\<close> divide_less_eq floor_eq_iff)
have "?n * a \<le> a + y"
apply (simp add: algebra_simps)
by (metis \<open>0 < a\<close> floor_correct less_irrefl nonzero_mult_div_cancel_left real_mult_le_cancel_iff2 times_divide_eq_right)
also have "... < x"
by (rule 2)
finally have "?n * a < x" .
then show ?thesis
using Ints_1 Ints_add Ints_of_int that y by blast
qed
qed
lemma cell_subdivision_lemma:
assumes "finite \<F>"
and "\<And>X. X \<in> \<F> \<Longrightarrow> polytope X"
and "\<And>X. X \<in> \<F> \<Longrightarrow> aff_dim X \<le> d"
and "\<And>X Y. \<lbrakk>X \<in> \<F>; Y \<in> \<F>\<rbrakk> \<Longrightarrow> (X \<inter> Y) face_of X \<and> (X \<inter> Y) face_of Y"
and "finite I"
shows "\<exists>\<G>. \<Union>\<G> = \<Union>\<F> \<and>
finite \<G> \<and>
(\<forall>C \<in> \<G>. \<exists>D. D \<in> \<F> \<and> C \<subseteq> D) \<and>
(\<forall>C \<in> \<F>. \<forall>x \<in> C. \<exists>D. D \<in> \<G> \<and> x \<in> D \<and> D \<subseteq> C) \<and>
(\<forall>X \<in> \<G>. polytope X) \<and>
(\<forall>X \<in> \<G>. aff_dim X \<le> d) \<and>
(\<forall>X \<in> \<G>. \<forall>Y \<in> \<G>. X \<inter> Y face_of X \<and> X \<inter> Y face_of Y) \<and>
(\<forall>X \<in> \<G>. \<forall>x \<in> X. \<forall>y \<in> X. \<forall>a b.
(a,b) \<in> I \<longrightarrow> a \<bullet> x \<le> b \<and> a \<bullet> y \<le> b \<or>
a \<bullet> x \<ge> b \<and> a \<bullet> y \<ge> b)"
using \<open>finite I\<close>
proof induction
case empty
then show ?case
by (rule_tac x="\<F>" in exI) (auto simp: assms)
next
case (insert ab I)
then obtain \<G> where eq: "\<Union>\<G> = \<Union>\<F>" and "finite \<G>"
and sub1: "\<And>C. C \<in> \<G> \<Longrightarrow> \<exists>D. D \<in> \<F> \<and> C \<subseteq> D"
and sub2: "\<And>C x. C \<in> \<F> \<and> x \<in> C \<Longrightarrow> \<exists>D. D \<in> \<G> \<and> x \<in> D \<and> D \<subseteq> C"
and poly: "\<And>X. X \<in> \<G> \<Longrightarrow> polytope X"
and aff: "\<And>X. X \<in> \<G> \<Longrightarrow> aff_dim X \<le> d"
and face: "\<And>X Y. \<lbrakk>X \<in> \<G>; Y \<in> \<G>\<rbrakk> \<Longrightarrow> X \<inter> Y face_of X \<and> X \<inter> Y face_of Y"
and I: "\<And>X x y a b. \<lbrakk>X \<in> \<G>; x \<in> X; y \<in> X; (a,b) \<in> I\<rbrakk> \<Longrightarrow>
a \<bullet> x \<le> b \<and> a \<bullet> y \<le> b \<or> a \<bullet> x \<ge> b \<and> a \<bullet> y \<ge> b"
by (auto simp: that)
obtain a b where "ab = (a,b)"
by fastforce
let ?\<G> = "(\<lambda>X. X \<inter> {x. a \<bullet> x \<le> b}) ` \<G> \<union> (\<lambda>X. X \<inter> {x. a \<bullet> x \<ge> b}) ` \<G>"
have eqInt: "(S \<inter> Collect P) \<inter> (T \<inter> Collect Q) = (S \<inter> T) \<inter> (Collect P \<inter> Collect Q)" for S T::"'a set" and P Q
by blast
show ?case
proof (intro conjI exI)
show "\<Union>?\<G> = \<Union>\<F>"
by (force simp: eq [symmetric])
show "finite ?\<G>"
using \<open>finite \<G>\<close> by force
show "\<forall>X \<in> ?\<G>. polytope X"
by (force simp: poly polytope_Int_polyhedron polyhedron_halfspace_le polyhedron_halfspace_ge)
show "\<forall>X \<in> ?\<G>. aff_dim X \<le> d"
by (auto; metis order_trans aff aff_dim_subset inf_le1)
show "\<forall>X \<in> ?\<G>. \<forall>x \<in> X. \<forall>y \<in> X. \<forall>a b.
(a,b) \<in> insert ab I \<longrightarrow> a \<bullet> x \<le> b \<and> a \<bullet> y \<le> b \<or>
a \<bullet> x \<ge> b \<and> a \<bullet> y \<ge> b"
using \<open>ab = (a, b)\<close> I by fastforce
show "\<forall>X \<in> ?\<G>. \<forall>Y \<in> ?\<G>. X \<inter> Y face_of X \<and> X \<inter> Y face_of Y"
by (auto simp: eqInt halfspace_Int_eq face_of_Int_Int face face_of_halfspace_le face_of_halfspace_ge)
show "\<forall>C \<in> ?\<G>. \<exists>D. D \<in> \<F> \<and> C \<subseteq> D"
using sub1 by force
show "\<forall>C\<in>\<F>. \<forall>x\<in>C. \<exists>D. D \<in> ?\<G> \<and> x \<in> D \<and> D \<subseteq> C"
proof (intro ballI)
fix C z
assume "C \<in> \<F>" "z \<in> C"
with sub2 obtain D where D: "D \<in> \<G>" "z \<in> D" "D \<subseteq> C" by blast
have "D \<in> \<G> \<and> z \<in> D \<inter> {x. a \<bullet> x \<le> b} \<and> D \<inter> {x. a \<bullet> x \<le> b} \<subseteq> C \<or>
D \<in> \<G> \<and> z \<in> D \<inter> {x. a \<bullet> x \<ge> b} \<and> D \<inter> {x. a \<bullet> x \<ge> b} \<subseteq> C"
using linorder_class.linear [of "a \<bullet> z" b] D by blast
then show "\<exists>D. D \<in> ?\<G> \<and> z \<in> D \<and> D \<subseteq> C"
by blast
qed
qed
qed
proposition cell_complex_subdivision_exists:
fixes \<F> :: "'a::euclidean_space set set"
assumes "0 < e" "finite \<F>"
and poly: "\<And>X. X \<in> \<F> \<Longrightarrow> polytope X"
and aff: "\<And>X. X \<in> \<F> \<Longrightarrow> aff_dim X \<le> d"
and face: "\<And>X Y. \<lbrakk>X \<in> \<F>; Y \<in> \<F>\<rbrakk> \<Longrightarrow> X \<inter> Y face_of X \<and> X \<inter> Y face_of Y"
obtains "\<F>'" where "finite \<F>'" "\<Union>\<F>' = \<Union>\<F>" "\<And>X. X \<in> \<F>' \<Longrightarrow> diameter X < e"
"\<And>X. X \<in> \<F>' \<Longrightarrow> polytope X" "\<And>X. X \<in> \<F>' \<Longrightarrow> aff_dim X \<le> d"
"\<And>X Y. \<lbrakk>X \<in> \<F>'; Y \<in> \<F>'\<rbrakk> \<Longrightarrow> X \<inter> Y face_of X \<and> X \<inter> Y face_of Y"
"\<And>C. C \<in> \<F>' \<Longrightarrow> \<exists>D. D \<in> \<F> \<and> C \<subseteq> D"
"\<And>C x. C \<in> \<F> \<and> x \<in> C \<Longrightarrow> \<exists>D. D \<in> \<F>' \<and> x \<in> D \<and> D \<subseteq> C"
proof -
have "bounded(\<Union>\<F>)"
by (simp add: \<open>finite \<F>\<close> poly bounded_Union polytope_imp_bounded)
then obtain B where "B > 0" and B: "\<And>x. x \<in> \<Union>\<F> \<Longrightarrow> norm x < B"
by (meson bounded_pos_less)
define C where "C \<equiv> {z \<in> \<int>. \<bar>z * e / 2 / real DIM('a)\<bar> \<le> B}"
define I where "I \<equiv> \<Union>i \<in> Basis. \<Union>j \<in> C. { (i::'a, j * e / 2 / DIM('a)) }"
have "finite C"
using finite_int_segment [of "-B / (e / 2 / DIM('a))" "B / (e / 2 / DIM('a))"]
apply (simp add: C_def)
apply (erule rev_finite_subset)
using \<open>0 < e\<close>
apply (auto simp: field_split_simps)
done
then have "finite I"
by (simp add: I_def)
obtain \<F>' where eq: "\<Union>\<F>' = \<Union>\<F>" and "finite \<F>'"
and poly: "\<And>X. X \<in> \<F>' \<Longrightarrow> polytope X"
and aff: "\<And>X. X \<in> \<F>' \<Longrightarrow> aff_dim X \<le> d"
and face: "\<And>X Y. \<lbrakk>X \<in> \<F>'; Y \<in> \<F>'\<rbrakk> \<Longrightarrow> X \<inter> Y face_of X \<and> X \<inter> Y face_of Y"
and I: "\<And>X x y a b. \<lbrakk>X \<in> \<F>'; x \<in> X; y \<in> X; (a,b) \<in> I\<rbrakk> \<Longrightarrow>
a \<bullet> x \<le> b \<and> a \<bullet> y \<le> b \<or> a \<bullet> x \<ge> b \<and> a \<bullet> y \<ge> b"
and sub1: "\<And>C. C \<in> \<F>' \<Longrightarrow> \<exists>D. D \<in> \<F> \<and> C \<subseteq> D"
and sub2: "\<And>C x. C \<in> \<F> \<and> x \<in> C \<Longrightarrow> \<exists>D. D \<in> \<F>' \<and> x \<in> D \<and> D \<subseteq> C"
apply (rule exE [OF cell_subdivision_lemma])
using assms \<open>finite I\<close> apply auto
done
show ?thesis
proof (rule_tac \<F>'="\<F>'" in that)
show "diameter X < e" if "X \<in> \<F>'" for X
proof -
have "diameter X \<le> e/2"
proof (rule diameter_le)
show "norm (x - y) \<le> e / 2" if "x \<in> X" "y \<in> X" for x y
proof -
have "norm x < B" "norm y < B"
using B \<open>X \<in> \<F>'\<close> eq that by fastforce+
have "norm (x - y) \<le> (\<Sum>b\<in>Basis. \<bar>(x-y) \<bullet> b\<bar>)"
by (rule norm_le_l1)
also have "... \<le> of_nat (DIM('a)) * (e / 2 / DIM('a))"
proof (rule sum_bounded_above)
fix i::'a
assume "i \<in> Basis"
then have I': "\<And>z b. \<lbrakk>z \<in> C; b = z * e / (2 * real DIM('a))\<rbrakk> \<Longrightarrow> i \<bullet> x \<le> b \<and> i \<bullet> y \<le> b \<or> i \<bullet> x \<ge> b \<and> i \<bullet> y \<ge> b"
using I \<open>X \<in> \<F>'\<close> that
by (fastforce simp: I_def)
show "\<bar>(x - y) \<bullet> i\<bar> \<le> e / 2 / real DIM('a)"
proof (rule ccontr)
assume "\<not> \<bar>(x - y) \<bullet> i\<bar> \<le> e / 2 / real DIM('a)"
then have xyi: "\<bar>i \<bullet> x - i \<bullet> y\<bar> > e / 2 / real DIM('a)"
by (simp add: inner_commute inner_diff_right)
obtain n where "n \<in> \<int>" and n: "i \<bullet> x < n * (e / 2 / real DIM('a)) \<and> n * (e / 2 / real DIM('a)) < i \<bullet> y \<or> i \<bullet> y < n * (e / 2 / real DIM('a)) \<and> n * (e / 2 / real DIM('a)) < i \<bullet> x"
using subdivide_interval [OF xyi] DIM_positive \<open>0 < e\<close>
by (auto simp: zero_less_divide_iff)
have "\<bar>i \<bullet> x\<bar> < B"
by (metis \<open>i \<in> Basis\<close> \<open>norm x < B\<close> inner_commute norm_bound_Basis_lt)
have "\<bar>i \<bullet> y\<bar> < B"
by (metis \<open>i \<in> Basis\<close> \<open>norm y < B\<close> inner_commute norm_bound_Basis_lt)
have *: "\<bar>n * e\<bar> \<le> B * (2 * real DIM('a))"
if "\<bar>ix\<bar> < B" "\<bar>iy\<bar> < B"
and ix: "ix * (2 * real DIM('a)) < n * e"
and iy: "n * e < iy * (2 * real DIM('a))" for ix iy
proof (rule abs_leI)
have "iy * (2 * real DIM('a)) \<le> B * (2 * real DIM('a))"
by (rule mult_right_mono) (use \<open>\<bar>iy\<bar> < B\<close> in linarith)+
then show "n * e \<le> B * (2 * real DIM('a))"
using iy by linarith
next
have "- ix * (2 * real DIM('a)) \<le> B * (2 * real DIM('a))"
by (rule mult_right_mono) (use \<open>\<bar>ix\<bar> < B\<close> in linarith)+
then show "- (n * e) \<le> B * (2 * real DIM('a))"
using ix by linarith
qed
have "n \<in> C"
using \<open>n \<in> \<int>\<close> n by (auto simp: C_def divide_simps intro: * \<open>\<bar>i \<bullet> x\<bar> < B\<close> \<open>\<bar>i \<bullet> y\<bar> < B\<close>)
show False
using I' [OF \<open>n \<in> C\<close> refl] n by auto
qed
qed
also have "... = e / 2"
by simp
finally show ?thesis .
qed
qed (use \<open>0 < e\<close> in force)
also have "... < e"
by (simp add: \<open>0 < e\<close>)
finally show ?thesis .
qed
qed (auto simp: eq poly aff face sub1 sub2 \<open>finite \<F>'\<close>)
qed
subsection\<open>Simplexes\<close>
text\<open>The notion of n-simplex for integer \<^term>\<open>n \<ge> -1\<close>\<close>
definition\<^marker>\<open>tag important\<close> simplex :: "int \<Rightarrow> 'a::euclidean_space set \<Rightarrow> bool" (infix "simplex" 50)
where "n simplex S \<equiv> \<exists>C. \<not> affine_dependent C \<and> int(card C) = n + 1 \<and> S = convex hull C"
lemma simplex:
"n simplex S \<longleftrightarrow> (\<exists>C. finite C \<and>
\<not> affine_dependent C \<and>
int(card C) = n + 1 \<and>
S = convex hull C)"
by (auto simp add: simplex_def intro: aff_independent_finite)
lemma simplex_convex_hull:
"\<not> affine_dependent C \<and> int(card C) = n + 1 \<Longrightarrow> n simplex (convex hull C)"
by (auto simp add: simplex_def)
lemma convex_simplex: "n simplex S \<Longrightarrow> convex S"
by (metis convex_convex_hull simplex_def)
lemma compact_simplex: "n simplex S \<Longrightarrow> compact S"
unfolding simplex
using finite_imp_compact_convex_hull by blast
lemma closed_simplex: "n simplex S \<Longrightarrow> closed S"
by (simp add: compact_imp_closed compact_simplex)
lemma simplex_imp_polytope:
"n simplex S \<Longrightarrow> polytope S"
unfolding simplex_def polytope_def
using aff_independent_finite by blast
lemma simplex_imp_polyhedron:
"n simplex S \<Longrightarrow> polyhedron S"
by (simp add: polytope_imp_polyhedron simplex_imp_polytope)
lemma simplex_dim_ge: "n simplex S \<Longrightarrow> -1 \<le> n"
by (metis (no_types, hide_lams) aff_dim_geq affine_independent_iff_card diff_add_cancel diff_diff_eq2 simplex_def)
lemma simplex_empty [simp]: "n simplex {} \<longleftrightarrow> n = -1"
proof
assume "n simplex {}"
then show "n = -1"
unfolding simplex by (metis card_empty convex_hull_eq_empty diff_0 diff_eq_eq of_nat_0)
next
assume "n = -1" then show "n simplex {}"
by (fastforce simp: simplex)
qed
lemma simplex_minus_1 [simp]: "-1 simplex S \<longleftrightarrow> S = {}"
by (metis simplex cancel_comm_monoid_add_class.diff_cancel card_0_eq diff_minus_eq_add of_nat_eq_0_iff simplex_empty)
lemma aff_dim_simplex:
"n simplex S \<Longrightarrow> aff_dim S = n"
by (metis simplex add.commute add_diff_cancel_left' aff_dim_convex_hull affine_independent_iff_card)
lemma zero_simplex_sing: "0 simplex {a}"
apply (simp add: simplex_def)
by (metis affine_independent_1 card_empty card_insert_disjoint convex_hull_singleton empty_iff finite.emptyI)
lemma simplex_sing [simp]: "n simplex {a} \<longleftrightarrow> n = 0"
using aff_dim_simplex aff_dim_sing zero_simplex_sing by blast
lemma simplex_zero: "0 simplex S \<longleftrightarrow> (\<exists>a. S = {a})"
apply (auto simp: )
using aff_dim_eq_0 aff_dim_simplex by blast
lemma one_simplex_segment: "a \<noteq> b \<Longrightarrow> 1 simplex closed_segment a b"
apply (simp add: simplex_def)
apply (rule_tac x="{a,b}" in exI)
apply (auto simp: segment_convex_hull)
done
lemma simplex_segment_cases:
"(if a = b then 0 else 1) simplex closed_segment a b"
by (auto simp: one_simplex_segment)
lemma simplex_segment:
"\<exists>n. n simplex closed_segment a b"
using simplex_segment_cases by metis
lemma polytope_lowdim_imp_simplex:
assumes "polytope P" "aff_dim P \<le> 1"
obtains n where "n simplex P"
proof (cases "P = {}")
case True
then show ?thesis
by (simp add: that)
next
case False
then show ?thesis
by (metis assms compact_convex_collinear_segment collinear_aff_dim polytope_imp_compact polytope_imp_convex simplex_segment_cases that)
qed
lemma simplex_insert_dimplus1:
fixes n::int
assumes "n simplex S" and a: "a \<notin> affine hull S"
shows "(n+1) simplex (convex hull (insert a S))"
proof -
obtain C where C: "finite C" "\<not> affine_dependent C" "int(card C) = n+1" and S: "S = convex hull C"
using assms unfolding simplex by force
show ?thesis
unfolding simplex
proof (intro exI conjI)
have "aff_dim S = n"
using aff_dim_simplex assms(1) by blast
moreover have "a \<notin> affine hull C"
using S a affine_hull_convex_hull by blast
moreover have "a \<notin> C"
using S a hull_inc by fastforce
ultimately show "\<not> affine_dependent (insert a C)"
by (simp add: C S aff_dim_convex_hull aff_dim_insert affine_independent_iff_card)
next
have "a \<notin> C"
using S a hull_inc by fastforce
then show "int (card (insert a C)) = n + 1 + 1"
by (simp add: C)
next
show "convex hull insert a S = convex hull (insert a C)"
by (simp add: S convex_hull_insert_segments)
qed (use C in auto)
qed
subsection \<open>Simplicial complexes and triangulations\<close>
definition\<^marker>\<open>tag important\<close> simplicial_complex where
"simplicial_complex \<C> \<equiv>
finite \<C> \<and>
(\<forall>S \<in> \<C>. \<exists>n. n simplex S) \<and>
(\<forall>F S. S \<in> \<C> \<and> F face_of S \<longrightarrow> F \<in> \<C>) \<and>
(\<forall>S S'. S \<in> \<C> \<and> S' \<in> \<C>
\<longrightarrow> (S \<inter> S') face_of S \<and> (S \<inter> S') face_of S')"
definition\<^marker>\<open>tag important\<close> triangulation where
"triangulation \<T> \<equiv>
finite \<T> \<and>
(\<forall>T \<in> \<T>. \<exists>n. n simplex T) \<and>
(\<forall>T T'. T \<in> \<T> \<and> T' \<in> \<T>
\<longrightarrow> (T \<inter> T') face_of T \<and> (T \<inter> T') face_of T')"
subsection\<open>Refining a cell complex to a simplicial complex\<close>
proposition convex_hull_insert_Int_eq:
fixes z :: "'a :: euclidean_space"
assumes z: "z \<in> rel_interior S"
and T: "T \<subseteq> rel_frontier S"
and U: "U \<subseteq> rel_frontier S"
and "convex S" "convex T" "convex U"
shows "convex hull (insert z T) \<inter> convex hull (insert z U) = convex hull (insert z (T \<inter> U))"
(is "?lhs = ?rhs")
proof
show "?lhs \<subseteq> ?rhs"
proof (cases "T={} \<or> U={}")
case True then show ?thesis by auto
next
case False
then have "T \<noteq> {}" "U \<noteq> {}" by auto
have TU: "convex (T \<inter> U)"
by (simp add: \<open>convex T\<close> \<open>convex U\<close> convex_Int)
have "(\<Union>x\<in>T. closed_segment z x) \<inter> (\<Union>x\<in>U. closed_segment z x)
\<subseteq> (if T \<inter> U = {} then {z} else \<Union>((closed_segment z) ` (T \<inter> U)))" (is "_ \<subseteq> ?IF")
proof clarify
fix x t u
assume xt: "x \<in> closed_segment z t"
and xu: "x \<in> closed_segment z u"
and "t \<in> T" "u \<in> U"
then have ne: "t \<noteq> z" "u \<noteq> z"
using T U z unfolding rel_frontier_def by blast+
show "x \<in> ?IF"
proof (cases "x = z")
case True then show ?thesis by auto
next
case False
have t: "t \<in> closure S"
using T \<open>t \<in> T\<close> rel_frontier_def by auto
have u: "u \<in> closure S"
using U \<open>u \<in> U\<close> rel_frontier_def by auto
show ?thesis
proof (cases "t = u")
case True
then show ?thesis
using \<open>t \<in> T\<close> \<open>u \<in> U\<close> xt by auto
next
case False
have tnot: "t \<notin> closed_segment u z"
proof -
have "t \<in> closure S - rel_interior S"
using T \<open>t \<in> T\<close> rel_frontier_def by blast
then have "t \<notin> open_segment z u"
by (meson DiffD2 rel_interior_closure_convex_segment [OF \<open>convex S\<close> z u] subsetD)
then show ?thesis
by (simp add: \<open>t \<noteq> u\<close> \<open>t \<noteq> z\<close> open_segment_commute open_segment_def)
qed
moreover have "u \<notin> closed_segment z t"
using rel_interior_closure_convex_segment [OF \<open>convex S\<close> z t] \<open>u \<in> U\<close> \<open>u \<noteq> z\<close>
U [unfolded rel_frontier_def] tnot
by (auto simp: closed_segment_eq_open)
ultimately
have "\<not>(between (t,u) z | between (u,z) t | between (z,t) u)" if "x \<noteq> z"
using that xt xu
apply (simp add: between_mem_segment [symmetric])
by (metis between_commute between_trans_2 between_antisym)
then have "\<not> collinear {t, z, u}" if "x \<noteq> z"
by (auto simp: that collinear_between_cases between_commute)
moreover have "collinear {t, z, x}"
by (metis closed_segment_commute collinear_2 collinear_closed_segment collinear_triples ends_in_segment(1) insert_absorb insert_absorb2 xt)
moreover have "collinear {z, x, u}"
by (metis closed_segment_commute collinear_2 collinear_closed_segment collinear_triples ends_in_segment(1) insert_absorb insert_absorb2 xu)
ultimately have False
using collinear_3_trans [of t z x u] \<open>x \<noteq> z\<close> by blast
then show ?thesis by metis
qed
qed
qed
then show ?thesis
using False \<open>convex T\<close> \<open>convex U\<close> TU
by (simp add: convex_hull_insert_segments hull_same split: if_split_asm)
qed
show "?rhs \<subseteq> ?lhs"
by (metis inf_greatest hull_mono inf.cobounded1 inf.cobounded2 insert_mono)
qed
lemma simplicial_subdivision_aux:
assumes "finite \<M>"
and "\<And>C. C \<in> \<M> \<Longrightarrow> polytope C"
and "\<And>C. C \<in> \<M> \<Longrightarrow> aff_dim C \<le> of_nat n"
and "\<And>C F. \<lbrakk>C \<in> \<M>; F face_of C\<rbrakk> \<Longrightarrow> F \<in> \<M>"
and "\<And>C1 C2. \<lbrakk>C1 \<in> \<M>; C2 \<in> \<M>\<rbrakk> \<Longrightarrow> C1 \<inter> C2 face_of C1 \<and> C1 \<inter> C2 face_of C2"
shows "\<exists>\<T>. simplicial_complex \<T> \<and>
(\<forall>K \<in> \<T>. aff_dim K \<le> of_nat n) \<and>
\<Union>\<T> = \<Union>\<M> \<and>
(\<forall>C \<in> \<M>. \<exists>F. finite F \<and> F \<subseteq> \<T> \<and> C = \<Union>F) \<and>
(\<forall>K \<in> \<T>. \<exists>C. C \<in> \<M> \<and> K \<subseteq> C)"
using assms
proof (induction n arbitrary: \<M> rule: less_induct)
case (less n)
then have poly\<M>: "\<And>C. C \<in> \<M> \<Longrightarrow> polytope C"
and aff\<M>: "\<And>C. C \<in> \<M> \<Longrightarrow> aff_dim C \<le> of_nat n"
and face\<M>: "\<And>C F. \<lbrakk>C \<in> \<M>; F face_of C\<rbrakk> \<Longrightarrow> F \<in> \<M>"
and intface\<M>: "\<And>C1 C2. \<lbrakk>C1 \<in> \<M>; C2 \<in> \<M>\<rbrakk> \<Longrightarrow> C1 \<inter> C2 face_of C1 \<and> C1 \<inter> C2 face_of C2"
by metis+
show ?case
proof (cases "n \<le> 1")
case True
have "\<And>s. \<lbrakk>n \<le> 1; s \<in> \<M>\<rbrakk> \<Longrightarrow> \<exists>m. m simplex s"
using poly\<M> aff\<M> by (force intro: polytope_lowdim_imp_simplex)
then show ?thesis
unfolding simplicial_complex_def
apply (rule_tac x="\<M>" in exI)
using True by (auto simp: less.prems)
next
case False
define \<S> where "\<S> \<equiv> {C \<in> \<M>. aff_dim C < n}"
have "finite \<S>" "\<And>C. C \<in> \<S> \<Longrightarrow> polytope C" "\<And>C. C \<in> \<S> \<Longrightarrow> aff_dim C \<le> int (n - 1)"
"\<And>C F. \<lbrakk>C \<in> \<S>; F face_of C\<rbrakk> \<Longrightarrow> F \<in> \<S>"
"\<And>C1 C2. \<lbrakk>C1 \<in> \<S>; C2 \<in> \<S>\<rbrakk> \<Longrightarrow> C1 \<inter> C2 face_of C1 \<and> C1 \<inter> C2 face_of C2"
using less.prems
apply (auto simp: \<S>_def)
by (metis aff_dim_subset face_of_imp_subset less_le not_le)
with less.IH [of "n-1" \<S>] False
obtain \<U> where "simplicial_complex \<U>"
and aff_dim\<U>: "\<And>K. K \<in> \<U> \<Longrightarrow> aff_dim K \<le> int (n - 1)"
and "\<Union>\<U> = \<Union>\<S>"
and fin\<U>: "\<And>C. C \<in> \<S> \<Longrightarrow> \<exists>F. finite F \<and> F \<subseteq> \<U> \<and> C = \<Union>F"
and C\<U>: "\<And>K. K \<in> \<U> \<Longrightarrow> \<exists>C. C \<in> \<S> \<and> K \<subseteq> C"
by auto
then have "finite \<U>"
and simpl\<U>: "\<And>S. S \<in> \<U> \<Longrightarrow> \<exists>n. n simplex S"
and face\<U>: "\<And>F S. \<lbrakk>S \<in> \<U>; F face_of S\<rbrakk> \<Longrightarrow> F \<in> \<U>"
and faceI\<U>: "\<And>S S'. \<lbrakk>S \<in> \<U>; S' \<in> \<U>\<rbrakk> \<Longrightarrow> (S \<inter> S') face_of S \<and> (S \<inter> S') face_of S'"
by (auto simp: simplicial_complex_def)
define \<N> where "\<N> \<equiv> {C \<in> \<M>. aff_dim C = n}"
have "finite \<N>"
by (simp add: \<N>_def less.prems(1))
have poly\<N>: "\<And>C. C \<in> \<N> \<Longrightarrow> polytope C"
and convex\<N>: "\<And>C. C \<in> \<N> \<Longrightarrow> convex C"
and closed\<N>: "\<And>C. C \<in> \<N> \<Longrightarrow> closed C"
by (auto simp: \<N>_def poly\<M> polytope_imp_convex polytope_imp_closed)
have in_rel_interior: "(SOME z. z \<in> rel_interior C) \<in> rel_interior C" if "C \<in> \<N>" for C
using that poly\<M> polytope_imp_convex rel_interior_aff_dim some_in_eq by (fastforce simp: \<N>_def)
have *: "\<exists>T. \<not> affine_dependent T \<and> card T \<le> n \<and> aff_dim K < n \<and> K = convex hull T"
if "K \<in> \<U>" for K
proof -
obtain r where r: "r simplex K"
using \<open>K \<in> \<U>\<close> simpl\<U> by blast
have "r = aff_dim K"
using \<open>r simplex K\<close> aff_dim_simplex by blast
with r
show ?thesis
unfolding simplex_def
using False \<open>\<And>K. K \<in> \<U> \<Longrightarrow> aff_dim K \<le> int (n - 1)\<close> that by fastforce
qed
have ahK_C_disjoint: "affine hull K \<inter> rel_interior C = {}"
if "C \<in> \<N>" "K \<in> \<U>" "K \<subseteq> rel_frontier C" for C K
proof -
have "convex C" "closed C"
by (auto simp: convex\<N> closed\<N> \<open>C \<in> \<N>\<close>)
obtain F where F: "F face_of C" and "F \<noteq> C" "K \<subseteq> F"
proof -
obtain L where "L \<in> \<S>" "K \<subseteq> L"
using \<open>K \<in> \<U>\<close> C\<U> by blast
have "K \<le> rel_frontier C"
by (simp add: \<open>K \<subseteq> rel_frontier C\<close>)
also have "... \<le> C"
by (simp add: \<open>closed C\<close> rel_frontier_def subset_iff)
finally have "K \<subseteq> C" .
have "L \<inter> C face_of C"
using \<N>_def \<S>_def \<open>C \<in> \<N>\<close> \<open>L \<in> \<S>\<close> intface\<M> by auto
moreover have "L \<inter> C \<noteq> C"
using \<open>C \<in> \<N>\<close> \<open>L \<in> \<S>\<close>
apply (clarsimp simp: \<N>_def \<S>_def)
by (metis aff_dim_subset inf_le1 not_le)
moreover have "K \<subseteq> L \<inter> C"
using \<open>C \<in> \<N>\<close> \<open>L \<in> \<S>\<close> \<open>K \<subseteq> C\<close> \<open>K \<subseteq> L\<close>
by (auto simp: \<N>_def \<S>_def)
ultimately show ?thesis using that by metis
qed
have "affine hull F \<inter> rel_interior C = {}"
by (rule affine_hull_face_of_disjoint_rel_interior [OF \<open>convex C\<close> F \<open>F \<noteq> C\<close>])
with hull_mono [OF \<open>K \<subseteq> F\<close>]
show "affine hull K \<inter> rel_interior C = {}"
by fastforce
qed
let ?\<T> = "(\<Union>C \<in> \<N>. \<Union>K \<in> \<U> \<inter> Pow (rel_frontier C).
{convex hull (insert (SOME z. z \<in> rel_interior C) K)})"
have "\<exists>\<T>. simplicial_complex \<T> \<and>
(\<forall>K \<in> \<T>. aff_dim K \<le> of_nat n) \<and>
(\<forall>C \<in> \<M>. \<exists>F. F \<subseteq> \<T> \<and> C = \<Union>F) \<and>
(\<forall>K \<in> \<T>. \<exists>C. C \<in> \<M> \<and> K \<subseteq> C)"
proof (rule exI, intro conjI ballI)
show "simplicial_complex (\<U> \<union> ?\<T>)"
unfolding simplicial_complex_def
proof (intro conjI impI ballI allI)
show "finite (\<U> \<union> ?\<T>)"
using \<open>finite \<U>\<close> \<open>finite \<N>\<close> by simp
show "\<exists>n. n simplex S" if "S \<in> \<U> \<union> ?\<T>" for S
using that ahK_C_disjoint in_rel_interior simpl\<U> simplex_insert_dimplus1 by fastforce
show "F \<in> \<U> \<union> ?\<T>" if S: "S \<in> \<U> \<union> ?\<T> \<and> F face_of S" for F S
proof -
have "F \<in> \<U>" if "S \<in> \<U>"
using S face\<U> that by blast
moreover have "F \<in> \<U> \<union> ?\<T>"
if "F face_of S" "C \<in> \<N>" "K \<in> \<U>" and "K \<subseteq> rel_frontier C"
and S: "S = convex hull insert (SOME z. z \<in> rel_interior C) K" for C K
proof -
let ?z = "SOME z. z \<in> rel_interior C"
have "?z \<in> rel_interior C"
by (simp add: in_rel_interior \<open>C \<in> \<N>\<close>)
moreover
obtain I where "\<not> affine_dependent I" "card I \<le> n" "aff_dim K < int n" "K = convex hull I"
using * [OF \<open>K \<in> \<U>\<close>] by auto
ultimately have "?z \<notin> affine hull I"
using ahK_C_disjoint affine_hull_convex_hull that by blast
have "compact I" "finite I"
by (auto simp: \<open>\<not> affine_dependent I\<close> aff_independent_finite finite_imp_compact)
moreover have "F face_of convex hull insert ?z I"
by (metis S \<open>F face_of S\<close> \<open>K = convex hull I\<close> convex_hull_eq_empty convex_hull_insert_segments hull_hull)
ultimately obtain J where "J \<subseteq> insert ?z I" "F = convex hull J"
using face_of_convex_hull_subset [of "insert ?z I" F] by auto
show ?thesis
proof (cases "?z \<in> J")
case True
have "F \<in> (\<Union>K\<in>\<U> \<inter> Pow (rel_frontier C). {convex hull insert ?z K})"
proof
have "convex hull (J - {?z}) face_of K"
by (metis True \<open>J \<subseteq> insert ?z I\<close> \<open>K = convex hull I\<close> \<open>\<not> affine_dependent I\<close> face_of_convex_hull_affine_independent subset_insert_iff)
then have "convex hull (J - {?z}) \<in> \<U>"
by (rule face\<U> [OF \<open>K \<in> \<U>\<close>])
moreover
have "\<And>x. x \<in> convex hull (J - {?z}) \<Longrightarrow> x \<in> rel_frontier C"
by (metis True \<open>J \<subseteq> insert ?z I\<close> \<open>K = convex hull I\<close> subsetD hull_mono subset_insert_iff that(4))
ultimately show "convex hull (J - {?z}) \<in> \<U> \<inter> Pow (rel_frontier C)" by auto
let ?F = "convex hull insert ?z (convex hull (J - {?z}))"
have "F \<subseteq> ?F"
apply (clarsimp simp: \<open>F = convex hull J\<close>)
by (metis True subsetD hull_mono hull_subset subset_insert_iff)
moreover have "?F \<subseteq> F"
apply (clarsimp simp: \<open>F = convex hull J\<close>)
by (metis (no_types, lifting) True convex_hull_eq_empty convex_hull_insert_segments hull_hull insert_Diff)
ultimately
show "F \<in> {?F}" by auto
qed
with \<open>C\<in>\<N>\<close> show ?thesis by auto
next
case False
then have "F \<in> \<U>"
using face_of_convex_hull_affine_independent [OF \<open>\<not> affine_dependent I\<close>]
by (metis Int_absorb2 Int_insert_right_if0 \<open>F = convex hull J\<close> \<open>J \<subseteq> insert ?z I\<close> \<open>K = convex hull I\<close> face\<U> inf_le2 \<open>K \<in> \<U>\<close>)
then show "F \<in> \<U> \<union> ?\<T>"
by blast
qed
qed
ultimately show ?thesis
using that by auto
qed
have "(S \<inter> S' face_of S) \<and> (S \<inter> S' face_of S')"
if "S \<in> \<U> \<union> ?\<T>" "S' \<in> \<U> \<union> ?\<T>" for S S'
proof -
have symmy: "\<lbrakk>\<And>X Y. R X Y \<Longrightarrow> R Y X;
\<And>X Y. \<lbrakk>X \<in> \<U>; Y \<in> \<U>\<rbrakk> \<Longrightarrow> R X Y;
\<And>X Y. \<lbrakk>X \<in> \<U>; Y \<in> ?\<T>\<rbrakk> \<Longrightarrow> R X Y;
\<And>X Y. \<lbrakk>X \<in> ?\<T>; Y \<in> ?\<T>\<rbrakk> \<Longrightarrow> R X Y\<rbrakk> \<Longrightarrow> R S S'" for R
using that by (metis (no_types, lifting) Un_iff)
show ?thesis
proof (rule symmy)
show "Y \<inter> X face_of Y \<and> Y \<inter> X face_of X"
if "X \<inter> Y face_of X \<and> X \<inter> Y face_of Y" for X Y :: "'a set"
by (simp add: inf_commute that)
next
show "X \<inter> Y face_of X \<and> X \<inter> Y face_of Y"
if "X \<in> \<U>" and "Y \<in> \<U>" for X Y
by (simp add: faceI\<U> that)
next
show "X \<inter> Y face_of X \<and> X \<inter> Y face_of Y"
if XY: "X \<in> \<U>" "Y \<in> ?\<T>" for X Y
proof -
obtain C K
where "C \<in> \<N>" "K \<in> \<U>" "K \<subseteq> rel_frontier C"
and Y: "Y = convex hull insert (SOME z. z \<in> rel_interior C) K"
using XY by blast
have "convex C"
by (simp add: \<open>C \<in> \<N>\<close> convex\<N>)
have "K \<subseteq> C"
by (metis DiffE \<open>C \<in> \<N>\<close> \<open>K \<subseteq> rel_frontier C\<close> closed\<N> closure_closed rel_frontier_def subset_iff)
let ?z = "(SOME z. z \<in> rel_interior C)"
have z: "?z \<in> rel_interior C"
using \<open>C \<in> \<N>\<close> in_rel_interior by blast
obtain D where "D \<in> \<S>" "X \<subseteq> D"
using C\<U> \<open>X \<in> \<U>\<close> by blast
have "D \<inter> rel_interior C = (C \<inter> D) \<inter> rel_interior C"
using rel_interior_subset by blast
also have "(C \<inter> D) \<inter> rel_interior C = {}"
proof (rule face_of_disjoint_rel_interior)
show "C \<inter> D face_of C"
using \<N>_def \<S>_def \<open>C \<in> \<N>\<close> \<open>D \<in> \<S>\<close> intface\<M> by blast
show "C \<inter> D \<noteq> C"
by (metis (mono_tags, lifting) Int_lower2 \<N>_def \<S>_def \<open>C \<in> \<N>\<close> \<open>D \<in> \<S>\<close> aff_dim_subset mem_Collect_eq not_le)
qed
finally have DC: "D \<inter> rel_interior C = {}" .
have eq: "X \<inter> convex hull (insert ?z K) = X \<inter> convex hull K"
apply (rule Int_convex_hull_insert_rel_exterior [OF \<open>convex C\<close> \<open>K \<subseteq> C\<close> z])
using DC by (meson \<open>X \<subseteq> D\<close> disjnt_def disjnt_subset1)
obtain I where I: "\<not> affine_dependent I"
and Keq: "K = convex hull I" and [simp]: "convex hull K = K"
using "*" \<open>K \<in> \<U>\<close> by force
then have "?z \<notin> affine hull I"
using ahK_C_disjoint \<open>C \<in> \<N>\<close> \<open>K \<in> \<U>\<close> \<open>K \<subseteq> rel_frontier C\<close> affine_hull_convex_hull z by blast
have "X \<inter> K face_of K"
by (simp add: \<open>K \<in> \<U>\<close> faceI\<U> \<open>X \<in> \<U>\<close>)
also have "... face_of convex hull insert ?z K"
by (metis I Keq \<open>?z \<notin> affine hull I\<close> aff_independent_finite convex_convex_hull face_of_convex_hull_insert face_of_refl hull_insert)
finally have "X \<inter> K face_of convex hull insert ?z K" .
then show ?thesis
using "*" \<open>K \<in> \<U>\<close> faceI\<U> that(1) by (fastforce simp add: Y eq)
qed
next
show "X \<inter> Y face_of X \<and> X \<inter> Y face_of Y"
if XY: "X \<in> ?\<T>" "Y \<in> ?\<T>" for X Y
proof -
obtain C K D L
where "C \<in> \<N>" "K \<in> \<U>" "K \<subseteq> rel_frontier C"
and X: "X = convex hull insert (SOME z. z \<in> rel_interior C) K"
and "D \<in> \<N>" "L \<in> \<U>" "L \<subseteq> rel_frontier D"
and Y: "Y = convex hull insert (SOME z. z \<in> rel_interior D) L"
using XY by blast
let ?z = "(SOME z. z \<in> rel_interior C)"
have z: "?z \<in> rel_interior C"
using \<open>C \<in> \<N>\<close> in_rel_interior by blast
have "convex C"
by (simp add: \<open>C \<in> \<N>\<close> convex\<N>)
have "convex K"
using "*" \<open>K \<in> \<U>\<close> by blast
have "convex L"
by (meson \<open>L \<in> \<U>\<close> convex_simplex simpl\<U>)
show ?thesis
proof (cases "D=C")
case True
then have "L \<subseteq> rel_frontier C"
using \<open>L \<subseteq> rel_frontier D\<close> by auto
show ?thesis
apply (simp add: X Y True)
apply (simp add: convex_hull_insert_Int_eq [OF z] \<open>K \<subseteq> rel_frontier C\<close> \<open>L \<subseteq> rel_frontier C\<close> \<open>convex C\<close> \<open>convex K\<close> \<open>convex L\<close>)
using face_of_polytope_insert2
by (metis "*" IntI \<open>C \<in> \<N>\<close> \<open>K \<in> \<U>\<close> \<open>L \<in> \<U>\<close>\<open>K \<subseteq> rel_frontier C\<close> \<open>L \<subseteq> rel_frontier C\<close> aff_independent_finite ahK_C_disjoint empty_iff faceI\<U> polytope_convex_hull z)
next
case False
have "convex D"
by (simp add: \<open>D \<in> \<N>\<close> convex\<N>)
have "K \<subseteq> C"
by (metis DiffE \<open>C \<in> \<N>\<close> \<open>K \<subseteq> rel_frontier C\<close> closed\<N> closure_closed rel_frontier_def subset_eq)
have "L \<subseteq> D"
by (metis DiffE \<open>D \<in> \<N>\<close> \<open>L \<subseteq> rel_frontier D\<close> closed\<N> closure_closed rel_frontier_def subset_eq)
let ?w = "(SOME w. w \<in> rel_interior D)"
have w: "?w \<in> rel_interior D"
using \<open>D \<in> \<N>\<close> in_rel_interior by blast
have "C \<inter> rel_interior D = (D \<inter> C) \<inter> rel_interior D"
using rel_interior_subset by blast
also have "(D \<inter> C) \<inter> rel_interior D = {}"
proof (rule face_of_disjoint_rel_interior)
show "D \<inter> C face_of D"
using \<N>_def \<open>C \<in> \<N>\<close> \<open>D \<in> \<N>\<close> intface\<M> by blast
have "D \<in> \<M> \<and> aff_dim D = int n"
using \<N>_def \<open>D \<in> \<N>\<close> by blast
moreover have "C \<in> \<M> \<and> aff_dim C = int n"
using \<N>_def \<open>C \<in> \<N>\<close> by blast
ultimately show "D \<inter> C \<noteq> D"
by (metis False face_of_aff_dim_lt inf.idem inf_le1 intface\<M> not_le poly\<M> polytope_imp_convex)
qed
finally have CD: "C \<inter> (rel_interior D) = {}" .
have zKC: "(convex hull insert ?z K) \<subseteq> C"
by (metis DiffE \<open>C \<in> \<N>\<close> \<open>K \<subseteq> rel_frontier C\<close> closed\<N> closure_closed convex\<N> hull_minimal insert_subset rel_frontier_def rel_interior_subset subset_iff z)
have eq: "convex hull (insert ?z K) \<inter> convex hull (insert ?w L) =
convex hull (insert ?z K) \<inter> convex hull L"
apply (rule Int_convex_hull_insert_rel_exterior [OF \<open>convex D\<close> \<open>L \<subseteq> D\<close> w])
using zKC CD apply (force simp: disjnt_def)
done
have ch_id: "convex hull K = K" "convex hull L = L"
using "*" \<open>K \<in> \<U>\<close> \<open>L \<in> \<U>\<close> hull_same by auto
have "convex C"
by (simp add: \<open>C \<in> \<N>\<close> convex\<N>)
have "convex hull (insert ?z K) \<inter> L = L \<inter> convex hull (insert ?z K)"
by blast
also have "... = convex hull K \<inter> L"
proof (subst Int_convex_hull_insert_rel_exterior [OF \<open>convex C\<close> \<open>K \<subseteq> C\<close> z])
have "(C \<inter> D) \<inter> rel_interior C = {}"
proof (rule face_of_disjoint_rel_interior)
show "C \<inter> D face_of C"
using \<N>_def \<open>C \<in> \<N>\<close> \<open>D \<in> \<N>\<close> intface\<M> by blast
have "D \<in> \<M>" "aff_dim D = int n"
using \<N>_def \<open>D \<in> \<N>\<close> by fastforce+
moreover have "C \<in> \<M>" "aff_dim C = int n"
using \<N>_def \<open>C \<in> \<N>\<close> by fastforce+
ultimately have "aff_dim D + - 1 * aff_dim C \<le> 0"
by fastforce
then have "\<not> C face_of D"
using False \<open>convex D\<close> face_of_aff_dim_lt by fastforce
show "C \<inter> D \<noteq> C"
using \<open>C \<in> \<M>\<close> \<open>D \<in> \<M>\<close> \<open>\<not> C face_of D\<close> intface\<M> by fastforce
qed
then have "D \<inter> rel_interior C = {}"
by (metis inf.absorb_iff2 inf_assoc inf_sup_aci(1) rel_interior_subset)
then show "disjnt L (rel_interior C)"
by (meson \<open>L \<subseteq> D\<close> disjnt_def disjnt_subset1)
next
show "L \<inter> convex hull K = convex hull K \<inter> L"
by force
qed
finally have chKL: "convex hull (insert ?z K) \<inter> L = convex hull K \<inter> L" .
have "convex hull insert ?z K \<inter> convex hull L face_of K"
by (simp add: \<open>K \<in> \<U>\<close> \<open>L \<in> \<U>\<close> ch_id chKL faceI\<U>)
also have "... face_of convex hull insert ?z K"
proof -
obtain I where I: "\<not> affine_dependent I" "K = convex hull I"
using * [OF \<open>K \<in> \<U>\<close>] by auto
then have "\<And>a. a \<notin> rel_interior C \<or> a \<notin> affine hull I"
using ahK_C_disjoint \<open>C \<in> \<N>\<close> \<open>K \<in> \<U>\<close> \<open>K \<subseteq> rel_frontier C\<close> affine_hull_convex_hull by blast
then show ?thesis
by (metis I affine_independent_insert face_of_convex_hull_affine_independent hull_insert subset_insertI z)
qed
finally have 1: "convex hull insert ?z K \<inter> convex hull L face_of convex hull insert ?z K" .
have "convex hull insert ?z K \<inter> convex hull L face_of L"
by (simp add: \<open>K \<in> \<U>\<close> \<open>L \<in> \<U>\<close> ch_id chKL faceI\<U>)
also have "... face_of convex hull insert ?w L"
proof -
obtain I where I: "\<not> affine_dependent I" "L = convex hull I"
using * [OF \<open>L \<in> \<U>\<close>] by auto
then have "\<And>a. a \<notin> rel_interior D \<or> a \<notin> affine hull I"
using \<open>D \<in> \<N>\<close> \<open>L \<in> \<U>\<close> \<open>L \<subseteq> rel_frontier D\<close> affine_hull_convex_hull ahK_C_disjoint by blast
then show ?thesis
by (metis I aff_independent_finite convex_convex_hull face_of_convex_hull_insert face_of_refl hull_insert w)
qed
finally have 2: "convex hull insert ?z K \<inter> convex hull L face_of convex hull insert ?w L" .
show ?thesis
by (simp add: X Y eq 1 2)
qed
qed
qed
qed
then
show "S \<inter> S' face_of S" "S \<inter> S' face_of S'" if "S \<in> \<U> \<union> ?\<T> \<and> S' \<in> \<U> \<union> ?\<T>" for S S'
using that by auto
qed
show "\<exists>F \<subseteq> \<U> \<union> ?\<T>. C = \<Union>F" if "C \<in> \<M>" for C
proof (cases "C \<in> \<S>")
case True
then show ?thesis
by (meson UnCI fin\<U> subsetD subsetI)
next
case False
then have "C \<in> \<N>"
by (simp add: \<N>_def \<S>_def aff\<M> less_le that)
let ?z = "SOME z. z \<in> rel_interior C"
have z: "?z \<in> rel_interior C"
using \<open>C \<in> \<N>\<close> in_rel_interior by blast
let ?F = "\<Union>K \<in> \<U> \<inter> Pow (rel_frontier C). {convex hull (insert ?z K)}"
have "?F \<subseteq> ?\<T>"
using \<open>C \<in> \<N>\<close> by blast
moreover have "C \<subseteq> \<Union>?F"
proof
fix x
assume "x \<in> C"
have "convex C"
using \<open>C \<in> \<N>\<close> convex\<N> by blast
have "bounded C"
using \<open>C \<in> \<N>\<close> by (simp add: poly\<M> polytope_imp_bounded that)
have "polytope C"
using \<open>C \<in> \<N>\<close> poly\<N> by auto
have "\<not> (?z = x \<and> C = {?z})"
using \<open>C \<in> \<N>\<close> aff_dim_sing [of ?z] \<open>\<not> n \<le> 1\<close> by (force simp: \<N>_def)
then obtain y where y: "y \<in> rel_frontier C" and xzy: "x \<in> closed_segment ?z y"
and sub: "open_segment ?z y \<subseteq> rel_interior C"
by (blast intro: segment_to_rel_frontier [OF \<open>convex C\<close> \<open>bounded C\<close> z \<open>x \<in> C\<close>])
then obtain F where "y \<in> F" "F face_of C" "F \<noteq> C"
by (auto simp: rel_frontier_of_polyhedron_alt [OF polytope_imp_polyhedron [OF \<open>polytope C\<close>]])
then obtain \<G> where "finite \<G>" "\<G> \<subseteq> \<U>" "F = \<Union>\<G>"
by (metis (mono_tags, lifting) \<S>_def \<open>C \<in> \<M>\<close> \<open>convex C\<close> aff\<M> face\<M> face_of_aff_dim_lt fin\<U> le_less_trans mem_Collect_eq not_less)
then obtain K where "y \<in> K" "K \<in> \<G>"
using \<open>y \<in> F\<close> by blast
moreover have x: "x \<in> convex hull {?z,y}"
using segment_convex_hull xzy by auto
moreover have "convex hull {?z,y} \<subseteq> convex hull insert ?z K"
by (metis (full_types) \<open>y \<in> K\<close> hull_mono empty_subsetI insertCI insert_subset)
moreover have "K \<in> \<U>"
using \<open>K \<in> \<G>\<close> \<open>\<G> \<subseteq> \<U>\<close> by blast
moreover have "K \<subseteq> rel_frontier C"
using \<open>F = \<Union>\<G>\<close> \<open>F \<noteq> C\<close> \<open>F face_of C\<close> \<open>K \<in> \<G>\<close> face_of_subset_rel_frontier by fastforce
ultimately show "x \<in> \<Union>?F"
by force
qed
moreover
have "convex hull insert (SOME z. z \<in> rel_interior C) K \<subseteq> C"
if "K \<in> \<U>" "K \<subseteq> rel_frontier C" for K
proof (rule hull_minimal)
show "insert (SOME z. z \<in> rel_interior C) K \<subseteq> C"
using that \<open>C \<in> \<N>\<close> in_rel_interior rel_interior_subset
by (force simp: closure_eq rel_frontier_def closed\<N>)
show "convex C"
by (simp add: \<open>C \<in> \<N>\<close> convex\<N>)
qed
then have "\<Union>?F \<subseteq> C"
by auto
ultimately show ?thesis
by blast
qed
have "(\<exists>C. C \<in> \<M> \<and> L \<subseteq> C) \<and> aff_dim L \<le> int n" if "L \<in> \<U> \<union> ?\<T>" for L
using that
proof
assume "L \<in> \<U>"
then show ?thesis
using C\<U> \<S>_def "*" by fastforce
next
assume "L \<in> ?\<T>"
then obtain C K where "C \<in> \<N>"
and L: "L = convex hull insert (SOME z. z \<in> rel_interior C) K"
and K: "K \<in> \<U>" "K \<subseteq> rel_frontier C"
by auto
then have "convex hull C = C"
by (meson convex\<N> convex_hull_eq)
then have "convex C"
by (metis (no_types) convex_convex_hull)
have "rel_frontier C \<subseteq> C"
by (metis DiffE closed\<N> \<open>C \<in> \<N>\<close> closure_closed rel_frontier_def subsetI)
have "K \<subseteq> C"
using K \<open>rel_frontier C \<subseteq> C\<close> by blast
have "C \<in> \<M>"
using \<N>_def \<open>C \<in> \<N>\<close> by auto
moreover have "L \<subseteq> C"
using K L \<open>C \<in> \<N>\<close>
by (metis \<open>K \<subseteq> C\<close> \<open>convex hull C = C\<close> contra_subsetD hull_mono in_rel_interior insert_subset rel_interior_subset)
ultimately show ?thesis
using \<open>rel_frontier C \<subseteq> C\<close> \<open>L \<subseteq> C\<close> aff\<M> aff_dim_subset \<open>C \<in> \<M>\<close> dual_order.trans by blast
qed
then show "\<exists>C. C \<in> \<M> \<and> L \<subseteq> C" "aff_dim L \<le> int n" if "L \<in> \<U> \<union> ?\<T>" for L
using that by auto
qed
then show ?thesis
apply (rule ex_forward, safe)
apply (meson Union_iff subsetCE, fastforce)
by (meson infinite_super simplicial_complex_def)
qed
qed
lemma simplicial_subdivision_of_cell_complex_lowdim:
assumes "finite \<M>"
and poly: "\<And>C. C \<in> \<M> \<Longrightarrow> polytope C"
and face: "\<And>C1 C2. \<lbrakk>C1 \<in> \<M>; C2 \<in> \<M>\<rbrakk> \<Longrightarrow> C1 \<inter> C2 face_of C1 \<and> C1 \<inter> C2 face_of C2"
and aff: "\<And>C. C \<in> \<M> \<Longrightarrow> aff_dim C \<le> d"
obtains \<T> where "simplicial_complex \<T>" "\<And>K. K \<in> \<T> \<Longrightarrow> aff_dim K \<le> d"
"\<Union>\<T> = \<Union>\<M>"
"\<And>C. C \<in> \<M> \<Longrightarrow> \<exists>F. finite F \<and> F \<subseteq> \<T> \<and> C = \<Union>F"
"\<And>K. K \<in> \<T> \<Longrightarrow> \<exists>C. C \<in> \<M> \<and> K \<subseteq> C"
proof (cases "d \<ge> 0")
case True
then obtain n where n: "d = of_nat n"
using zero_le_imp_eq_int by blast
have "\<exists>\<T>. simplicial_complex \<T> \<and>
(\<forall>K\<in>\<T>. aff_dim K \<le> int n) \<and>
\<Union>\<T> = \<Union>(\<Union>C\<in>\<M>. {F. F face_of C}) \<and>
(\<forall>C\<in>\<Union>C\<in>\<M>. {F. F face_of C}.
\<exists>F. finite F \<and> F \<subseteq> \<T> \<and> C = \<Union>F) \<and>
(\<forall>K\<in>\<T>. \<exists>C. C \<in> (\<Union>C\<in>\<M>. {F. F face_of C}) \<and> K \<subseteq> C)"
proof (rule simplicial_subdivision_aux)
show "finite (\<Union>C\<in>\<M>. {F. F face_of C})"
using \<open>finite \<M>\<close> poly polyhedron_eq_finite_faces polytope_imp_polyhedron by fastforce
show "polytope F" if "F \<in> (\<Union>C\<in>\<M>. {F. F face_of C})" for F
using poly that face_of_polytope_polytope by blast
show "aff_dim F \<le> int n" if "F \<in> (\<Union>C\<in>\<M>. {F. F face_of C})" for F
using that
by clarify (metis n aff_dim_subset aff face_of_imp_subset order_trans)
show "F \<in> (\<Union>C\<in>\<M>. {F. F face_of C})"
if "G \<in> (\<Union>C\<in>\<M>. {F. F face_of C})" and "F face_of G" for F G
using that face_of_trans by blast
next
show "F1 \<inter> F2 face_of F1 \<and> F1 \<inter> F2 face_of F2"
if "F1 \<in> (\<Union>C\<in>\<M>. {F. F face_of C})" and "F2 \<in> (\<Union>C\<in>\<M>. {F. F face_of C})" for F1 F2
using that
by safe (meson face face_of_Int_subface)+
qed
moreover
have "\<Union>(\<Union>C\<in>\<M>. {F. F face_of C}) = \<Union>\<M>"
using face_of_imp_subset face by blast
ultimately show ?thesis
apply clarify
apply (rule that, assumption+)
using n apply blast
apply (simp_all add: poly face_of_refl polytope_imp_convex)
using face_of_imp_subset by fastforce
next
case False
then have m1: "\<And>C. C \<in> \<M> \<Longrightarrow> aff_dim C = -1"
by (metis aff aff_dim_empty_eq aff_dim_negative_iff dual_order.trans not_less)
then have face\<M>: "\<And>F S. \<lbrakk>S \<in> \<M>; F face_of S\<rbrakk> \<Longrightarrow> F \<in> \<M>"
by (metis aff_dim_empty face_of_empty)
show ?thesis
proof
have "\<And>S. S \<in> \<M> \<Longrightarrow> \<exists>n. n simplex S"
by (metis (no_types) m1 aff_dim_empty simplex_minus_1)
then show "simplicial_complex \<M>"
by (auto simp: simplicial_complex_def \<open>finite \<M>\<close> face intro: face\<M>)
show "aff_dim K \<le> d" if "K \<in> \<M>" for K
by (simp add: that aff)
show "\<exists>F. finite F \<and> F \<subseteq> \<M> \<and> C = \<Union>F" if "C \<in> \<M>" for C
using \<open>C \<in> \<M>\<close> equals0I by auto
show "\<exists>C. C \<in> \<M> \<and> K \<subseteq> C" if "K \<in> \<M>" for K
using \<open>K \<in> \<M>\<close> by blast
qed auto
qed
proposition simplicial_subdivision_of_cell_complex:
assumes "finite \<M>"
and poly: "\<And>C. C \<in> \<M> \<Longrightarrow> polytope C"
and face: "\<And>C1 C2. \<lbrakk>C1 \<in> \<M>; C2 \<in> \<M>\<rbrakk> \<Longrightarrow> C1 \<inter> C2 face_of C1 \<and> C1 \<inter> C2 face_of C2"
obtains \<T> where "simplicial_complex \<T>"
"\<Union>\<T> = \<Union>\<M>"
"\<And>C. C \<in> \<M> \<Longrightarrow> \<exists>F. finite F \<and> F \<subseteq> \<T> \<and> C = \<Union>F"
"\<And>K. K \<in> \<T> \<Longrightarrow> \<exists>C. C \<in> \<M> \<and> K \<subseteq> C"
by (blast intro: simplicial_subdivision_of_cell_complex_lowdim [OF assms aff_dim_le_DIM])
corollary fine_simplicial_subdivision_of_cell_complex:
assumes "0 < e" "finite \<M>"
and poly: "\<And>C. C \<in> \<M> \<Longrightarrow> polytope C"
and face: "\<And>C1 C2. \<lbrakk>C1 \<in> \<M>; C2 \<in> \<M>\<rbrakk> \<Longrightarrow> C1 \<inter> C2 face_of C1 \<and> C1 \<inter> C2 face_of C2"
obtains \<T> where "simplicial_complex \<T>"
"\<And>K. K \<in> \<T> \<Longrightarrow> diameter K < e"
"\<Union>\<T> = \<Union>\<M>"
"\<And>C. C \<in> \<M> \<Longrightarrow> \<exists>F. finite F \<and> F \<subseteq> \<T> \<and> C = \<Union>F"
"\<And>K. K \<in> \<T> \<Longrightarrow> \<exists>C. C \<in> \<M> \<and> K \<subseteq> C"
proof -
obtain \<N> where \<N>: "finite \<N>" "\<Union>\<N> = \<Union>\<M>"
and diapoly: "\<And>X. X \<in> \<N> \<Longrightarrow> diameter X < e" "\<And>X. X \<in> \<N> \<Longrightarrow> polytope X"
and "\<And>X Y. \<lbrakk>X \<in> \<N>; Y \<in> \<N>\<rbrakk> \<Longrightarrow> X \<inter> Y face_of X \<and> X \<inter> Y face_of Y"
and \<N>covers: "\<And>C x. C \<in> \<M> \<and> x \<in> C \<Longrightarrow> \<exists>D. D \<in> \<N> \<and> x \<in> D \<and> D \<subseteq> C"
and \<N>covered: "\<And>C. C \<in> \<N> \<Longrightarrow> \<exists>D. D \<in> \<M> \<and> C \<subseteq> D"
by (blast intro: cell_complex_subdivision_exists [OF \<open>0 < e\<close> \<open>finite \<M>\<close> poly aff_dim_le_DIM face])
then obtain \<T> where \<T>: "simplicial_complex \<T>" "\<Union>\<T> = \<Union>\<N>"
and \<T>covers: "\<And>C. C \<in> \<N> \<Longrightarrow> \<exists>F. finite F \<and> F \<subseteq> \<T> \<and> C = \<Union>F"
and \<T>covered: "\<And>K. K \<in> \<T> \<Longrightarrow> \<exists>C. C \<in> \<N> \<and> K \<subseteq> C"
using simplicial_subdivision_of_cell_complex [OF \<open>finite \<N>\<close>] by metis
show ?thesis
proof
show "simplicial_complex \<T>"
by (rule \<T>)
show "diameter K < e" if "K \<in> \<T>" for K
by (metis le_less_trans diapoly \<T>covered diameter_subset polytope_imp_bounded that)
show "\<Union>\<T> = \<Union>\<M>"
by (simp add: \<N>(2) \<open>\<Union>\<T> = \<Union>\<N>\<close>)
show "\<exists>F. finite F \<and> F \<subseteq> \<T> \<and> C = \<Union>F" if "C \<in> \<M>" for C
proof -
{ fix x
assume "x \<in> C"
then obtain D where "D \<in> \<T>" "x \<in> D" "D \<subseteq> C"
using \<N>covers \<open>C \<in> \<M>\<close> \<T>covers by force
then have "\<exists>X\<in>\<T> \<inter> Pow C. x \<in> X"
using \<open>D \<in> \<T>\<close> \<open>D \<subseteq> C\<close> \<open>x \<in> D\<close> by blast
}
moreover
have "finite (\<T> \<inter> Pow C)"
using \<open>simplicial_complex \<T>\<close> simplicial_complex_def by auto
ultimately show ?thesis
by (rule_tac x="(\<T> \<inter> Pow C)" in exI) auto
qed
show "\<exists>C. C \<in> \<M> \<and> K \<subseteq> C" if "K \<in> \<T>" for K
by (meson \<N>covered \<T>covered order_trans that)
qed
qed
subsection\<open>Some results on cell division with full-dimensional cells only\<close>
lemma convex_Union_fulldim_cells:
assumes "finite \<S>" and clo: "\<And>C. C \<in> \<S> \<Longrightarrow> closed C" and con: "\<And>C. C \<in> \<S> \<Longrightarrow> convex C"
and eq: "\<Union>\<S> = U"and "convex U"
shows "\<Union>{C \<in> \<S>. aff_dim C = aff_dim U} = U" (is "?lhs = U")
proof -
have "closed U"
using \<open>finite \<S>\<close> clo eq by blast
have "?lhs \<subseteq> U"
using eq by blast
moreover have "U \<subseteq> ?lhs"
proof (cases "\<forall>C \<in> \<S>. aff_dim C = aff_dim U")
case True
then show ?thesis
using eq by blast
next
case False
have "closed ?lhs"
by (simp add: \<open>finite \<S>\<close> clo closed_Union)
moreover have "U \<subseteq> closure ?lhs"
proof -
have "U \<subseteq> closure(\<Inter>{U - C |C. C \<in> \<S> \<and> aff_dim C < aff_dim U})"
proof (rule Baire [OF \<open>closed U\<close>])
show "countable {U - C |C. C \<in> \<S> \<and> aff_dim C < aff_dim U}"
using \<open>finite \<S>\<close> uncountable_infinite by fastforce
have "\<And>C. C \<in> \<S> \<Longrightarrow> openin (top_of_set U) (U-C)"
by (metis Sup_upper clo closed_limpt closedin_limpt eq openin_diff openin_subtopology_self)
then show "openin (top_of_set U) T \<and> U \<subseteq> closure T"
if "T \<in> {U - C |C. C \<in> \<S> \<and> aff_dim C < aff_dim U}" for T
using that dense_complement_convex_closed \<open>closed U\<close> \<open>convex U\<close> by auto
qed
also have "... \<subseteq> closure ?lhs"
proof -
obtain C where "C \<in> \<S>" "aff_dim C < aff_dim U"
by (metis False Sup_upper aff_dim_subset eq eq_iff not_le)
have "\<exists>X. X \<in> \<S> \<and> aff_dim X = aff_dim U \<and> x \<in> X"
if "\<And>V. (\<exists>C. V = U - C \<and> C \<in> \<S> \<and> aff_dim C < aff_dim U) \<Longrightarrow> x \<in> V" for x
proof -
have "x \<in> U \<and> x \<in> \<Union>\<S>"
using \<open>C \<in> \<S>\<close> \<open>aff_dim C < aff_dim U\<close> eq that by blast
then show ?thesis
by (metis Diff_iff Sup_upper Union_iff aff_dim_subset dual_order.order_iff_strict eq that)
qed
then show ?thesis
by (auto intro!: closure_mono)
qed
finally show ?thesis .
qed
ultimately show ?thesis
using closure_subset_eq by blast
qed
ultimately show ?thesis by blast
qed
proposition fine_triangular_subdivision_of_cell_complex:
assumes "0 < e" "finite \<M>"
and poly: "\<And>C. C \<in> \<M> \<Longrightarrow> polytope C"
and aff: "\<And>C. C \<in> \<M> \<Longrightarrow> aff_dim C = d"
and face: "\<And>C1 C2. \<lbrakk>C1 \<in> \<M>; C2 \<in> \<M>\<rbrakk> \<Longrightarrow> C1 \<inter> C2 face_of C1 \<and> C1 \<inter> C2 face_of C2"
obtains \<T> where "triangulation \<T>" "\<And>k. k \<in> \<T> \<Longrightarrow> diameter k < e"
"\<And>k. k \<in> \<T> \<Longrightarrow> aff_dim k = d" "\<Union>\<T> = \<Union>\<M>"
"\<And>C. C \<in> \<M> \<Longrightarrow> \<exists>f. finite f \<and> f \<subseteq> \<T> \<and> C = \<Union>f"
"\<And>k. k \<in> \<T> \<Longrightarrow> \<exists>C. C \<in> \<M> \<and> k \<subseteq> C"
proof -
obtain \<T> where "simplicial_complex \<T>"
and dia\<T>: "\<And>K. K \<in> \<T> \<Longrightarrow> diameter K < e"
and "\<Union>\<T> = \<Union>\<M>"
and in\<M>: "\<And>C. C \<in> \<M> \<Longrightarrow> \<exists>F. finite F \<and> F \<subseteq> \<T> \<and> C = \<Union>F"
and in\<T>: "\<And>K. K \<in> \<T> \<Longrightarrow> \<exists>C. C \<in> \<M> \<and> K \<subseteq> C"
by (blast intro: fine_simplicial_subdivision_of_cell_complex [OF \<open>e > 0\<close> \<open>finite \<M>\<close> poly face])
let ?\<T> = "{K \<in> \<T>. aff_dim K = d}"
show thesis
proof
show "triangulation ?\<T>"
using \<open>simplicial_complex \<T>\<close> by (auto simp: triangulation_def simplicial_complex_def)
show "diameter L < e" if "L \<in> {K \<in> \<T>. aff_dim K = d}" for L
using that by (auto simp: dia\<T>)
show "aff_dim L = d" if "L \<in> {K \<in> \<T>. aff_dim K = d}" for L
using that by auto
show "\<exists>F. finite F \<and> F \<subseteq> {K \<in> \<T>. aff_dim K = d} \<and> C = \<Union>F" if "C \<in> \<M>" for C
proof -
obtain F where "finite F" "F \<subseteq> \<T>" "C = \<Union>F"
using in\<M> [OF \<open>C \<in> \<M>\<close>] by auto
show ?thesis
proof (intro exI conjI)
show "finite {K \<in> F. aff_dim K = d}"
by (simp add: \<open>finite F\<close>)
show "{K \<in> F. aff_dim K = d} \<subseteq> {K \<in> \<T>. aff_dim K = d}"
using \<open>F \<subseteq> \<T>\<close> by blast
have "d = aff_dim C"
by (simp add: aff that)
moreover have "\<And>K. K \<in> F \<Longrightarrow> closed K \<and> convex K"
using \<open>simplicial_complex \<T>\<close> \<open>F \<subseteq> \<T>\<close>
unfolding simplicial_complex_def by (metis subsetCE \<open>F \<subseteq> \<T>\<close> closed_simplex convex_simplex)
moreover have "convex (\<Union>F)"
using \<open>C = \<Union>F\<close> poly polytope_imp_convex that by blast
ultimately show "C = \<Union>{K \<in> F. aff_dim K = d}"
by (simp add: convex_Union_fulldim_cells \<open>C = \<Union>F\<close> \<open>finite F\<close>)
qed
qed
then show "\<Union>{K \<in> \<T>. aff_dim K = d} = \<Union>\<M>"
by auto (meson in\<T> subsetCE)
show "\<exists>C. C \<in> \<M> \<and> L \<subseteq> C"
if "L \<in> {K \<in> \<T>. aff_dim K = d}" for L
using that by (auto simp: in\<T>)
qed
qed
end
diff --git a/src/HOL/Analysis/Product_Topology.thy b/src/HOL/Analysis/Product_Topology.thy
--- a/src/HOL/Analysis/Product_Topology.thy
+++ b/src/HOL/Analysis/Product_Topology.thy
@@ -1,577 +1,577 @@
section\<open>The binary product topology\<close>
theory Product_Topology
imports Function_Topology
begin
section \<open>Product Topology\<close>
subsection\<open>Definition\<close>
definition prod_topology :: "'a topology \<Rightarrow> 'b topology \<Rightarrow> ('a \<times> 'b) topology" where
"prod_topology X Y \<equiv> topology (arbitrary union_of (\<lambda>U. U \<in> {S \<times> T |S T. openin X S \<and> openin Y T}))"
lemma open_product_open:
assumes "open A"
shows "\<exists>\<U>. \<U> \<subseteq> {S \<times> T |S T. open S \<and> open T} \<and> \<Union> \<U> = A"
proof -
obtain f g where *: "\<And>u. u \<in> A \<Longrightarrow> open (f u) \<and> open (g u) \<and> u \<in> (f u) \<times> (g u) \<and> (f u) \<times> (g u) \<subseteq> A"
using open_prod_def [of A] assms by metis
let ?\<U> = "(\<lambda>u. f u \<times> g u) ` A"
show ?thesis
by (rule_tac x="?\<U>" in exI) (auto simp: dest: *)
qed
lemma open_product_open_eq: "(arbitrary union_of (\<lambda>U. \<exists>S T. U = S \<times> T \<and> open S \<and> open T)) = open"
by (force simp: union_of_def arbitrary_def intro: open_product_open open_Times)
lemma openin_prod_topology:
"openin (prod_topology X Y) = arbitrary union_of (\<lambda>U. U \<in> {S \<times> T |S T. openin X S \<and> openin Y T})"
unfolding prod_topology_def
proof (rule topology_inverse')
show "istopology (arbitrary union_of (\<lambda>U. U \<in> {S \<times> T |S T. openin X S \<and> openin Y T}))"
apply (rule istopology_base, simp)
by (metis openin_Int Times_Int_Times)
qed
lemma topspace_prod_topology [simp]:
"topspace (prod_topology X Y) = topspace X \<times> topspace Y"
proof -
have "topspace(prod_topology X Y) = \<Union> (Collect (openin (prod_topology X Y)))" (is "_ = ?Z")
unfolding topspace_def ..
also have "\<dots> = topspace X \<times> topspace Y"
proof
show "?Z \<subseteq> topspace X \<times> topspace Y"
apply (auto simp: openin_prod_topology union_of_def arbitrary_def)
using openin_subset by force+
next
have *: "\<exists>A B. topspace X \<times> topspace Y = A \<times> B \<and> openin X A \<and> openin Y B"
by blast
show "topspace X \<times> topspace Y \<subseteq> ?Z"
apply (rule Union_upper)
using * by (simp add: openin_prod_topology arbitrary_union_of_inc)
qed
finally show ?thesis .
qed
lemma subtopology_Times:
shows "subtopology (prod_topology X Y) (S \<times> T) = prod_topology (subtopology X S) (subtopology Y T)"
proof -
have "((\<lambda>U. \<exists>S T. U = S \<times> T \<and> openin X S \<and> openin Y T) relative_to S \<times> T) =
(\<lambda>U. \<exists>S' T'. U = S' \<times> T' \<and> (openin X relative_to S) S' \<and> (openin Y relative_to T) T')"
by (auto simp: relative_to_def Times_Int_Times fun_eq_iff) metis
then show ?thesis
by (simp add: topology_eq openin_prod_topology arbitrary_union_of_relative_to flip: openin_relative_to)
qed
lemma prod_topology_subtopology:
"prod_topology (subtopology X S) Y = subtopology (prod_topology X Y) (S \<times> topspace Y)"
"prod_topology X (subtopology Y T) = subtopology (prod_topology X Y) (topspace X \<times> T)"
by (auto simp: subtopology_Times)
lemma prod_topology_discrete_topology:
"discrete_topology (S \<times> T) = prod_topology (discrete_topology S) (discrete_topology T)"
by (auto simp: discrete_topology_unique openin_prod_topology intro: arbitrary_union_of_inc)
lemma prod_topology_euclidean [simp]: "prod_topology euclidean euclidean = euclidean"
by (simp add: prod_topology_def open_product_open_eq)
lemma prod_topology_subtopology_eu [simp]:
"prod_topology (subtopology euclidean S) (subtopology euclidean T) = subtopology euclidean (S \<times> T)"
by (simp add: prod_topology_subtopology subtopology_subtopology Times_Int_Times)
lemma openin_prod_topology_alt:
"openin (prod_topology X Y) S \<longleftrightarrow>
(\<forall>x y. (x,y) \<in> S \<longrightarrow> (\<exists>U V. openin X U \<and> openin Y V \<and> x \<in> U \<and> y \<in> V \<and> U \<times> V \<subseteq> S))"
apply (auto simp: openin_prod_topology arbitrary_union_of_alt, fastforce)
by (metis mem_Sigma_iff)
lemma open_map_fst: "open_map (prod_topology X Y) X fst"
unfolding open_map_def openin_prod_topology_alt
by (force simp: openin_subopen [of X "fst ` _"] intro: subset_fst_imageI)
lemma open_map_snd: "open_map (prod_topology X Y) Y snd"
unfolding open_map_def openin_prod_topology_alt
by (force simp: openin_subopen [of Y "snd ` _"] intro: subset_snd_imageI)
lemma openin_prod_Times_iff:
"openin (prod_topology X Y) (S \<times> T) \<longleftrightarrow> S = {} \<or> T = {} \<or> openin X S \<and> openin Y T"
proof (cases "S = {} \<or> T = {}")
case False
then show ?thesis
apply (simp add: openin_prod_topology_alt openin_subopen [of X S] openin_subopen [of Y T] times_subset_iff, safe)
apply (meson|force)+
done
qed force
lemma closure_of_Times:
"(prod_topology X Y) closure_of (S \<times> T) = (X closure_of S) \<times> (Y closure_of T)" (is "?lhs = ?rhs")
proof
show "?lhs \<subseteq> ?rhs"
by (clarsimp simp: closure_of_def openin_prod_topology_alt) blast
show "?rhs \<subseteq> ?lhs"
by (clarsimp simp: closure_of_def openin_prod_topology_alt) (meson SigmaI subsetD)
qed
lemma closedin_prod_Times_iff:
"closedin (prod_topology X Y) (S \<times> T) \<longleftrightarrow> S = {} \<or> T = {} \<or> closedin X S \<and> closedin Y T"
by (auto simp: closure_of_Times times_eq_iff simp flip: closure_of_eq)
lemma interior_of_Times: "(prod_topology X Y) interior_of (S \<times> T) = (X interior_of S) \<times> (Y interior_of T)"
proof (rule interior_of_unique)
show "(X interior_of S) \<times> Y interior_of T \<subseteq> S \<times> T"
by (simp add: Sigma_mono interior_of_subset)
show "openin (prod_topology X Y) ((X interior_of S) \<times> Y interior_of T)"
by (simp add: openin_prod_Times_iff)
next
show "T' \<subseteq> (X interior_of S) \<times> Y interior_of T" if "T' \<subseteq> S \<times> T" "openin (prod_topology X Y) T'" for T'
proof (clarsimp; intro conjI)
fix a :: "'a" and b :: "'b"
assume "(a, b) \<in> T'"
with that obtain U V where UV: "openin X U" "openin Y V" "a \<in> U" "b \<in> V" "U \<times> V \<subseteq> T'"
by (metis openin_prod_topology_alt)
then show "a \<in> X interior_of S"
using interior_of_maximal_eq that(1) by fastforce
show "b \<in> Y interior_of T"
using UV interior_of_maximal_eq that(1)
by (metis SigmaI mem_Sigma_iff subset_eq)
qed
qed
subsection \<open>Continuity\<close>
lemma continuous_map_pairwise:
"continuous_map Z (prod_topology X Y) f \<longleftrightarrow> continuous_map Z X (fst \<circ> f) \<and> continuous_map Z Y (snd \<circ> f)"
(is "?lhs = ?rhs")
proof -
let ?g = "fst \<circ> f" and ?h = "snd \<circ> f"
have f: "f x = (?g x, ?h x)" for x
by auto
show ?thesis
proof (cases "(\<forall>x \<in> topspace Z. ?g x \<in> topspace X) \<and> (\<forall>x \<in> topspace Z. ?h x \<in> topspace Y)")
case True
show ?thesis
proof safe
assume "continuous_map Z (prod_topology X Y) f"
then have "openin Z {x \<in> topspace Z. fst (f x) \<in> U}" if "openin X U" for U
unfolding continuous_map_def using True that
apply clarify
apply (drule_tac x="U \<times> topspace Y" in spec)
by (simp add: openin_prod_Times_iff mem_Times_iff cong: conj_cong)
with True show "continuous_map Z X (fst \<circ> f)"
by (auto simp: continuous_map_def)
next
assume "continuous_map Z (prod_topology X Y) f"
then have "openin Z {x \<in> topspace Z. snd (f x) \<in> V}" if "openin Y V" for V
unfolding continuous_map_def using True that
apply clarify
apply (drule_tac x="topspace X \<times> V" in spec)
by (simp add: openin_prod_Times_iff mem_Times_iff cong: conj_cong)
with True show "continuous_map Z Y (snd \<circ> f)"
by (auto simp: continuous_map_def)
next
assume Z: "continuous_map Z X (fst \<circ> f)" "continuous_map Z Y (snd \<circ> f)"
have *: "openin Z {x \<in> topspace Z. f x \<in> W}"
if "\<And>w. w \<in> W \<Longrightarrow> \<exists>U V. openin X U \<and> openin Y V \<and> w \<in> U \<times> V \<and> U \<times> V \<subseteq> W" for W
proof (subst openin_subopen, clarify)
fix x :: "'a"
assume "x \<in> topspace Z" and "f x \<in> W"
with that [OF \<open>f x \<in> W\<close>]
obtain U V where UV: "openin X U" "openin Y V" "f x \<in> U \<times> V" "U \<times> V \<subseteq> W"
by auto
with Z UV show "\<exists>T. openin Z T \<and> x \<in> T \<and> T \<subseteq> {x \<in> topspace Z. f x \<in> W}"
apply (rule_tac x="{x \<in> topspace Z. ?g x \<in> U} \<inter> {x \<in> topspace Z. ?h x \<in> V}" in exI)
apply (auto simp: \<open>x \<in> topspace Z\<close> continuous_map_def)
done
qed
show "continuous_map Z (prod_topology X Y) f"
using True by (simp add: continuous_map_def openin_prod_topology_alt mem_Times_iff *)
qed
qed (auto simp: continuous_map_def)
qed
lemma continuous_map_paired:
"continuous_map Z (prod_topology X Y) (\<lambda>x. (f x,g x)) \<longleftrightarrow> continuous_map Z X f \<and> continuous_map Z Y g"
by (simp add: continuous_map_pairwise o_def)
lemma continuous_map_pairedI [continuous_intros]:
"\<lbrakk>continuous_map Z X f; continuous_map Z Y g\<rbrakk> \<Longrightarrow> continuous_map Z (prod_topology X Y) (\<lambda>x. (f x,g x))"
by (simp add: continuous_map_pairwise o_def)
lemma continuous_map_fst [continuous_intros]: "continuous_map (prod_topology X Y) X fst"
using continuous_map_pairwise [of "prod_topology X Y" X Y id]
by (simp add: continuous_map_pairwise)
lemma continuous_map_snd [continuous_intros]: "continuous_map (prod_topology X Y) Y snd"
using continuous_map_pairwise [of "prod_topology X Y" X Y id]
by (simp add: continuous_map_pairwise)
lemma continuous_map_fst_of [continuous_intros]:
"continuous_map Z (prod_topology X Y) f \<Longrightarrow> continuous_map Z X (fst \<circ> f)"
by (simp add: continuous_map_pairwise)
lemma continuous_map_snd_of [continuous_intros]:
"continuous_map Z (prod_topology X Y) f \<Longrightarrow> continuous_map Z Y (snd \<circ> f)"
by (simp add: continuous_map_pairwise)
lemma continuous_map_prod_fst:
"i \<in> I \<Longrightarrow> continuous_map (prod_topology (product_topology (\<lambda>i. Y) I) X) Y (\<lambda>x. fst x i)"
using continuous_map_componentwise_UNIV continuous_map_fst by fastforce
lemma continuous_map_prod_snd:
"i \<in> I \<Longrightarrow> continuous_map (prod_topology X (product_topology (\<lambda>i. Y) I)) Y (\<lambda>x. snd x i)"
using continuous_map_componentwise_UNIV continuous_map_snd by fastforce
lemma continuous_map_if_iff [simp]: "continuous_map X Y (\<lambda>x. if P then f x else g x) \<longleftrightarrow> continuous_map X Y (if P then f else g)"
by simp
lemma continuous_map_if [continuous_intros]: "\<lbrakk>P \<Longrightarrow> continuous_map X Y f; ~P \<Longrightarrow> continuous_map X Y g\<rbrakk>
\<Longrightarrow> continuous_map X Y (\<lambda>x. if P then f x else g x)"
by simp
lemma continuous_map_subtopology_fst [continuous_intros]: "continuous_map (subtopology (prod_topology X Y) Z) X fst"
using continuous_map_from_subtopology continuous_map_fst by force
lemma continuous_map_subtopology_snd [continuous_intros]: "continuous_map (subtopology (prod_topology X Y) Z) Y snd"
using continuous_map_from_subtopology continuous_map_snd by force
lemma quotient_map_fst [simp]:
"quotient_map(prod_topology X Y) X fst \<longleftrightarrow> (topspace Y = {} \<longrightarrow> topspace X = {})"
by (auto simp: continuous_open_quotient_map open_map_fst continuous_map_fst)
lemma quotient_map_snd [simp]:
"quotient_map(prod_topology X Y) Y snd \<longleftrightarrow> (topspace X = {} \<longrightarrow> topspace Y = {})"
by (auto simp: continuous_open_quotient_map open_map_snd continuous_map_snd)
lemma retraction_map_fst:
"retraction_map (prod_topology X Y) X fst \<longleftrightarrow> (topspace Y = {} \<longrightarrow> topspace X = {})"
proof (cases "topspace Y = {}")
case True
then show ?thesis
using continuous_map_image_subset_topspace
by (fastforce simp: retraction_map_def retraction_maps_def continuous_map_fst continuous_map_on_empty)
next
case False
have "\<exists>g. continuous_map X (prod_topology X Y) g \<and> (\<forall>x\<in>topspace X. fst (g x) = x)"
if y: "y \<in> topspace Y" for y
by (rule_tac x="\<lambda>x. (x,y)" in exI) (auto simp: y continuous_map_paired)
with False have "retraction_map (prod_topology X Y) X fst"
by (fastforce simp: retraction_map_def retraction_maps_def continuous_map_fst)
with False show ?thesis
by simp
qed
lemma retraction_map_snd:
"retraction_map (prod_topology X Y) Y snd \<longleftrightarrow> (topspace X = {} \<longrightarrow> topspace Y = {})"
proof (cases "topspace X = {}")
case True
then show ?thesis
using continuous_map_image_subset_topspace
by (fastforce simp: retraction_map_def retraction_maps_def continuous_map_fst continuous_map_on_empty)
next
case False
have "\<exists>g. continuous_map Y (prod_topology X Y) g \<and> (\<forall>y\<in>topspace Y. snd (g y) = y)"
if x: "x \<in> topspace X" for x
by (rule_tac x="\<lambda>y. (x,y)" in exI) (auto simp: x continuous_map_paired)
with False have "retraction_map (prod_topology X Y) Y snd"
by (fastforce simp: retraction_map_def retraction_maps_def continuous_map_snd)
with False show ?thesis
by simp
qed
lemma continuous_map_of_fst:
"continuous_map (prod_topology X Y) Z (f \<circ> fst) \<longleftrightarrow> topspace Y = {} \<or> continuous_map X Z f"
proof (cases "topspace Y = {}")
case True
then show ?thesis
by (simp add: continuous_map_on_empty)
next
case False
then show ?thesis
- by (simp add: continuous_compose_quotient_map_eq quotient_map_fst)
+ by (simp add: continuous_compose_quotient_map_eq)
qed
lemma continuous_map_of_snd:
"continuous_map (prod_topology X Y) Z (f \<circ> snd) \<longleftrightarrow> topspace X = {} \<or> continuous_map Y Z f"
proof (cases "topspace X = {}")
case True
then show ?thesis
by (simp add: continuous_map_on_empty)
next
case False
then show ?thesis
- by (simp add: continuous_compose_quotient_map_eq quotient_map_snd)
+ by (simp add: continuous_compose_quotient_map_eq)
qed
lemma continuous_map_prod_top:
"continuous_map (prod_topology X Y) (prod_topology X' Y') (\<lambda>(x,y). (f x, g y)) \<longleftrightarrow>
topspace (prod_topology X Y) = {} \<or> continuous_map X X' f \<and> continuous_map Y Y' g"
proof (cases "topspace (prod_topology X Y) = {}")
case True
then show ?thesis
by (simp add: continuous_map_on_empty)
next
case False
then show ?thesis
by (simp add: continuous_map_paired case_prod_unfold continuous_map_of_fst [unfolded o_def] continuous_map_of_snd [unfolded o_def])
qed
lemma in_prod_topology_closure_of:
assumes "z \<in> (prod_topology X Y) closure_of S"
shows "fst z \<in> X closure_of (fst ` S)" "snd z \<in> Y closure_of (snd ` S)"
using assms continuous_map_eq_image_closure_subset continuous_map_fst apply fastforce
using assms continuous_map_eq_image_closure_subset continuous_map_snd apply fastforce
done
proposition compact_space_prod_topology:
"compact_space(prod_topology X Y) \<longleftrightarrow> topspace(prod_topology X Y) = {} \<or> compact_space X \<and> compact_space Y"
proof (cases "topspace(prod_topology X Y) = {}")
case True
then show ?thesis
using compact_space_topspace_empty by blast
next
case False
then have non_mt: "topspace X \<noteq> {}" "topspace Y \<noteq> {}"
by auto
have "compact_space X" "compact_space Y" if "compact_space(prod_topology X Y)"
proof -
have "compactin X (fst ` (topspace X \<times> topspace Y))"
by (metis compact_space_def continuous_map_fst image_compactin that topspace_prod_topology)
moreover
have "compactin Y (snd ` (topspace X \<times> topspace Y))"
by (metis compact_space_def continuous_map_snd image_compactin that topspace_prod_topology)
ultimately show "compact_space X" "compact_space Y"
by (simp_all add: non_mt compact_space_def)
qed
moreover
define \<X> where "\<X> \<equiv> (\<lambda>V. topspace X \<times> V) ` Collect (openin Y)"
define \<Y> where "\<Y> \<equiv> (\<lambda>U. U \<times> topspace Y) ` Collect (openin X)"
have "compact_space(prod_topology X Y)" if "compact_space X" "compact_space Y"
proof (rule Alexander_subbase_alt)
show toptop: "topspace X \<times> topspace Y \<subseteq> \<Union>(\<X> \<union> \<Y>)"
unfolding \<X>_def \<Y>_def by auto
fix \<C> :: "('a \<times> 'b) set set"
assume \<C>: "\<C> \<subseteq> \<X> \<union> \<Y>" "topspace X \<times> topspace Y \<subseteq> \<Union>\<C>"
then obtain \<X>' \<Y>' where XY: "\<X>' \<subseteq> \<X>" "\<Y>' \<subseteq> \<Y>" and \<C>eq: "\<C> = \<X>' \<union> \<Y>'"
using subset_UnE by metis
then have sub: "topspace X \<times> topspace Y \<subseteq> \<Union>(\<X>' \<union> \<Y>')"
using \<C> by simp
obtain \<U> \<V> where \<U>: "\<And>U. U \<in> \<U> \<Longrightarrow> openin X U" "\<Y>' = (\<lambda>U. U \<times> topspace Y) ` \<U>"
and \<V>: "\<And>V. V \<in> \<V> \<Longrightarrow> openin Y V" "\<X>' = (\<lambda>V. topspace X \<times> V) ` \<V>"
using XY by (clarsimp simp add: \<X>_def \<Y>_def subset_image_iff) (force simp add: subset_iff)
have "\<exists>\<D>. finite \<D> \<and> \<D> \<subseteq> \<X>' \<union> \<Y>' \<and> topspace X \<times> topspace Y \<subseteq> \<Union> \<D>"
proof -
have "topspace X \<subseteq> \<Union>\<U> \<or> topspace Y \<subseteq> \<Union>\<V>"
using \<U> \<V> \<C> \<C>eq by auto
then have *: "\<exists>\<D>. finite \<D> \<and>
(\<forall>x \<in> \<D>. x \<in> (\<lambda>V. topspace X \<times> V) ` \<V> \<or> x \<in> (\<lambda>U. U \<times> topspace Y) ` \<U>) \<and>
(topspace X \<times> topspace Y \<subseteq> \<Union>\<D>)"
proof
assume "topspace X \<subseteq> \<Union>\<U>"
with \<open>compact_space X\<close> \<U> obtain \<F> where "finite \<F>" "\<F> \<subseteq> \<U>" "topspace X \<subseteq> \<Union>\<F>"
by (meson compact_space_alt)
with that show ?thesis
by (rule_tac x="(\<lambda>D. D \<times> topspace Y) ` \<F>" in exI) auto
next
assume "topspace Y \<subseteq> \<Union>\<V>"
with \<open>compact_space Y\<close> \<V> obtain \<F> where "finite \<F>" "\<F> \<subseteq> \<V>" "topspace Y \<subseteq> \<Union>\<F>"
by (meson compact_space_alt)
with that show ?thesis
by (rule_tac x="(\<lambda>C. topspace X \<times> C) ` \<F>" in exI) auto
qed
then show ?thesis
using that \<U> \<V> by blast
qed
then show "\<exists>\<D>. finite \<D> \<and> \<D> \<subseteq> \<C> \<and> topspace X \<times> topspace Y \<subseteq> \<Union> \<D>"
using \<C> \<C>eq by blast
next
have "(finite intersection_of (\<lambda>x. x \<in> \<X> \<or> x \<in> \<Y>) relative_to topspace X \<times> topspace Y)
= (\<lambda>U. \<exists>S T. U = S \<times> T \<and> openin X S \<and> openin Y T)"
(is "?lhs = ?rhs")
proof -
have "?rhs U" if "?lhs U" for U
proof -
have "topspace X \<times> topspace Y \<inter> \<Inter> T \<in> {A \<times> B |A B. A \<in> Collect (openin X) \<and> B \<in> Collect (openin Y)}"
if "finite T" "T \<subseteq> \<X> \<union> \<Y>" for T
using that
proof induction
case (insert B \<B>)
then show ?case
unfolding \<X>_def \<Y>_def
apply (simp add: Int_ac subset_eq image_def)
apply (metis (no_types) openin_Int openin_topspace Times_Int_Times)
done
qed auto
then show ?thesis
using that
by (auto simp: subset_eq elim!: relative_toE intersection_ofE)
qed
moreover
have "?lhs Z" if Z: "?rhs Z" for Z
proof -
obtain U V where "Z = U \<times> V" "openin X U" "openin Y V"
using Z by blast
then have UV: "U \<times> V = (topspace X \<times> topspace Y) \<inter> (U \<times> V)"
by (simp add: Sigma_mono inf_absorb2 openin_subset)
moreover
have "?lhs ((topspace X \<times> topspace Y) \<inter> (U \<times> V))"
proof (rule relative_to_inc)
show "(finite intersection_of (\<lambda>x. x \<in> \<X> \<or> x \<in> \<Y>)) (U \<times> V)"
apply (simp add: intersection_of_def \<X>_def \<Y>_def)
apply (rule_tac x="{(U \<times> topspace Y),(topspace X \<times> V)}" in exI)
using \<open>openin X U\<close> \<open>openin Y V\<close> openin_subset UV apply (fastforce simp add:)
done
qed
ultimately show ?thesis
using \<open>Z = U \<times> V\<close> by auto
qed
ultimately show ?thesis
by meson
qed
then show "topology (arbitrary union_of (finite intersection_of (\<lambda>x. x \<in> \<X> \<union> \<Y>)
relative_to (topspace X \<times> topspace Y))) =
prod_topology X Y"
by (simp add: prod_topology_def)
qed
ultimately show ?thesis
using False by blast
qed
lemma compactin_Times:
"compactin (prod_topology X Y) (S \<times> T) \<longleftrightarrow> S = {} \<or> T = {} \<or> compactin X S \<and> compactin Y T"
by (auto simp: compactin_subspace subtopology_Times compact_space_prod_topology)
subsection\<open>Homeomorphic maps\<close>
lemma homeomorphic_maps_prod:
"homeomorphic_maps (prod_topology X Y) (prod_topology X' Y') (\<lambda>(x,y). (f x, g y)) (\<lambda>(x,y). (f' x, g' y)) \<longleftrightarrow>
topspace(prod_topology X Y) = {} \<and>
topspace(prod_topology X' Y') = {} \<or>
homeomorphic_maps X X' f f' \<and>
homeomorphic_maps Y Y' g g'"
unfolding homeomorphic_maps_def continuous_map_prod_top
by (auto simp: continuous_map_def homeomorphic_maps_def continuous_map_prod_top)
lemma homeomorphic_maps_swap:
"homeomorphic_maps (prod_topology X Y) (prod_topology Y X)
(\<lambda>(x,y). (y,x)) (\<lambda>(y,x). (x,y))"
by (auto simp: homeomorphic_maps_def case_prod_unfold continuous_map_fst continuous_map_pairedI continuous_map_snd)
lemma homeomorphic_map_swap:
"homeomorphic_map (prod_topology X Y) (prod_topology Y X) (\<lambda>(x,y). (y,x))"
using homeomorphic_map_maps homeomorphic_maps_swap by metis
lemma embedding_map_graph:
"embedding_map X (prod_topology X Y) (\<lambda>x. (x, f x)) \<longleftrightarrow> continuous_map X Y f"
(is "?lhs = ?rhs")
proof
assume L: ?lhs
have "snd \<circ> (\<lambda>x. (x, f x)) = f"
by force
moreover have "continuous_map X Y (snd \<circ> (\<lambda>x. (x, f x)))"
using L
unfolding embedding_map_def
by (meson continuous_map_in_subtopology continuous_map_snd_of homeomorphic_imp_continuous_map)
ultimately show ?rhs
by simp
next
assume R: ?rhs
then show ?lhs
unfolding homeomorphic_map_maps embedding_map_def homeomorphic_maps_def
by (rule_tac x=fst in exI)
(auto simp: continuous_map_in_subtopology continuous_map_paired continuous_map_from_subtopology
continuous_map_fst)
qed
lemma homeomorphic_space_prod_topology:
"\<lbrakk>X homeomorphic_space X''; Y homeomorphic_space Y'\<rbrakk>
\<Longrightarrow> prod_topology X Y homeomorphic_space prod_topology X'' Y'"
using homeomorphic_maps_prod unfolding homeomorphic_space_def by blast
lemma prod_topology_homeomorphic_space_left:
"topspace Y = {b} \<Longrightarrow> prod_topology X Y homeomorphic_space X"
unfolding homeomorphic_space_def
by (rule_tac x=fst in exI) (simp add: homeomorphic_map_def inj_on_def flip: homeomorphic_map_maps)
lemma prod_topology_homeomorphic_space_right:
"topspace X = {a} \<Longrightarrow> prod_topology X Y homeomorphic_space Y"
unfolding homeomorphic_space_def
by (rule_tac x=snd in exI) (simp add: homeomorphic_map_def inj_on_def flip: homeomorphic_map_maps)
lemma homeomorphic_space_prod_topology_sing1:
"b \<in> topspace Y \<Longrightarrow> X homeomorphic_space (prod_topology X (subtopology Y {b}))"
by (metis empty_subsetI homeomorphic_space_sym inf.absorb_iff2 insert_subset prod_topology_homeomorphic_space_left topspace_subtopology)
lemma homeomorphic_space_prod_topology_sing2:
"a \<in> topspace X \<Longrightarrow> Y homeomorphic_space (prod_topology (subtopology X {a}) Y)"
by (metis empty_subsetI homeomorphic_space_sym inf.absorb_iff2 insert_subset prod_topology_homeomorphic_space_right topspace_subtopology)
lemma topological_property_of_prod_component:
assumes major: "P(prod_topology X Y)"
and X: "\<And>x. \<lbrakk>x \<in> topspace X; P(prod_topology X Y)\<rbrakk> \<Longrightarrow> P(subtopology (prod_topology X Y) ({x} \<times> topspace Y))"
and Y: "\<And>y. \<lbrakk>y \<in> topspace Y; P(prod_topology X Y)\<rbrakk> \<Longrightarrow> P(subtopology (prod_topology X Y) (topspace X \<times> {y}))"
and PQ: "\<And>X X'. X homeomorphic_space X' \<Longrightarrow> (P X \<longleftrightarrow> Q X')"
and PR: "\<And>X X'. X homeomorphic_space X' \<Longrightarrow> (P X \<longleftrightarrow> R X')"
shows "topspace(prod_topology X Y) = {} \<or> Q X \<and> R Y"
proof -
have "Q X \<and> R Y" if "topspace(prod_topology X Y) \<noteq> {}"
proof -
from that obtain a b where a: "a \<in> topspace X" and b: "b \<in> topspace Y"
by force
show ?thesis
using X [OF a major] and Y [OF b major] homeomorphic_space_prod_topology_sing1 [OF b, of X] homeomorphic_space_prod_topology_sing2 [OF a, of Y]
by (simp add: subtopology_Times) (meson PQ PR homeomorphic_space_prod_topology_sing2 homeomorphic_space_sym)
qed
then show ?thesis by metis
qed
lemma limitin_pairwise:
"limitin (prod_topology X Y) f l F \<longleftrightarrow> limitin X (fst \<circ> f) (fst l) F \<and> limitin Y (snd \<circ> f) (snd l) F"
(is "?lhs = ?rhs")
proof
assume ?lhs
then obtain a b where ev: "\<And>U. \<lbrakk>(a,b) \<in> U; openin (prod_topology X Y) U\<rbrakk> \<Longrightarrow> \<forall>\<^sub>F x in F. f x \<in> U"
and a: "a \<in> topspace X" and b: "b \<in> topspace Y" and l: "l = (a,b)"
by (auto simp: limitin_def)
moreover have "\<forall>\<^sub>F x in F. fst (f x) \<in> U" if "openin X U" "a \<in> U" for U
proof -
have "\<forall>\<^sub>F c in F. f c \<in> U \<times> topspace Y"
using b that ev [of "U \<times> topspace Y"] by (auto simp: openin_prod_topology_alt)
then show ?thesis
by (rule eventually_mono) (metis (mono_tags, lifting) SigmaE2 prod.collapse)
qed
moreover have "\<forall>\<^sub>F x in F. snd (f x) \<in> U" if "openin Y U" "b \<in> U" for U
proof -
have "\<forall>\<^sub>F c in F. f c \<in> topspace X \<times> U"
using a that ev [of "topspace X \<times> U"] by (auto simp: openin_prod_topology_alt)
then show ?thesis
by (rule eventually_mono) (metis (mono_tags, lifting) SigmaE2 prod.collapse)
qed
ultimately show ?rhs
by (simp add: limitin_def)
next
have "limitin (prod_topology X Y) f (a,b) F"
if "limitin X (fst \<circ> f) a F" "limitin Y (snd \<circ> f) b F" for a b
using that
proof (clarsimp simp: limitin_def)
fix Z :: "('a \<times> 'b) set"
assume a: "a \<in> topspace X" "\<forall>U. openin X U \<and> a \<in> U \<longrightarrow> (\<forall>\<^sub>F x in F. fst (f x) \<in> U)"
and b: "b \<in> topspace Y" "\<forall>U. openin Y U \<and> b \<in> U \<longrightarrow> (\<forall>\<^sub>F x in F. snd (f x) \<in> U)"
and Z: "openin (prod_topology X Y) Z" "(a, b) \<in> Z"
then obtain U V where "openin X U" "openin Y V" "a \<in> U" "b \<in> V" "U \<times> V \<subseteq> Z"
using Z by (force simp: openin_prod_topology_alt)
then have "\<forall>\<^sub>F x in F. fst (f x) \<in> U" "\<forall>\<^sub>F x in F. snd (f x) \<in> V"
by (simp_all add: a b)
then show "\<forall>\<^sub>F x in F. f x \<in> Z"
by (rule eventually_elim2) (use \<open>U \<times> V \<subseteq> Z\<close> subsetD in auto)
qed
then show "?rhs \<Longrightarrow> ?lhs"
by (metis prod.collapse)
qed
end
diff --git a/src/HOL/Analysis/Regularity.thy b/src/HOL/Analysis/Regularity.thy
--- a/src/HOL/Analysis/Regularity.thy
+++ b/src/HOL/Analysis/Regularity.thy
@@ -1,391 +1,391 @@
(* Title: HOL/Analysis/Regularity.thy
Author: Fabian Immler, TU München
*)
section \<open>Regularity of Measures\<close>
theory Regularity (* FIX suggestion to rename e.g. RegularityMeasures and/ or move as
this theory consists of 1 result only *)
imports Measure_Space Borel_Space
begin
theorem
fixes M::"'a::{second_countable_topology, complete_space} measure"
assumes sb: "sets M = sets borel"
assumes "emeasure M (space M) \<noteq> \<infinity>"
assumes "B \<in> sets borel"
shows inner_regular: "emeasure M B =
(SUP K \<in> {K. K \<subseteq> B \<and> compact K}. emeasure M K)" (is "?inner B")
and outer_regular: "emeasure M B =
(INF U \<in> {U. B \<subseteq> U \<and> open U}. emeasure M U)" (is "?outer B")
proof -
have Us: "UNIV = space M" by (metis assms(1) sets_eq_imp_space_eq space_borel)
hence sU: "space M = UNIV" by simp
interpret finite_measure M by rule fact
have approx_inner: "\<And>A. A \<in> sets M \<Longrightarrow>
(\<And>e. e > 0 \<Longrightarrow> \<exists>K. K \<subseteq> A \<and> compact K \<and> emeasure M A \<le> emeasure M K + ennreal e) \<Longrightarrow> ?inner A"
by (rule ennreal_approx_SUP)
(force intro!: emeasure_mono simp: compact_imp_closed emeasure_eq_measure)+
have approx_outer: "\<And>A. A \<in> sets M \<Longrightarrow>
(\<And>e. e > 0 \<Longrightarrow> \<exists>B. A \<subseteq> B \<and> open B \<and> emeasure M B \<le> emeasure M A + ennreal e) \<Longrightarrow> ?outer A"
by (rule ennreal_approx_INF)
(force intro!: emeasure_mono simp: emeasure_eq_measure sb)+
from countable_dense_setE guess X::"'a set" . note X = this
{
fix r::real assume "r > 0" hence "\<And>y. open (ball y r)" "\<And>y. ball y r \<noteq> {}" by auto
with X(2)[OF this]
have x: "space M = (\<Union>x\<in>X. cball x r)"
by (auto simp add: sU) (metis dist_commute order_less_imp_le)
let ?U = "\<Union>k. (\<Union>n\<in>{0..k}. cball (from_nat_into X n) r)"
have "(\<lambda>k. emeasure M (\<Union>n\<in>{0..k}. cball (from_nat_into X n) r)) \<longlonglongrightarrow> M ?U"
by (rule Lim_emeasure_incseq) (auto intro!: borel_closed bexI simp: incseq_def Us sb)
also have "?U = space M"
proof safe
fix x from X(2)[OF open_ball[of x r]] \<open>r > 0\<close> obtain d where d: "d\<in>X" "d \<in> ball x r" by auto
show "x \<in> ?U"
using X(1) d
by simp (auto intro!: exI [where x = "to_nat_on X d"] simp: dist_commute Bex_def)
qed (simp add: sU)
finally have "(\<lambda>k. M (\<Union>n\<in>{0..k}. cball (from_nat_into X n) r)) \<longlonglongrightarrow> M (space M)" .
} note M_space = this
{
fix e ::real and n :: nat assume "e > 0" "n > 0"
hence "1/n > 0" "e * 2 powr - n > 0" by (auto)
from M_space[OF \<open>1/n>0\<close>]
have "(\<lambda>k. measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n))) \<longlonglongrightarrow> measure M (space M)"
- unfolding emeasure_eq_measure by (auto simp: measure_nonneg)
+ unfolding emeasure_eq_measure by (auto)
from metric_LIMSEQ_D[OF this \<open>0 < e * 2 powr -n\<close>]
obtain k where "dist (measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n))) (measure M (space M)) <
e * 2 powr -n"
by auto
hence "measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n)) \<ge>
measure M (space M) - e * 2 powr -real n"
by (auto simp: dist_real_def)
hence "\<exists>k. measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n)) \<ge>
measure M (space M) - e * 2 powr - real n" ..
} note k=this
hence "\<forall>e\<in>{0<..}. \<forall>(n::nat)\<in>{0<..}. \<exists>k.
measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n)) \<ge> measure M (space M) - e * 2 powr - real n"
by blast
then obtain k where k: "\<forall>e\<in>{0<..}. \<forall>n\<in>{0<..}. measure M (space M) - e * 2 powr - real (n::nat)
\<le> measure M (\<Union>i\<in>{0..k e n}. cball (from_nat_into X i) (1 / n))"
by metis
hence k: "\<And>e n. e > 0 \<Longrightarrow> n > 0 \<Longrightarrow> measure M (space M) - e * 2 powr - n
\<le> measure M (\<Union>i\<in>{0..k e n}. cball (from_nat_into X i) (1 / n))"
unfolding Ball_def by blast
have approx_space:
"\<exists>K \<in> {K. K \<subseteq> space M \<and> compact K}. emeasure M (space M) \<le> emeasure M K + ennreal e"
(is "?thesis e") if "0 < e" for e :: real
proof -
define B where [abs_def]:
"B n = (\<Union>i\<in>{0..k e (Suc n)}. cball (from_nat_into X i) (1 / Suc n))" for n
have "\<And>n. closed (B n)" by (auto simp: B_def)
hence [simp]: "\<And>n. B n \<in> sets M" by (simp add: sb)
from k[OF \<open>e > 0\<close> zero_less_Suc]
have "\<And>n. measure M (space M) - measure M (B n) \<le> e * 2 powr - real (Suc n)"
by (simp add: algebra_simps B_def finite_measure_compl)
hence B_compl_le: "\<And>n::nat. measure M (space M - B n) \<le> e * 2 powr - real (Suc n)"
by (simp add: finite_measure_compl)
define K where "K = (\<Inter>n. B n)"
from \<open>closed (B _)\<close> have "closed K" by (auto simp: K_def)
hence [simp]: "K \<in> sets M" by (simp add: sb)
have "measure M (space M) - measure M K = measure M (space M - K)"
by (simp add: finite_measure_compl)
also have "\<dots> = emeasure M (\<Union>n. space M - B n)" by (auto simp: K_def emeasure_eq_measure)
also have "\<dots> \<le> (\<Sum>n. emeasure M (space M - B n))"
by (rule emeasure_subadditive_countably) (auto simp: summable_def)
also have "\<dots> \<le> (\<Sum>n. ennreal (e*2 powr - real (Suc n)))"
- using B_compl_le by (intro suminf_le) (simp_all add: measure_nonneg emeasure_eq_measure ennreal_leI)
+ using B_compl_le by (intro suminf_le) (simp_all add: emeasure_eq_measure ennreal_leI)
also have "\<dots> \<le> (\<Sum>n. ennreal (e * (1 / 2) ^ Suc n))"
by (simp add: powr_minus powr_realpow field_simps del: of_nat_Suc)
also have "\<dots> = ennreal e * (\<Sum>n. ennreal ((1 / 2) ^ Suc n))"
unfolding ennreal_power[symmetric]
using \<open>0 < e\<close>
by (simp add: ac_simps ennreal_mult' divide_ennreal[symmetric] divide_ennreal_def
ennreal_power[symmetric])
also have "\<dots> = e"
by (subst suminf_ennreal_eq[OF zero_le_power power_half_series]) auto
finally have "measure M (space M) \<le> measure M K + e"
using \<open>0 < e\<close> by simp
hence "emeasure M (space M) \<le> emeasure M K + e"
using \<open>0 < e\<close> by (simp add: emeasure_eq_measure flip: ennreal_plus)
moreover have "compact K"
unfolding compact_eq_totally_bounded
proof safe
show "complete K" using \<open>closed K\<close> by (simp add: complete_eq_closed)
fix e'::real assume "0 < e'"
from nat_approx_posE[OF this] guess n . note n = this
let ?k = "from_nat_into X ` {0..k e (Suc n)}"
have "finite ?k" by simp
moreover have "K \<subseteq> (\<Union>x\<in>?k. ball x e')" unfolding K_def B_def using n by force
ultimately show "\<exists>k. finite k \<and> K \<subseteq> (\<Union>x\<in>k. ball x e')" by blast
qed
ultimately
show ?thesis by (auto simp: sU)
qed
{ fix A::"'a set" assume "closed A" hence "A \<in> sets borel" by (simp add: compact_imp_closed)
hence [simp]: "A \<in> sets M" by (simp add: sb)
have "?inner A"
proof (rule approx_inner)
fix e::real assume "e > 0"
from approx_space[OF this] obtain K where
K: "K \<subseteq> space M" "compact K" "emeasure M (space M) \<le> emeasure M K + e"
by (auto simp: emeasure_eq_measure)
hence [simp]: "K \<in> sets M" by (simp add: sb compact_imp_closed)
have "measure M A - measure M (A \<inter> K) = measure M (A - A \<inter> K)"
by (subst finite_measure_Diff) auto
also have "A - A \<inter> K = A \<union> K - K" by auto
also have "measure M \<dots> = measure M (A \<union> K) - measure M K"
by (subst finite_measure_Diff) auto
also have "\<dots> \<le> measure M (space M) - measure M K"
by (simp add: emeasure_eq_measure sU sb finite_measure_mono)
also have "\<dots> \<le> e"
using K \<open>0 < e\<close> by (simp add: emeasure_eq_measure flip: ennreal_plus)
finally have "emeasure M A \<le> emeasure M (A \<inter> K) + ennreal e"
using \<open>0<e\<close> by (simp add: emeasure_eq_measure algebra_simps flip: ennreal_plus)
moreover have "A \<inter> K \<subseteq> A" "compact (A \<inter> K)" using \<open>closed A\<close> \<open>compact K\<close> by auto
ultimately show "\<exists>K \<subseteq> A. compact K \<and> emeasure M A \<le> emeasure M K + ennreal e"
by blast
qed simp
have "?outer A"
proof cases
assume "A \<noteq> {}"
let ?G = "\<lambda>d. {x. infdist x A < d}"
{
fix d
have "?G d = (\<lambda>x. infdist x A) -` {..<d}" by auto
also have "open \<dots>"
by (intro continuous_open_vimage) (auto intro!: continuous_infdist continuous_ident)
finally have "open (?G d)" .
} note open_G = this
from in_closed_iff_infdist_zero[OF \<open>closed A\<close> \<open>A \<noteq> {}\<close>]
have "A = {x. infdist x A = 0}" by auto
also have "\<dots> = (\<Inter>i. ?G (1/real (Suc i)))"
proof (auto simp del: of_nat_Suc, rule ccontr)
fix x
assume "infdist x A \<noteq> 0"
hence pos: "infdist x A > 0" using infdist_nonneg[of x A] by simp
from nat_approx_posE[OF this] guess n .
moreover
assume "\<forall>i. infdist x A < 1 / real (Suc i)"
hence "infdist x A < 1 / real (Suc n)" by auto
ultimately show False by simp
qed
also have "M \<dots> = (INF n. emeasure M (?G (1 / real (Suc n))))"
proof (rule INF_emeasure_decseq[symmetric], safe)
fix i::nat
from open_G[of "1 / real (Suc i)"]
show "?G (1 / real (Suc i)) \<in> sets M" by (simp add: sb borel_open)
next
show "decseq (\<lambda>i. {x. infdist x A < 1 / real (Suc i)})"
by (auto intro: less_trans intro!: divide_strict_left_mono
simp: decseq_def le_eq_less_or_eq)
qed simp
finally
have "emeasure M A = (INF n. emeasure M {x. infdist x A < 1 / real (Suc n)})" .
moreover
have "\<dots> \<ge> (INF U\<in>{U. A \<subseteq> U \<and> open U}. emeasure M U)"
proof (intro INF_mono)
fix m
have "?G (1 / real (Suc m)) \<in> {U. A \<subseteq> U \<and> open U}" using open_G by auto
moreover have "M (?G (1 / real (Suc m))) \<le> M (?G (1 / real (Suc m)))" by simp
ultimately show "\<exists>U\<in>{U. A \<subseteq> U \<and> open U}.
emeasure M U \<le> emeasure M {x. infdist x A < 1 / real (Suc m)}"
by blast
qed
moreover
have "emeasure M A \<le> (INF U\<in>{U. A \<subseteq> U \<and> open U}. emeasure M U)"
by (rule INF_greatest) (auto intro!: emeasure_mono simp: sb)
ultimately show ?thesis by simp
qed (auto intro!: INF_eqI)
note \<open>?inner A\<close> \<open>?outer A\<close> }
note closed_in_D = this
from \<open>B \<in> sets borel\<close>
have "Int_stable (Collect closed)" "Collect closed \<subseteq> Pow UNIV" "B \<in> sigma_sets UNIV (Collect closed)"
by (auto simp: Int_stable_def borel_eq_closed)
then show "?inner B" "?outer B"
proof (induct B rule: sigma_sets_induct_disjoint)
case empty
{ case 1 show ?case by (intro SUP_eqI[symmetric]) auto }
{ case 2 show ?case by (intro INF_eqI[symmetric]) (auto elim!: meta_allE[of _ "{}"]) }
next
case (basic B)
{ case 1 from basic closed_in_D show ?case by auto }
{ case 2 from basic closed_in_D show ?case by auto }
next
case (compl B)
note inner = compl(2) and outer = compl(3)
from compl have [simp]: "B \<in> sets M" by (auto simp: sb borel_eq_closed)
case 2
have "M (space M - B) = M (space M) - emeasure M B" by (auto simp: emeasure_compl)
also have "\<dots> = (INF K\<in>{K. K \<subseteq> B \<and> compact K}. M (space M) - M K)"
by (subst ennreal_SUP_const_minus) (auto simp: less_top[symmetric] inner)
also have "\<dots> = (INF U\<in>{U. U \<subseteq> B \<and> compact U}. M (space M - U))"
by (auto simp add: emeasure_compl sb compact_imp_closed)
also have "\<dots> \<ge> (INF U\<in>{U. U \<subseteq> B \<and> closed U}. M (space M - U))"
by (rule INF_superset_mono) (auto simp add: compact_imp_closed)
also have "(INF U\<in>{U. U \<subseteq> B \<and> closed U}. M (space M - U)) =
(INF U\<in>{U. space M - B \<subseteq> U \<and> open U}. emeasure M U)"
apply (rule arg_cong [of _ _ Inf])
using sU
apply (auto simp add: image_iff)
apply (rule exI [of _ "UNIV - y" for y])
apply safe
apply (auto simp add: double_diff)
done
finally have
"(INF U\<in>{U. space M - B \<subseteq> U \<and> open U}. emeasure M U) \<le> emeasure M (space M - B)" .
moreover have
"(INF U\<in>{U. space M - B \<subseteq> U \<and> open U}. emeasure M U) \<ge> emeasure M (space M - B)"
by (auto simp: sb sU intro!: INF_greatest emeasure_mono)
ultimately show ?case by (auto intro!: antisym simp: sets_eq_imp_space_eq[OF sb])
case 1
have "M (space M - B) = M (space M) - emeasure M B" by (auto simp: emeasure_compl)
also have "\<dots> = (SUP U\<in> {U. B \<subseteq> U \<and> open U}. M (space M) - M U)"
unfolding outer by (subst ennreal_INF_const_minus) auto
also have "\<dots> = (SUP U\<in>{U. B \<subseteq> U \<and> open U}. M (space M - U))"
by (auto simp add: emeasure_compl sb compact_imp_closed)
also have "\<dots> = (SUP K\<in>{K. K \<subseteq> space M - B \<and> closed K}. emeasure M K)"
unfolding SUP_image [of _ "\<lambda>u. space M - u" _, symmetric, unfolded comp_def]
apply (rule arg_cong [of _ _ Sup])
using sU apply (auto intro!: imageI)
done
also have "\<dots> = (SUP K\<in>{K. K \<subseteq> space M - B \<and> compact K}. emeasure M K)"
proof (safe intro!: antisym SUP_least)
fix K assume "closed K" "K \<subseteq> space M - B"
from closed_in_D[OF \<open>closed K\<close>]
have K_inner: "emeasure M K = (SUP K\<in>{Ka. Ka \<subseteq> K \<and> compact Ka}. emeasure M K)" by simp
show "emeasure M K \<le> (SUP K\<in>{K. K \<subseteq> space M - B \<and> compact K}. emeasure M K)"
unfolding K_inner using \<open>K \<subseteq> space M - B\<close>
by (auto intro!: SUP_upper SUP_least)
qed (fastforce intro!: SUP_least SUP_upper simp: compact_imp_closed)
finally show ?case by (auto intro!: antisym simp: sets_eq_imp_space_eq[OF sb])
next
case (union D)
then have "range D \<subseteq> sets M" by (auto simp: sb borel_eq_closed)
with union have M[symmetric]: "(\<Sum>i. M (D i)) = M (\<Union>i. D i)" by (intro suminf_emeasure)
also have "(\<lambda>n. \<Sum>i<n. M (D i)) \<longlonglongrightarrow> (\<Sum>i. M (D i))"
by (intro summable_LIMSEQ) auto
finally have measure_LIMSEQ: "(\<lambda>n. \<Sum>i<n. measure M (D i)) \<longlonglongrightarrow> measure M (\<Union>i. D i)"
- by (simp add: emeasure_eq_measure measure_nonneg sum_nonneg)
+ by (simp add: emeasure_eq_measure sum_nonneg)
have "(\<Union>i. D i) \<in> sets M" using \<open>range D \<subseteq> sets M\<close> by auto
case 1
show ?case
proof (rule approx_inner)
fix e::real assume "e > 0"
with measure_LIMSEQ
have "\<exists>no. \<forall>n\<ge>no. \<bar>(\<Sum>i<n. measure M (D i)) -measure M (\<Union>x. D x)\<bar> < e/2"
by (auto simp: lim_sequentially dist_real_def simp del: less_divide_eq_numeral1)
hence "\<exists>n0. \<bar>(\<Sum>i<n0. measure M (D i)) - measure M (\<Union>x. D x)\<bar> < e/2" by auto
then obtain n0 where n0: "\<bar>(\<Sum>i<n0. measure M (D i)) - measure M (\<Union>i. D i)\<bar> < e/2"
unfolding choice_iff by blast
have "ennreal (\<Sum>i<n0. measure M (D i)) = (\<Sum>i<n0. M (D i))"
- by (auto simp add: emeasure_eq_measure sum_nonneg measure_nonneg)
+ by (auto simp add: emeasure_eq_measure)
also have "\<dots> \<le> (\<Sum>i. M (D i))" by (rule sum_le_suminf) auto
also have "\<dots> = M (\<Union>i. D i)" by (simp add: M)
also have "\<dots> = measure M (\<Union>i. D i)" by (simp add: emeasure_eq_measure)
finally have n0: "measure M (\<Union>i. D i) - (\<Sum>i<n0. measure M (D i)) < e/2"
- using n0 by (auto simp: measure_nonneg sum_nonneg)
+ using n0 by (auto simp: sum_nonneg)
have "\<forall>i. \<exists>K. K \<subseteq> D i \<and> compact K \<and> emeasure M (D i) \<le> emeasure M K + e/(2*Suc n0)"
proof
fix i
from \<open>0 < e\<close> have "0 < e/(2*Suc n0)" by simp
have "emeasure M (D i) = (SUP K\<in>{K. K \<subseteq> (D i) \<and> compact K}. emeasure M K)"
using union by blast
from SUP_approx_ennreal[OF \<open>0 < e/(2*Suc n0)\<close> _ this]
show "\<exists>K. K \<subseteq> D i \<and> compact K \<and> emeasure M (D i) \<le> emeasure M K + e/(2*Suc n0)"
by (auto simp: emeasure_eq_measure intro: less_imp_le compact_empty)
qed
then obtain K where K: "\<And>i. K i \<subseteq> D i" "\<And>i. compact (K i)"
"\<And>i. emeasure M (D i) \<le> emeasure M (K i) + e/(2*Suc n0)"
unfolding choice_iff by blast
let ?K = "\<Union>i\<in>{..<n0}. K i"
have "disjoint_family_on K {..<n0}" using K \<open>disjoint_family D\<close>
unfolding disjoint_family_on_def by blast
hence mK: "measure M ?K = (\<Sum>i<n0. measure M (K i))" using K
by (intro finite_measure_finite_Union) (auto simp: sb compact_imp_closed)
have "measure M (\<Union>i. D i) < (\<Sum>i<n0. measure M (D i)) + e/2" using n0 by simp
also have "(\<Sum>i<n0. measure M (D i)) \<le> (\<Sum>i<n0. measure M (K i) + e/(2*Suc n0))"
using K \<open>0 < e\<close>
by (auto intro: sum_mono simp: emeasure_eq_measure simp flip: ennreal_plus)
also have "\<dots> = (\<Sum>i<n0. measure M (K i)) + (\<Sum>i<n0. e/(2*Suc n0))"
by (simp add: sum.distrib)
also have "\<dots> \<le> (\<Sum>i<n0. measure M (K i)) + e / 2" using \<open>0 < e\<close>
by (auto simp: field_simps intro!: mult_left_mono)
finally
have "measure M (\<Union>i. D i) < (\<Sum>i<n0. measure M (K i)) + e / 2 + e / 2"
by auto
hence "M (\<Union>i. D i) < M ?K + e"
using \<open>0<e\<close> by (auto simp: mK emeasure_eq_measure sum_nonneg ennreal_less_iff simp flip: ennreal_plus)
moreover
have "?K \<subseteq> (\<Union>i. D i)" using K by auto
moreover
have "compact ?K" using K by auto
ultimately
have "?K\<subseteq>(\<Union>i. D i) \<and> compact ?K \<and> emeasure M (\<Union>i. D i) \<le> emeasure M ?K + ennreal e" by simp
thus "\<exists>K\<subseteq>\<Union>i. D i. compact K \<and> emeasure M (\<Union>i. D i) \<le> emeasure M K + ennreal e" ..
qed fact
case 2
show ?case
proof (rule approx_outer[OF \<open>(\<Union>i. D i) \<in> sets M\<close>])
fix e::real assume "e > 0"
have "\<forall>i::nat. \<exists>U. D i \<subseteq> U \<and> open U \<and> e/(2 powr Suc i) > emeasure M U - emeasure M (D i)"
proof
fix i::nat
from \<open>0 < e\<close> have "0 < e/(2 powr Suc i)" by simp
have "emeasure M (D i) = (INF U\<in>{U. (D i) \<subseteq> U \<and> open U}. emeasure M U)"
using union by blast
from INF_approx_ennreal[OF \<open>0 < e/(2 powr Suc i)\<close> this]
show "\<exists>U. D i \<subseteq> U \<and> open U \<and> e/(2 powr Suc i) > emeasure M U - emeasure M (D i)"
using \<open>0<e\<close>
by (auto simp: emeasure_eq_measure sum_nonneg ennreal_less_iff ennreal_minus
finite_measure_mono sb
simp flip: ennreal_plus)
qed
then obtain U where U: "\<And>i. D i \<subseteq> U i" "\<And>i. open (U i)"
"\<And>i. e/(2 powr Suc i) > emeasure M (U i) - emeasure M (D i)"
unfolding choice_iff by blast
let ?U = "\<Union>i. U i"
have "ennreal (measure M ?U - measure M (\<Union>i. D i)) = M ?U - M (\<Union>i. D i)"
using U(1,2)
by (subst ennreal_minus[symmetric])
- (auto intro!: finite_measure_mono simp: sb measure_nonneg emeasure_eq_measure)
+ (auto intro!: finite_measure_mono simp: sb emeasure_eq_measure)
also have "\<dots> = M (?U - (\<Union>i. D i))" using U \<open>(\<Union>i. D i) \<in> sets M\<close>
by (subst emeasure_Diff) (auto simp: sb)
also have "\<dots> \<le> M (\<Union>i. U i - D i)" using U \<open>range D \<subseteq> sets M\<close>
by (intro emeasure_mono) (auto simp: sb intro!: sets.countable_nat_UN sets.Diff)
also have "\<dots> \<le> (\<Sum>i. M (U i - D i))" using U \<open>range D \<subseteq> sets M\<close>
by (intro emeasure_subadditive_countably) (auto intro!: sets.Diff simp: sb)
also have "\<dots> \<le> (\<Sum>i. ennreal e/(2 powr Suc i))" using U \<open>range D \<subseteq> sets M\<close>
using \<open>0<e\<close>
by (intro suminf_le, subst emeasure_Diff)
- (auto simp: emeasure_Diff emeasure_eq_measure sb measure_nonneg ennreal_minus
+ (auto simp: emeasure_Diff emeasure_eq_measure sb ennreal_minus
finite_measure_mono divide_ennreal ennreal_less_iff
intro: less_imp_le)
also have "\<dots> \<le> (\<Sum>n. ennreal (e * (1 / 2) ^ Suc n))"
using \<open>0<e\<close>
by (simp add: powr_minus powr_realpow field_simps divide_ennreal del: of_nat_Suc)
also have "\<dots> = ennreal e * (\<Sum>n. ennreal ((1 / 2) ^ Suc n))"
unfolding ennreal_power[symmetric]
using \<open>0 < e\<close>
by (simp add: ac_simps ennreal_mult' divide_ennreal[symmetric] divide_ennreal_def
ennreal_power[symmetric])
also have "\<dots> = ennreal e"
by (subst suminf_ennreal_eq[OF zero_le_power power_half_series]) auto
finally have "emeasure M ?U \<le> emeasure M (\<Union>i. D i) + ennreal e"
using \<open>0<e\<close> by (simp add: emeasure_eq_measure flip: ennreal_plus)
moreover
have "(\<Union>i. D i) \<subseteq> ?U" using U by auto
moreover
have "open ?U" using U by auto
ultimately
have "(\<Union>i. D i) \<subseteq> ?U \<and> open ?U \<and> emeasure M ?U \<le> emeasure M (\<Union>i. D i) + ennreal e" by simp
thus "\<exists>B. (\<Union>i. D i) \<subseteq> B \<and> open B \<and> emeasure M B \<le> emeasure M (\<Union>i. D i) + ennreal e" ..
qed
qed
qed
end
diff --git a/src/HOL/Analysis/Simplex_Content.thy b/src/HOL/Analysis/Simplex_Content.thy
--- a/src/HOL/Analysis/Simplex_Content.thy
+++ b/src/HOL/Analysis/Simplex_Content.thy
@@ -1,264 +1,264 @@
(*
File: Analysis/Simplex_Content.thy
Author: Manuel Eberl <eberlm@in.tum.de>
The content of an n-dimensional simplex, including the formula for the content of a
triangle and Heron's formula.
*)
section \<open>Volume of a Simplex\<close>
theory Simplex_Content
imports Change_Of_Vars
begin
lemma fact_neq_top_ennreal [simp]: "fact n \<noteq> (top :: ennreal)"
by (induction n) (auto simp: ennreal_mult_eq_top_iff)
lemma ennreal_fact: "ennreal (fact n) = fact n"
by (induction n) (auto simp: ennreal_mult algebra_simps ennreal_of_nat_eq_real_of_nat)
context
fixes S :: "'a set \<Rightarrow> real \<Rightarrow> ('a \<Rightarrow> real) set"
defines "S \<equiv> (\<lambda>A t. {x. (\<forall>i\<in>A. 0 \<le> x i) \<and> sum x A \<le> t})"
begin
lemma emeasure_std_simplex_aux_step:
assumes "b \<notin> A" "finite A"
shows "x(b := y) \<in> S (insert b A) t \<longleftrightarrow> y \<in> {0..t} \<and> x \<in> S A (t - y)"
using assms sum_nonneg[of A x] unfolding S_def
by (force simp: sum_delta_notmem algebra_simps)
lemma emeasure_std_simplex_aux:
fixes t :: real
assumes "finite (A :: 'a set)" "t \<ge> 0"
shows "emeasure (Pi\<^sub>M A (\<lambda>_. lborel))
(S A t \<inter> space (Pi\<^sub>M A (\<lambda>_. lborel))) = t ^ card A / fact (card A)"
using assms(1,2)
proof (induction arbitrary: t rule: finite_induct)
case (empty t)
thus ?case by (simp add: PiM_empty S_def)
next
case (insert b A t)
define n where "n = Suc (card A)"
have n_pos: "n > 0" by (simp add: n_def)
let ?M = "\<lambda>A. (Pi\<^sub>M A (\<lambda>_. lborel))"
{
fix A :: "'a set" and t :: real assume "finite A"
have "S A t \<inter> space (Pi\<^sub>M A (\<lambda>_. lborel)) =
Pi\<^sub>E A (\<lambda>_. {0..}) \<inter> (\<lambda>x. sum x A) -` {..t} \<inter> space (Pi\<^sub>M A (\<lambda>_. lborel))"
by (auto simp: S_def space_PiM)
also have "\<dots> \<in> sets (Pi\<^sub>M A (\<lambda>_. lborel))"
using \<open>finite A\<close> by measurable
finally have "S A t \<inter> space (Pi\<^sub>M A (\<lambda>_. lborel)) \<in> sets (Pi\<^sub>M A (\<lambda>_. lborel))" .
} note meas [measurable] = this
interpret product_sigma_finite "\<lambda>_. lborel"
by standard
have "emeasure (?M (insert b A)) (S (insert b A) t \<inter> space (?M (insert b A))) =
nn_integral (?M (insert b A))
(\<lambda>x. indicator (S (insert b A) t \<inter> space (?M (insert b A))) x)"
using insert.hyps by (subst nn_integral_indicator) auto
also have "\<dots> = (\<integral>\<^sup>+ y. \<integral>\<^sup>+ x. indicator (S (insert b A) t \<inter> space (?M (insert b A)))
(x(b := y)) \<partial>?M A \<partial>lborel)"
using insert.prems insert.hyps by (intro product_nn_integral_insert_rev) auto
also have "\<dots> = (\<integral>\<^sup>+ y. \<integral>\<^sup>+ x. indicator {0..t} y * indicator (S A (t - y) \<inter> space (?M A)) x
\<partial>?M A \<partial>lborel)"
using insert.hyps insert.prems emeasure_std_simplex_aux_step[of b A]
by (intro nn_integral_cong)
(auto simp: fun_eq_iff indicator_def space_PiM PiE_def extensional_def)
also have "\<dots> = (\<integral>\<^sup>+ y. indicator {0..t} y * (\<integral>\<^sup>+ x. indicator (S A (t - y) \<inter> space (?M A)) x
\<partial>?M A) \<partial>lborel)" using \<open>finite A\<close>
by (subst nn_integral_cmult) auto
also have "\<dots> = (\<integral>\<^sup>+ y. indicator {0..t} y * emeasure (?M A) (S A (t - y) \<inter> space (?M A)) \<partial>lborel)"
using \<open>finite A\<close> by (subst nn_integral_indicator) auto
also have "\<dots> = (\<integral>\<^sup>+ y. indicator {0..t} y * (t - y) ^ card A / ennreal (fact (card A)) \<partial>lborel)"
using insert.IH by (intro nn_integral_cong) (auto simp: indicator_def divide_ennreal)
also have "\<dots> = (\<integral>\<^sup>+ y. indicator {0..t} y * (t - y) ^ card A \<partial>lborel) / ennreal (fact (card A))"
using \<open>finite A\<close> by (subst nn_integral_divide) auto
also have "(\<integral>\<^sup>+ y. indicator {0..t} y * (t - y) ^ card A \<partial>lborel) =
(\<integral>\<^sup>+y\<in>{0..t}. ennreal ((t - y) ^ (n - 1)) \<partial>lborel)"
by (intro nn_integral_cong) (auto simp: indicator_def n_def)
also have "((\<lambda>x. - ((t - x) ^ n / n)) has_real_derivative (t - x) ^ (n - 1)) (at x)"
if "x \<in> {0..t}" for x by (rule derivative_eq_intros refl | simp add: n_pos)+
hence "(\<integral>\<^sup>+y\<in>{0..t}. ennreal ((t - y) ^ (n - 1)) \<partial>lborel) =
ennreal (-((t - t) ^ n / n) - (-((t - 0) ^ n / n)))"
using insert.prems insert.hyps by (intro nn_integral_FTC_Icc) auto
also have "\<dots> = ennreal (t ^ n / n)" using n_pos by (simp add: zero_power)
also have "\<dots> / ennreal (fact (card A)) = ennreal (t ^ n / n / fact (card A))"
using n_pos \<open>t \<ge> 0\<close> by (subst divide_ennreal) auto
also have "t ^ n / n / fact (card A) = t ^ n / fact n"
by (simp add: n_def)
also have "n = card (insert b A)"
using insert.hyps by (subst card_insert) (auto simp: n_def)
finally show ?case .
qed
end
lemma emeasure_std_simplex:
"emeasure lborel (convex hull (insert 0 Basis :: 'a :: euclidean_space set)) =
ennreal (1 / fact DIM('a))"
proof -
have "emeasure lborel {x::'a. (\<forall>i\<in>Basis. 0 \<le> x \<bullet> i) \<and> sum ((\<bullet>) x) Basis \<le> 1} =
emeasure (distr (Pi\<^sub>M Basis (\<lambda>b. lborel)) borel (\<lambda>f. \<Sum>b\<in>Basis. f b *\<^sub>R b))
{x::'a. (\<forall>i\<in>Basis. 0 \<le> x \<bullet> i) \<and> sum ((\<bullet>) x) Basis \<le> 1}"
by (subst lborel_eq) simp
also have "\<dots> = emeasure (Pi\<^sub>M Basis (\<lambda>b. lborel))
({y::'a \<Rightarrow> real. (\<forall>i\<in>Basis. 0 \<le> y i) \<and> sum y Basis \<le> 1} \<inter>
space (Pi\<^sub>M Basis (\<lambda>b. lborel)))"
by (subst emeasure_distr) auto
also have "\<dots> = ennreal (1 / fact DIM('a))"
by (subst emeasure_std_simplex_aux) auto
finally show ?thesis by (simp only: std_simplex)
qed
theorem content_std_simplex:
"measure lborel (convex hull (insert 0 Basis :: 'a :: euclidean_space set)) =
1 / fact DIM('a)"
by (simp add: measure_def emeasure_std_simplex)
(* TODO: move to Change_of_Vars *)
proposition measure_lebesgue_linear_transformation:
fixes A :: "(real ^ 'n :: {finite, wellorder}) set"
fixes f :: "_ \<Rightarrow> real ^ 'n :: {finite, wellorder}"
assumes "bounded A" "A \<in> sets lebesgue" "linear f"
shows "measure lebesgue (f ` A) = \<bar>det (matrix f)\<bar> * measure lebesgue A"
proof -
from assms have [intro]: "A \<in> lmeasurable"
by (intro bounded_set_imp_lmeasurable) auto
hence [intro]: "f ` A \<in> lmeasurable"
by (intro lmeasure_integral measurable_linear_image assms)
have "measure lebesgue (f ` A) = integral (f ` A) (\<lambda>_. 1)"
by (intro lmeasure_integral measurable_linear_image assms) auto
also have "\<dots> = integral (f ` A) (\<lambda>_. 1 :: real ^ 1) $ 0"
by (subst integral_component_eq_cart [symmetric]) (auto intro: integrable_on_const)
also have "\<dots> = \<bar>det (matrix f)\<bar> * integral A (\<lambda>x. 1 :: real ^ 1) $ 0"
using assms
by (subst integral_change_of_variables_linear)
(auto simp: o_def absolutely_integrable_on_def intro: integrable_on_const)
also have "integral A (\<lambda>x. 1 :: real ^ 1) $ 0 = integral A (\<lambda>x. 1)"
by (subst integral_component_eq_cart [symmetric]) (auto intro: integrable_on_const)
also have "\<dots> = measure lebesgue A"
by (intro lmeasure_integral [symmetric]) auto
finally show ?thesis .
qed
theorem content_simplex:
fixes X :: "(real ^ 'n :: {finite, wellorder}) set" and f :: "'n :: _ \<Rightarrow> real ^ ('n :: _)"
assumes "finite X" "card X = Suc CARD('n)" and x0: "x0 \<in> X" and bij: "bij_betw f UNIV (X - {x0})"
defines "M \<equiv> (\<chi> i. \<chi> j. f j $ i - x0 $ i)"
shows "content (convex hull X) = \<bar>det M\<bar> / fact (CARD('n))"
proof -
define g where "g = (\<lambda>x. M *v x)"
have [simp]: "M *v axis i 1 = f i - x0" for i :: 'n
by (simp add: M_def matrix_vector_mult_basis column_def vec_eq_iff)
define std where "std = (convex hull insert 0 Basis :: (real ^ 'n :: _) set)"
have compact: "compact std" unfolding std_def
by (intro finite_imp_compact_convex_hull) auto
have "measure lebesgue (convex hull X) = measure lebesgue (((+) (-x0)) ` (convex hull X))"
by (rule measure_translation [symmetric])
also have "((+) (-x0)) ` (convex hull X) = convex hull (((+) (-x0)) ` X)"
by (rule convex_hull_translation [symmetric])
also have "((+) (-x0)) ` X = insert 0 ((\<lambda>x. x - x0) ` (X - {x0}))"
using x0 by (auto simp: image_iff)
finally have eq: "measure lebesgue (convex hull X) = measure lebesgue (convex hull \<dots>)" .
from compact have "measure lebesgue (g ` std) = \<bar>det M\<bar> * measure lebesgue std"
by (subst measure_lebesgue_linear_transformation)
(auto intro: finite_imp_bounded_convex_hull dest: compact_imp_closed simp: g_def std_def)
also have "measure lebesgue std = content std" using compact
by (intro measure_completion) (auto dest: compact_imp_closed)
also have "content std = 1 / fact CARD('n)" unfolding std_def
by (simp add: content_std_simplex)
also have "g ` std = convex hull (g ` insert 0 Basis)" unfolding std_def
by (rule convex_hull_linear_image) (auto simp: g_def)
also have "g ` insert 0 Basis = insert 0 (g ` Basis)"
by (auto simp: g_def)
also have "g ` Basis = (\<lambda>x. x - x0) ` range f"
by (auto simp: g_def Basis_vec_def image_iff)
also have "range f = X - {x0}" using bij
using bij_betw_imp_surj_on by blast
also note eq [symmetric]
finally show ?thesis
using finite_imp_compact_convex_hull[OF \<open>finite X\<close>] by (auto dest: compact_imp_closed)
qed
theorem content_triangle:
fixes A B C :: "real ^ 2"
shows "content (convex hull {A, B, C}) =
\<bar>(C $ 1 - A $ 1) * (B $ 2 - A $ 2) - (B $ 1 - A $ 1) * (C $ 2 - A $ 2)\<bar> / 2"
proof -
define M :: "real ^ 2 ^ 2" where "M \<equiv> (\<chi> i. \<chi> j. (if j = 1 then B else C) $ i - A $ i)"
define g where "g = (\<lambda>x. M *v x)"
define std where "std = (convex hull insert 0 Basis :: (real ^ 2) set)"
have [simp]: "M *v axis i 1 = (if i = 1 then B - A else C - A)" for i
by (auto simp: M_def matrix_vector_mult_basis column_def vec_eq_iff)
have compact: "compact std" unfolding std_def
by (intro finite_imp_compact_convex_hull) auto
have "measure lebesgue (convex hull {A, B, C}) =
measure lebesgue (((+) (-A)) ` (convex hull {A, B, C}))"
by (rule measure_translation [symmetric])
also have "((+) (-A)) ` (convex hull {A, B, C}) = convex hull (((+) (-A)) ` {A, B, C})"
by (rule convex_hull_translation [symmetric])
also have "((+) (-A)) ` {A, B, C} = {0, B - A, C - A}"
by (auto simp: image_iff)
finally have eq: "measure lebesgue (convex hull {A, B, C}) =
measure lebesgue (convex hull {0, B - A, C - A})" .
from compact have "measure lebesgue (g ` std) = \<bar>det M\<bar> * measure lebesgue std"
by (subst measure_lebesgue_linear_transformation)
(auto intro: finite_imp_bounded_convex_hull dest: compact_imp_closed simp: g_def std_def)
also have "measure lebesgue std = content std" using compact
by (intro measure_completion) (auto dest: compact_imp_closed)
also have "content std = 1 / 2" unfolding std_def
by (simp add: content_std_simplex)
also have "g ` std = convex hull (g ` insert 0 Basis)" unfolding std_def
by (rule convex_hull_linear_image) (auto simp: g_def)
also have "g ` insert 0 Basis = insert 0 (g ` Basis)"
by (auto simp: g_def)
also have "(2 :: 2) \<noteq> 1" by auto
hence "\<not>(\<forall>y::2. y = 1)" by blast
hence "g ` Basis = {B - A, C - A}"
by (auto simp: g_def Basis_vec_def image_iff)
also note eq [symmetric]
finally show ?thesis
using finite_imp_compact_convex_hull[of "{A, B, C}"]
by (auto dest!: compact_imp_closed simp: det_2 M_def)
qed
theorem heron:
fixes A B C :: "real ^ 2"
defines "a \<equiv> dist B C" and "b \<equiv> dist A C" and "c \<equiv> dist A B"
defines "s \<equiv> (a + b + c) / 2"
shows "content (convex hull {A, B, C}) = sqrt (s * (s - a) * (s - b) * (s - c))"
proof -
have [simp]: "(UNIV :: 2 set) = {1, 2}"
using exhaust_2 by auto
have dist_eq: "dist (A :: real ^ 2) B ^ 2 = (A $ 1 - B $ 1) ^ 2 + (A $ 2 - B $ 2) ^ 2"
for A B by (simp add: dist_vec_def dist_real_def)
have nonneg: "s * (s - a) * (s - b) * (s - c) \<ge> 0"
using dist_triangle[of A B C] dist_triangle[of A C B] dist_triangle[of B C A]
by (intro mult_nonneg_nonneg) (auto simp: s_def a_def b_def c_def dist_commute)
have "16 * content (convex hull {A, B, C}) ^ 2 =
4 * ((C $ 1 - A $ 1) * (B $ 2 - A $ 2) - (B $ 1 - A $ 1) * (C $ 2 - A $ 2)) ^ 2"
by (subst content_triangle) (simp add: power_divide)
also have "\<dots> = (2 * (dist A B ^ 2 * dist A C ^ 2 + dist A B ^ 2 * dist B C ^ 2 +
dist A C ^ 2 * dist B C ^ 2) - (dist A B ^ 2) ^ 2 - (dist A C ^ 2) ^ 2 - (dist B C ^ 2) ^ 2)"
unfolding dist_eq unfolding power2_eq_square by algebra
also have "\<dots> = (a + b + c) * ((a + b + c) - 2 * a) * ((a + b + c) - 2 * b) *
((a + b + c) - 2 * c)"
unfolding power2_eq_square by (simp add: s_def a_def b_def c_def algebra_simps)
also have "\<dots> = 16 * s * (s - a) * (s - b) * (s - c)"
- by (simp add: s_def field_split_simps mult_ac)
+ by (simp add: s_def field_split_simps)
finally have "content (convex hull {A, B, C}) ^ 2 = s * (s - a) * (s - b) * (s - c)"
by simp
also have "\<dots> = sqrt (s * (s - a) * (s - b) * (s - c)) ^ 2"
by (intro real_sqrt_pow2 [symmetric] nonneg)
finally show ?thesis using nonneg
by (subst (asm) power2_eq_iff_nonneg) auto
qed
end
diff --git a/src/HOL/Analysis/Starlike.thy b/src/HOL/Analysis/Starlike.thy
--- a/src/HOL/Analysis/Starlike.thy
+++ b/src/HOL/Analysis/Starlike.thy
@@ -1,6588 +1,6588 @@
(* Title: HOL/Analysis/Starlike.thy
Author: L C Paulson, University of Cambridge
Author: Robert Himmelmann, TU Muenchen
Author: Bogdan Grechuk, University of Edinburgh
Author: Armin Heller, TU Muenchen
Author: Johannes Hoelzl, TU Muenchen
*)
chapter \<open>Unsorted\<close>
theory Starlike
imports
Convex_Euclidean_Space
Line_Segment
begin
lemma affine_hull_closed_segment [simp]:
"affine hull (closed_segment a b) = affine hull {a,b}"
by (simp add: segment_convex_hull)
lemma affine_hull_open_segment [simp]:
fixes a :: "'a::euclidean_space"
shows "affine hull (open_segment a b) = (if a = b then {} else affine hull {a,b})"
by (metis affine_hull_convex_hull affine_hull_empty closure_open_segment closure_same_affine_hull segment_convex_hull)
lemma rel_interior_closure_convex_segment:
fixes S :: "_::euclidean_space set"
assumes "convex S" "a \<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
lemma convex_hull_insert_segments:
"convex hull (insert a S) =
(if S = {} then {a} else \<Union>x \<in> convex hull S. closed_segment a x)"
by (force simp add: convex_hull_insert_alt in_segment)
lemma Int_convex_hull_insert_rel_exterior:
fixes z :: "'a::euclidean_space"
assumes "convex C" "T \<subseteq> C" and z: "z \<in> rel_interior C" and dis: "disjnt S (rel_interior C)"
shows "S \<inter> (convex hull (insert z T)) = S \<inter> (convex hull T)" (is "?lhs = ?rhs")
proof
have "T = {} \<Longrightarrow> z \<notin> S"
using dis z by (auto simp add: disjnt_def)
then show "?lhs \<subseteq> ?rhs"
proof (clarsimp simp add: convex_hull_insert_segments)
fix x y
assume "x \<in> S" and y: "y \<in> convex hull T" and "x \<in> closed_segment z y"
have "y \<in> closure C"
by (metis y \<open>convex C\<close> \<open>T \<subseteq> C\<close> closure_subset contra_subsetD convex_hull_eq hull_mono)
moreover have "x \<notin> rel_interior C"
by (meson \<open>x \<in> S\<close> dis disjnt_iff)
moreover have "x \<in> open_segment z y \<union> {z, y}"
using \<open>x \<in> closed_segment z y\<close> closed_segment_eq_open by blast
ultimately show "x \<in> convex hull T"
using rel_interior_closure_convex_segment [OF \<open>convex C\<close> z]
using y z by blast
qed
show "?rhs \<subseteq> ?lhs"
by (meson hull_mono inf_mono subset_insertI subset_refl)
qed
subsection\<^marker>\<open>tag unimportant\<close> \<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
proof (intro exI subsetI conjI)
fix y
assume "y \<in> ball (x - e *\<^sub>R (x - c)) (e*d)"
then have as: "dist (x - e *\<^sub>R (x - c)) y < e * d"
by simp
have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x"
using \<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)
by simp (simp add: field_simps norm_minus_commute)
then show ?thesis
unfolding *
using mem_interior_convex_shrink \<open>y \<in> S\<close> assms by blast
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
lemma convex_closure_interior:
fixes S :: "'a::euclidean_space set"
assumes "convex S" and int: "interior S \<noteq> {}"
shows "closure(interior S) = closure S"
proof -
obtain a where a: "a \<in> interior S"
using int by auto
have "closure S \<subseteq> closure(interior S)"
proof
fix x
assume x: "x \<in> closure S"
show "x \<in> closure (interior S)"
proof (cases "x=a")
case True
then show ?thesis
using \<open>a \<in> interior S\<close> closure_subset by blast
next
case False
show ?thesis
proof (clarsimp simp add: closure_def islimpt_approachable)
fix e::real
assume xnotS: "x \<notin> interior S" and "0 < e"
show "\<exists>x'\<in>interior S. x' \<noteq> x \<and> dist x' x < e"
proof (intro bexI conjI)
show "x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a) \<noteq> x"
using False \<open>0 < e\<close> by (auto simp: algebra_simps min_def)
show "dist (x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a)) x < e"
using \<open>0 < e\<close> by (auto simp: dist_norm min_def)
show "x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a) \<in> interior S"
apply (clarsimp simp add: min_def a)
apply (rule mem_interior_closure_convex_shrink [OF \<open>convex S\<close> a x])
using \<open>0 < e\<close> False apply (auto simp: field_split_simps)
done
qed
qed
qed
qed
then show ?thesis
by (simp add: closure_mono interior_subset subset_antisym)
qed
lemma closure_convex_Int_superset:
fixes S :: "'a::euclidean_space set"
assumes "convex S" "interior S \<noteq> {}" "interior S \<subseteq> closure T"
shows "closure(S \<inter> T) = closure S"
proof -
have "closure S \<subseteq> closure(interior S)"
by (simp add: convex_closure_interior assms)
also have "... \<subseteq> closure (S \<inter> T)"
using interior_subset [of S] assms
by (metis (no_types, lifting) Int_assoc Int_lower2 closure_mono closure_open_Int_superset inf.orderE open_interior)
finally show ?thesis
by (simp add: closure_mono dual_order.antisym)
qed
subsection\<^marker>\<open>tag unimportant\<close> \<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> sum u S \<le> 1 \<and> sum (\<lambda>x. u x *\<^sub>R x) S = y}"
proof (simp add: convex_hull_finite set_eq_iff assms, safe)
fix x and u :: "'a \<Rightarrow> real"
assume "0 \<le> u 0" "\<forall>x\<in>S. 0 \<le> u x" "u 0 + sum u S = 1"
then show "\<exists>v. (\<forall>x\<in>S. 0 \<le> v x) \<and> sum v S \<le> 1 \<and> (\<Sum>x\<in>S. v x *\<^sub>R x) = (\<Sum>x\<in>S. u x *\<^sub>R x)"
by force
next
fix x and u :: "'a \<Rightarrow> real"
assume "\<forall>x\<in>S. 0 \<le> u x" "sum u S \<le> 1"
then show "\<exists>v. 0 \<le> v 0 \<and> (\<forall>x\<in>S. 0 \<le> v x) \<and> v 0 + sum v S = 1 \<and> (\<Sum>x\<in>S. v x *\<^sub>R x) = (\<Sum>x\<in>S. u x *\<^sub>R x)"
by (rule_tac x="\<lambda>x. if x = 0 then 1 - sum u S else u x" in exI) (auto simp: sum_delta_notmem assms if_smult)
qed
lemma substd_simplex:
assumes d: "d \<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>]
proof (intro set_eqI; safe)
fix u :: "'a \<Rightarrow> real"
assume as: "\<forall>x\<in>?D. 0 \<le> u x" "sum u ?D \<le> 1"
let ?x = "(\<Sum>x\<in>?D. u x *\<^sub>R x)"
have ind: "\<forall>i\<in>Basis. i \<in> d \<longrightarrow> u i = ?x \<bullet> i"
and notind: "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> ?x \<bullet> i = 0)"
using substdbasis_expansion_unique[OF assms] by blast+
then have **: "sum u ?D = sum ((\<bullet>) ?x) ?D"
using assms by (auto intro!: sum.cong)
show "0 \<le> ?x \<bullet> i" if "i \<in> Basis" for i
using as(1) ind notind that by fastforce
show "sum ((\<bullet>) ?x) ?D \<le> 1"
using "**" as(2) by linarith
show "?x \<bullet> i = 0" if "i \<in> Basis" "i \<notin> d" for i
using notind that by blast
next
fix x
assume "\<forall>i\<in>Basis. 0 \<le> x \<bullet> i" "sum ((\<bullet>) x) ?D \<le> 1" "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)"
with d show "\<exists>u. (\<forall>x\<in>?D. 0 \<le> u x) \<and> sum u ?D \<le> 1 \<and> (\<Sum>x\<in>?D. u x *\<^sub>R x) = x"
unfolding substdbasis_expansion_unique[OF assms]
by (rule_tac x="inner x" in exI) auto
qed
qed
lemma std_simplex:
"convex hull (insert 0 Basis) =
{x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> sum (\<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> sum (\<lambda>i. x\<bullet>i) Basis < 1}"
unfolding set_eq_iff mem_interior std_simplex
proof (intro allI iffI CollectI; clarify)
fix x :: 'a
fix e
assume "e > 0" and as: "ball x e \<subseteq> {x. (\<forall>i\<in>Basis. 0 \<le> x \<bullet> i) \<and> sum ((\<bullet>) x) Basis \<le> 1}"
show "(\<forall>i\<in>Basis. 0 < x \<bullet> i) \<and> sum ((\<bullet>) x) Basis < 1"
proof safe
fix i :: 'a
assume i: "i \<in> Basis"
then show "0 < x \<bullet> i"
using as[THEN subsetD[where c="x - (e / 2) *\<^sub>R i"]] and \<open>e > 0\<close>
by (force simp add: 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 *: "sum ((\<bullet>) (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis))) Basis =
sum (\<lambda>i. x\<bullet>i + (if (SOME i. i\<in>Basis) = i then e/2 else 0)) Basis"
by (auto simp: intro!: sum.cong)
have "sum ((\<bullet>) x) Basis < sum ((\<bullet>) (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis))) Basis"
using \<open>e > 0\<close> DIM_positive by (auto simp: SOME_Basis sum.distrib *)
also have "\<dots> \<le> 1"
using ** as by force
finally show "sum ((\<bullet>) x) Basis < 1" by auto
qed
next
fix x :: 'a
assume as: "\<forall>i\<in>Basis. 0 < x \<bullet> i" "sum ((\<bullet>) x) Basis < 1"
obtain a :: 'b where "a \<in> UNIV" using UNIV_witness ..
let ?d = "(1 - sum ((\<bullet>) x) Basis) / real (DIM('a))"
show "\<exists>e>0. ball x e \<subseteq> {x. (\<forall>i\<in>Basis. 0 \<le> x \<bullet> i) \<and> sum ((\<bullet>) x) Basis \<le> 1}"
proof (rule_tac x="min (Min (((\<bullet>) x) ` Basis)) D" for D in exI, intro conjI subsetI CollectI)
fix y
assume y: "y \<in> ball x (min (Min ((\<bullet>) x ` Basis)) ?d)"
have "sum ((\<bullet>) y) Basis \<le> sum (\<lambda>i. x\<bullet>i + ?d) Basis"
proof (rule sum_mono)
fix i :: 'a
assume i: "i \<in> Basis"
have "\<bar>y\<bullet>i - x\<bullet>i\<bar> \<le> norm (y - x)"
by (metis Basis_le_norm i inner_commute inner_diff_right)
also have "... < ?d"
using y by (simp add: dist_norm norm_minus_commute)
finally have "\<bar>y\<bullet>i - x\<bullet>i\<bar> < ?d" .
then show "y \<bullet> i \<le> x \<bullet> i + ?d" by auto
qed
also have "\<dots> \<le> 1"
unfolding sum.distrib sum_constant
by (auto simp add: Suc_le_eq)
finally show "sum ((\<bullet>) y) Basis \<le> 1" .
show "(\<forall>i\<in>Basis. 0 \<le> y \<bullet> i)"
proof safe
fix i :: 'a
assume i: "i \<in> Basis"
have "norm (x - y) < Min (((\<bullet>) x) ` Basis)"
using y by (auto simp: dist_norm less_eq_real_def)
also have "... \<le> x\<bullet>i"
using i by auto
finally have "norm (x - y) < x\<bullet>i" .
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
next
have "Min (((\<bullet>) x) ` Basis) > 0"
using as by simp
moreover have "?d > 0"
using as by (auto simp: Suc_le_eq)
ultimately show "0 < min (Min ((\<bullet>) x ` Basis)) ((1 - sum ((\<bullet>) x) Basis) / real DIM('a))"
by linarith
qed
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 = "sum (\<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 _ "sum (\<lambda>j. if i = j then inverse (2 * real DIM('a)) else 0) ?D"])
(simp_all add: sum.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)
next
have "sum ((\<bullet>) ?a) ?D = sum (\<lambda>i. inverse (2 * real DIM('a))) ?D"
apply (rule sum.cong)
apply rule
apply auto
done
also have "\<dots> < 1"
unfolding sum_constant divide_inverse[symmetric]
by (auto simp add: field_simps)
finally show "sum ((\<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"
using D finite_Basis finite_subset by blast
show ?thesis
proof (cases "D = {}")
case True
then show ?thesis
using rel_interior_sing using euclidean_eq_iff[of _ 0] by auto
next
case False
have h0: "affine hull (convex hull (insert 0 ?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 "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 [rule_format]: "\<And>y. dist x y < e \<and> (\<forall>i\<in>Basis. i \<notin> D \<longrightarrow> y\<bullet>i = 0) \<longrightarrow>
(\<forall>i\<in>D. 0 \<le> y \<bullet> i) \<and> sum ((\<bullet>) y) 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> sum ((\<bullet>) x) D < 1 \<and> (\<forall>i\<in>Basis. i \<notin> D \<longrightarrow> x\<bullet>i = 0)"
proof (intro conjI ballI)
fix i :: 'a
assume "i \<in> D"
then have "\<forall>j\<in>D. 0 \<le> (x - (e / 2) *\<^sub>R i) \<bullet> j"
apply -
apply (rule as[THEN conjunct1])
using D \<open>e > 0\<close> x0
apply (auto simp: dist_norm inner_simps inner_Basis)
done
then show "0 < x \<bullet> i"
using \<open>e > 0\<close> \<open>i \<in> D\<close> D by (force simp: inner_simps inner_Basis)
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 *: "sum ((\<bullet>) (x + (e / 2) *\<^sub>R a)) D =
sum (\<lambda>i. x\<bullet>i + (if a = i then e/2 else 0)) D"
using D by (intro sum.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 "sum ((\<bullet>) x) D < sum ((\<bullet>) (x + (e / 2) *\<^sub>R a)) D"
using \<open>e > 0\<close> \<open>a \<in> D\<close> \<open>finite D\<close> by (auto simp add: * sum.distrib)
also have "\<dots> \<le> 1"
using ** h1 as[rule_format, of "x + (e / 2) *\<^sub>R a"]
by auto
finally show "sum ((\<bullet>) x) D < 1" "\<And>i. i\<in>Basis \<Longrightarrow> 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 - sum ((\<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 (((\<bullet>) x) ` D) > 0"
using as \<open>D \<noteq> {}\<close> \<open>finite D\<close> by (simp)
moreover have "?d > 0" using as using \<open>0 < card D\<close> by auto
ultimately have h3: "min (Min (((\<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 (((\<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 ((\<bullet>) x ` D)) ?d"
assume y2: "\<forall>i\<in>Basis. i \<notin> D \<longrightarrow> y\<bullet>i = 0"
have "sum ((\<bullet>) y) D \<le> sum (\<lambda>i. x\<bullet>i + ?d) D"
proof (rule sum_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> \<le> norm (y - x)"
by (metis i inner_commute inner_diff_right norm_bound_Basis_le order_refl)
also have "... < ?d"
by (metis dist_norm min_less_iff_conj norm_minus_commute y)
finally have "\<bar>y\<bullet>i - x\<bullet>i\<bar> < ?d" .
then show "y \<bullet> i \<le> x \<bullet> i + ?d" by auto
qed
also have "\<dots> \<le> 1"
unfolding sum.distrib sum_constant using \<open>0 < card D\<close>
by auto
finally show "sum ((\<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 "(\<bullet>) x ` D" "norm (x - y)"] \<open>0 < card D\<close> \<open>i \<in> 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> sum ((\<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 = "sum (\<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 _ "sum (\<lambda>j. if i = j then inverse (2 * real (card D)) else 0) ?D"])
unfolding inner_sum_left
apply (rule sum.cong)
using \<open>i \<in> D\<close> \<open>finite D\<close> sum.delta'[of D i "(\<lambda>k. inverse (2 * real (card D)))"]
d1 assms(2)
by (auto simp: inner_Basis rev_subsetD[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 "sum ((\<bullet>) ?a) ?D = sum (\<lambda>i. inverse (2 * real (card D))) ?D"
by (rule sum.cong) (rule refl, rule **)
also have "\<dots> < 1"
unfolding sum_constant divide_real_def[symmetric]
by (auto simp add: field_simps)
finally show "sum ((\<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_sum[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_base[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\<^marker>\<open>tag unimportant\<close> \<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 metis
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_superset 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> (+) (-a) ` S"
using assms exI[of "(\<lambda>x. x \<in> S \<and> - a + x = 0)" a] by auto
then have "rel_interior ((+) (-a) ` S) \<noteq> {}"
using rel_interior_convex_nonempty_aux[of "(+) (-a) ` S"]
convex_translation[of S "-a"] assms
by auto
then have "rel_interior S \<noteq> {}"
using rel_interior_translation [of "- a"] by simp
}
then show ?thesis by auto
qed
lemma interior_simplex_nonempty:
fixes S :: "'N :: euclidean_space set"
assumes "independent S" "finite S" "card S = DIM('N)"
obtains a where "a \<in> interior (convex hull (insert 0 S))"
proof -
have "affine hull (insert 0 S) = UNIV"
by (simp add: hull_inc affine_hull_span_0 dim_eq_full[symmetric]
assms(1) assms(3) dim_eq_card_independent)
moreover have "rel_interior (convex hull insert 0 S) \<noteq> {}"
using rel_interior_eq_empty [of "convex hull (insert 0 S)"] by auto
ultimately have "interior (convex hull insert 0 S) \<noteq> {}"
by (simp add: rel_interior_interior)
with that show ?thesis
by auto
qed
lemma convex_rel_interior:
fixes S :: "'n::euclidean_space set"
assumes "convex S"
shows "convex (rel_interior S)"
proof -
{
fix x y and u :: real
assume assm: "x \<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) \<in> 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 = {}" 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 (top_of_set (affine hull (rel_interior S))) (rel_interior S)"
using openin_rel_interior[of S] rel_interior_same_affine_hull[of S] assms by auto
then show ?thesis
using rel_interior_def by auto
qed
lemma rel_interior_rel_open:
fixes S :: "'n::euclidean_space set"
assumes "convex S"
shows "rel_open (rel_interior S)"
unfolding rel_open_def using rel_interior_rel_interior assms by auto
lemma convex_rel_interior_closure_aux:
fixes x y z :: "'n::euclidean_space"
assumes "0 < a" "0 < b" "(a + b) *\<^sub>R z = a *\<^sub>R x + b *\<^sub>R y"
obtains e where "0 < e" "e \<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)"
using assms by (simp add: eq_vector_fraction_iff)
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 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
subsection\<open>The relative frontier of a set\<close>
definition\<^marker>\<open>tag important\<close> "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"
unfolding rel_frontier_def
using rel_interior_subset_closure by (auto simp add: rel_interior_eq_closure [symmetric])
lemma rel_frontier_sing [simp]:
fixes a :: "'n::euclidean_space"
shows "rel_frontier {a} = {}"
by (simp add: rel_frontier_def)
lemma rel_frontier_affine_hull:
fixes S :: "'a::euclidean_space set"
shows "rel_frontier S \<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 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 closest_point_in_rel_frontier:
"\<lbrakk>closed S; S \<noteq> {}; x \<in> affine hull S - rel_interior S\<rbrakk>
\<Longrightarrow> closest_point S x \<in> rel_frontier S"
by (simp add: closest_point_in_rel_interior closest_point_in_set rel_frontier_def)
lemma closed_rel_frontier [iff]:
fixes S :: "'n::euclidean_space set"
shows "closed (rel_frontier S)"
proof -
have *: "closedin (top_of_set (affine hull S)) (closure S - rel_interior S)"
by (simp add: closed_subset closedin_diff closure_affine_hull openin_rel_interior)
show ?thesis
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)
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\<^marker>\<open>tag unimportant\<close> \<open>Relative interior and closure under common operations\<close>
lemma rel_interior_inter_aux: "\<Inter>{rel_interior S |S. S \<in> I} \<subseteq> \<Inter>I"
proof -
{
fix y
assume "y \<in> \<Inter>{rel_interior S |S. S \<in> 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 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 rel_interior_convex_Int_affine:
fixes S :: "'a::euclidean_space set"
assumes "convex S" "affine T" "interior S \<inter> T \<noteq> {}"
shows "rel_interior(S \<inter> T) = interior S \<inter> T"
proof -
obtain a where aS: "a \<in> interior S" and aT:"a \<in> T"
using assms by force
have "rel_interior S = interior S"
by (metis (no_types) aS affine_hull_nonempty_interior equals0D rel_interior_interior)
then show ?thesis
by (metis (no_types) affine_imp_convex assms convex_rel_interior_inter_two hull_same rel_interior_affine_hull)
qed
lemma closure_convex_Int_affine:
fixes S :: "'a::euclidean_space set"
assumes "convex S" "affine T" "rel_interior S \<inter> T \<noteq> {}"
shows "closure(S \<inter> T) = closure S \<inter> T"
proof
have "closure (S \<inter> T) \<subseteq> closure T"
by (simp add: closure_mono)
also have "... \<subseteq> T"
by (simp add: affine_closed assms)
finally show "closure(S \<inter> T) \<subseteq> closure S \<inter> T"
by (simp add: closure_mono)
next
obtain a where "a \<in> rel_interior S" "a \<in> T"
using assms by auto
then have ssT: "subspace ((\<lambda>x. (-a)+x) ` T)" and "a \<in> S"
using affine_diffs_subspace rel_interior_subset assms by blast+
show "closure S \<inter> T \<subseteq> closure (S \<inter> T)"
proof
fix x assume "x \<in> closure S \<inter> T"
show "x \<in> closure (S \<inter> T)"
proof (cases "x = a")
case True
then show ?thesis
using \<open>a \<in> S\<close> \<open>a \<in> T\<close> closure_subset by fastforce
next
case False
then have "x \<in> closure(open_segment a x)"
by auto
then show ?thesis
using \<open>x \<in> closure S \<inter> T\<close> assms convex_affine_closure_Int by blast
qed
qed
qed
lemma subset_rel_interior_convex:
fixes S T :: "'n::euclidean_space set"
assumes "convex S"
and "convex T"
and "S \<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 by auto
next
case False
interpret linear f by fact
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 \<in> 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 by (simp add: linear_add linear_scale \<open>linear f\<close>)
ultimately have "(1 - e) *\<^sub>R x + e *\<^sub>R z \<in> 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 \<in> 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 -
interpret linear f by fact
have "S \<noteq> {}"
using assms 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 (simp add: assms(2) convex_Int convex_linear_image linear_axioms)
{
fix z
assume "z \<in> f -` (rel_interior S)"
then have z: "f z \<in> 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 (simp add: linear_axioms linear_subspace_image subspace_imp_affine)
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 linear_fst \<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 linear_snd \<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 "((*\<^sub>R) c) ` (rel_interior S) = rel_interior (((*\<^sub>R) c) ` S)"
using rel_interior_injective_linear_image[of "((*\<^sub>R) c)" S]
linear_conv_bounded_linear[of "(*\<^sub>R) c"] linear_scaleR injective_scaleR[of c] assms
by auto
lemma rel_interior_convex_scaleR:
fixes S :: "'n::euclidean_space set"
assumes "convex S"
shows "((*\<^sub>R) c) ` (rel_interior S) = rel_interior (((*\<^sub>R) c) ` S)"
by (metis assms linear_scaleR rel_interior_convex_linear_image)
lemma convex_rel_open_scaleR:
fixes S :: "'n::euclidean_space set"
assumes "convex S"
and "rel_open S"
shows "convex (((*\<^sub>R) c) ` S) \<and> rel_open (((*\<^sub>R) c) ` S)"
by (metis assms convex_scaling rel_interior_convex_scaleR rel_open_def)
lemma convex_rel_open_finite_inter:
assumes "\<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 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 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 linear_fst 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}"] linear_snd 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 \<in> 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\<^marker>\<open>tag unimportant\<close> \<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: cone_0)
next
case False
then have *: "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (*\<^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> (*\<^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> (((*\<^sub>R) c) ` (rel_interior S))"
proof (cases "S = {}")
case True
then show ?thesis
by (simp add: cone_hull_empty)
next
case False
then obtain s where "s \<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 = ((*\<^sub>R) c ` S)"
using f_def cone_hull_expl[of "{1 :: real} \<times> S"] by auto
then have "rel_interior (f c) = (*\<^sub>R) c ` rel_interior S"
using rel_interior_convex_scaleR[of S c] assms by auto
}
then show ?thesis using * ** by auto
qed
lemma rel_interior_convex_cone:
fixes S :: "'m::euclidean_space set"
assumes "convex S"
shows "rel_interior (cone hull ({1 :: real} \<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
then have "z \<in> ?rhs"
using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms \<open>z \<in> ?lhs\<close> by fastforce
}
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)) =
{sum (\<lambda>i. c i *\<^sub>R s i) I | c s. (\<forall>i\<in>I. c i \<ge> 0) \<and> sum 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 \<in> ?rhs"
then obtain c s where *: "sum (\<lambda>i. c i *\<^sub>R s i) I = x" "sum 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_sum[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 *: "sum c I = 1"
using \<open>finite I\<close> \<open>i \<in> I\<close> sum.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 = sum (\<lambda>i. c i *\<^sub>R s i) I"
using s_def c_def \<open>finite I\<close> \<open>i \<in> I\<close> sum.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 = sum (\<lambda>i. c i *\<^sub>R s i) I \<and> (\<forall>i\<in>I. c i \<ge> 0) \<and> sum c I = 1 \<and> (\<forall>i\<in>I. s i \<in> S i)"
by auto
from xy obtain d t where
yc: "y = sum (\<lambda>i. d i *\<^sub>R t i) I \<and> (\<forall>i\<in>I. d i \<ge> 0) \<and> sum 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 "sum (\<lambda>i. u * c i) I = u * sum c I"
by (simp add: sum_distrib_left)
moreover have "sum (\<lambda>i. v * d i) I = v * sum d I"
by (simp add: sum_distrib_left)
ultimately have sum1: "sum e I = 1"
using e_def xc yc uv by (simp add: sum.distrib)
define q where "q i = (if e i = 0 then p i else (u * c i / e i) *\<^sub>R s i + (v * d i / e i) *\<^sub>R t i)"
for i
{
fix i
assume i: "i \<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 = sum (\<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.sum sum.distrib)
also have "\<dots> = sum (\<lambda>i. e i *\<^sub>R q i) I"
using * by auto
finally have "u *\<^sub>R x + v *\<^sub>R y = sum (\<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> sum 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> sum 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
proposition ray_to_rel_frontier:
fixes a :: "'a::real_inner"
assumes "bounded S"
and a: "a \<in> rel_interior S"
and aff: "(a + l) \<in> affine hull S"
and "l \<noteq> 0"
obtains d where "0 < d" "(a + d *\<^sub>R l) \<in> rel_frontier S"
"\<And>e. \<lbrakk>0 \<le> e; e < d\<rbrakk> \<Longrightarrow> (a + e *\<^sub>R l) \<in> rel_interior S"
proof -
have aaff: "a \<in> affine hull S"
by (meson a hull_subset rel_interior_subset rev_subsetD)
let ?D = "{d. 0 < d \<and> a + d *\<^sub>R l \<notin> rel_interior S}"
obtain B where "B > 0" and B: "S \<subseteq> ball a B"
using bounded_subset_ballD [OF \<open>bounded S\<close>] by blast
have "a + (B / norm l) *\<^sub>R l \<notin> ball a B"
by (simp add: dist_norm \<open>l \<noteq> 0\<close>)
with B have "a + (B / norm l) *\<^sub>R l \<notin> rel_interior S"
using rel_interior_subset subsetCE by blast
with \<open>B > 0\<close> \<open>l \<noteq> 0\<close> have nonMT: "?D \<noteq> {}"
using divide_pos_pos zero_less_norm_iff by fastforce
have bdd: "bdd_below ?D"
by (metis (no_types, lifting) bdd_belowI le_less mem_Collect_eq)
have relin_Ex: "\<And>x. x \<in> rel_interior S \<Longrightarrow>
\<exists>e>0. \<forall>x'\<in>affine hull S. dist x' x < e \<longrightarrow> x' \<in> rel_interior S"
using openin_rel_interior [of S] by (simp add: openin_euclidean_subtopology_iff)
define d where "d = Inf ?D"
obtain \<epsilon> where "0 < \<epsilon>" and \<epsilon>: "\<And>\<eta>. \<lbrakk>0 \<le> \<eta>; \<eta> < \<epsilon>\<rbrakk> \<Longrightarrow> (a + \<eta> *\<^sub>R l) \<in> rel_interior S"
proof -
obtain e where "e>0"
and e: "\<And>x'. x' \<in> affine hull S \<Longrightarrow> dist x' a < e \<Longrightarrow> x' \<in> rel_interior S"
using relin_Ex a by blast
show thesis
proof (rule_tac \<epsilon> = "e / norm l" in that)
show "0 < e / norm l" by (simp add: \<open>0 < e\<close> \<open>l \<noteq> 0\<close>)
next
show "a + \<eta> *\<^sub>R l \<in> rel_interior S" if "0 \<le> \<eta>" "\<eta> < e / norm l" for \<eta>
proof (rule e)
show "a + \<eta> *\<^sub>R l \<in> affine hull S"
by (metis (no_types) add_diff_cancel_left' aff affine_affine_hull mem_affine_3_minus aaff)
show "dist (a + \<eta> *\<^sub>R l) a < e"
using that by (simp add: \<open>l \<noteq> 0\<close> dist_norm pos_less_divide_eq)
qed
qed
qed
have inint: "\<And>e. \<lbrakk>0 \<le> e; e < d\<rbrakk> \<Longrightarrow> a + e *\<^sub>R l \<in> rel_interior S"
unfolding d_def using cInf_lower [OF _ bdd]
by (metis (no_types, lifting) a add.right_neutral le_less mem_Collect_eq not_less real_vector.scale_zero_left)
have "\<epsilon> \<le> d"
unfolding d_def
apply (rule cInf_greatest [OF nonMT])
using \<epsilon> dual_order.strict_implies_order le_less_linear by blast
with \<open>0 < \<epsilon>\<close> have "0 < d" by simp
have "a + d *\<^sub>R l \<notin> rel_interior S"
proof
assume adl: "a + d *\<^sub>R l \<in> rel_interior S"
obtain e where "e > 0"
and e: "\<And>x'. x' \<in> affine hull S \<Longrightarrow> dist x' (a + d *\<^sub>R l) < e \<Longrightarrow> x' \<in> rel_interior S"
using relin_Ex adl by blast
have "d + e / norm l \<le> Inf {d. 0 < d \<and> a + d *\<^sub>R l \<notin> rel_interior S}"
proof (rule cInf_greatest [OF nonMT], clarsimp)
fix x::real
assume "0 < x" and nonrel: "a + x *\<^sub>R l \<notin> rel_interior S"
show "d + e / norm l \<le> x"
proof (cases "x < d")
case True with inint nonrel \<open>0 < x\<close>
show ?thesis by auto
next
case False
then have dle: "x < d + e / norm l \<Longrightarrow> dist (a + x *\<^sub>R l) (a + d *\<^sub>R l) < e"
by (simp add: field_simps \<open>l \<noteq> 0\<close>)
have ain: "a + x *\<^sub>R l \<in> affine hull S"
by (metis add_diff_cancel_left' aff affine_affine_hull mem_affine_3_minus aaff)
show ?thesis
using e [OF ain] nonrel dle by force
qed
qed
then show False
using \<open>0 < e\<close> \<open>l \<noteq> 0\<close> by (simp add: d_def [symmetric] field_simps)
qed
moreover have "a + d *\<^sub>R l \<in> closure S"
proof (clarsimp simp: closure_approachable)
fix \<eta>::real assume "0 < \<eta>"
have 1: "a + (d - min d (\<eta> / 2 / norm l)) *\<^sub>R l \<in> S"
apply (rule subsetD [OF rel_interior_subset inint])
using \<open>l \<noteq> 0\<close> \<open>0 < d\<close> \<open>0 < \<eta>\<close> by auto
have "norm l * min d (\<eta> / (norm l * 2)) \<le> norm l * (\<eta> / (norm l * 2))"
by (metis min_def mult_left_mono norm_ge_zero order_refl)
also have "... < \<eta>"
using \<open>l \<noteq> 0\<close> \<open>0 < \<eta>\<close> by (simp add: field_simps)
finally have 2: "norm l * min d (\<eta> / (norm l * 2)) < \<eta>" .
show "\<exists>y\<in>S. dist y (a + d *\<^sub>R l) < \<eta>"
apply (rule_tac x="a + (d - min d (\<eta> / 2 / norm l)) *\<^sub>R l" in bexI)
using 1 2 \<open>0 < d\<close> \<open>0 < \<eta>\<close> apply (auto simp: algebra_simps)
done
qed
ultimately have infront: "a + d *\<^sub>R l \<in> rel_frontier S"
by (simp add: rel_frontier_def)
show ?thesis
by (rule that [OF \<open>0 < d\<close> infront inint])
qed
corollary ray_to_frontier:
fixes a :: "'a::euclidean_space"
assumes "bounded S"
and a: "a \<in> interior S"
and "l \<noteq> 0"
obtains d where "0 < d" "(a + d *\<^sub>R l) \<in> frontier S"
"\<And>e. \<lbrakk>0 \<le> e; e < d\<rbrakk> \<Longrightarrow> (a + e *\<^sub>R l) \<in> interior S"
proof -
have "interior S = rel_interior S"
using a rel_interior_nonempty_interior by auto
then have "a \<in> rel_interior S"
using a by simp
then show ?thesis
apply (rule ray_to_rel_frontier [OF \<open>bounded S\<close> _ _ \<open>l \<noteq> 0\<close>])
using a affine_hull_nonempty_interior apply blast
by (simp add: \<open>interior S = rel_interior S\<close> frontier_def rel_frontier_def that)
qed
lemma segment_to_rel_frontier_aux:
fixes x :: "'a::euclidean_space"
assumes "convex S" "bounded S" and x: "x \<in> rel_interior S" and y: "y \<in> S" and xy: "x \<noteq> y"
obtains z where "z \<in> rel_frontier S" "y \<in> closed_segment x z"
"open_segment x z \<subseteq> rel_interior S"
proof -
have "x + (y - x) \<in> affine hull S"
using hull_inc [OF y] by auto
then obtain d where "0 < d" and df: "(x + d *\<^sub>R (y-x)) \<in> rel_frontier S"
and di: "\<And>e. \<lbrakk>0 \<le> e; e < d\<rbrakk> \<Longrightarrow> (x + e *\<^sub>R (y-x)) \<in> rel_interior S"
by (rule ray_to_rel_frontier [OF \<open>bounded S\<close> x]) (use xy in auto)
show ?thesis
proof
show "x + d *\<^sub>R (y - x) \<in> rel_frontier S"
by (simp add: df)
next
have "open_segment x y \<subseteq> rel_interior S"
using rel_interior_closure_convex_segment [OF \<open>convex S\<close> x] closure_subset y by blast
moreover have "x + d *\<^sub>R (y - x) \<in> open_segment x y" if "d < 1"
using xy
apply (auto simp: in_segment)
apply (rule_tac x="d" in exI)
using \<open>0 < d\<close> that apply (auto simp: algebra_simps)
done
ultimately have "1 \<le> d"
using df rel_frontier_def by fastforce
moreover have "x = (1 / d) *\<^sub>R x + ((d - 1) / d) *\<^sub>R x"
by (metis \<open>0 < d\<close> add.commute add_divide_distrib diff_add_cancel divide_self_if less_irrefl scaleR_add_left scaleR_one)
ultimately show "y \<in> closed_segment x (x + d *\<^sub>R (y - x))"
apply (auto simp: in_segment)
apply (rule_tac x="1/d" in exI)
apply (auto simp: algebra_simps)
done
next
show "open_segment x (x + d *\<^sub>R (y - x)) \<subseteq> rel_interior S"
apply (rule rel_interior_closure_convex_segment [OF \<open>convex S\<close> x])
using df rel_frontier_def by auto
qed
qed
lemma segment_to_rel_frontier:
fixes x :: "'a::euclidean_space"
assumes S: "convex S" "bounded S" and x: "x \<in> rel_interior S"
and y: "y \<in> S" and xy: "\<not>(x = y \<and> S = {x})"
obtains z where "z \<in> rel_frontier S" "y \<in> closed_segment x z"
"open_segment x z \<subseteq> rel_interior S"
proof (cases "x=y")
case True
with xy have "S \<noteq> {x}"
by blast
with True show ?thesis
by (metis Set.set_insert all_not_in_conv ends_in_segment(1) insert_iff segment_to_rel_frontier_aux[OF S x] that y)
next
case False
then show ?thesis
using segment_to_rel_frontier_aux [OF S x y] that by blast
qed
proposition rel_frontier_not_sing:
fixes a :: "'a::euclidean_space"
assumes "bounded S"
shows "rel_frontier S \<noteq> {a}"
proof (cases "S = {}")
case True then show ?thesis by simp
next
case False
then obtain z where "z \<in> S"
by blast
then show ?thesis
proof (cases "S = {z}")
case True then show ?thesis by simp
next
case False
then obtain w where "w \<in> S" "w \<noteq> z"
using \<open>z \<in> S\<close> by blast
show ?thesis
proof
assume "rel_frontier S = {a}"
then consider "w \<notin> rel_frontier S" | "z \<notin> rel_frontier S"
using \<open>w \<noteq> z\<close> by auto
then show False
proof cases
case 1
then have w: "w \<in> rel_interior S"
using \<open>w \<in> S\<close> closure_subset rel_frontier_def by fastforce
have "w + (w - z) \<in> affine hull S"
by (metis \<open>w \<in> S\<close> \<open>z \<in> S\<close> affine_affine_hull hull_inc mem_affine_3_minus scaleR_one)
then obtain e where "0 < e" "(w + e *\<^sub>R (w - z)) \<in> rel_frontier S"
using \<open>w \<noteq> z\<close> \<open>z \<in> S\<close> by (metis assms ray_to_rel_frontier right_minus_eq w)
moreover obtain d where "0 < d" "(w + d *\<^sub>R (z - w)) \<in> rel_frontier S"
using ray_to_rel_frontier [OF \<open>bounded S\<close> w, of "1 *\<^sub>R (z - w)"] \<open>w \<noteq> z\<close> \<open>z \<in> S\<close>
by (metis add.commute add.right_neutral diff_add_cancel hull_inc scaleR_one)
ultimately have "d *\<^sub>R (z - w) = e *\<^sub>R (w - z)"
using \<open>rel_frontier S = {a}\<close> by force
moreover have "e \<noteq> -d "
using \<open>0 < e\<close> \<open>0 < d\<close> by force
ultimately show False
by (metis (no_types, lifting) \<open>w \<noteq> z\<close> eq_iff_diff_eq_0 minus_diff_eq real_vector.scale_cancel_right real_vector.scale_minus_right scaleR_left.minus)
next
case 2
then have z: "z \<in> rel_interior S"
using \<open>z \<in> S\<close> closure_subset rel_frontier_def by fastforce
have "z + (z - w) \<in> affine hull S"
by (metis \<open>z \<in> S\<close> \<open>w \<in> S\<close> affine_affine_hull hull_inc mem_affine_3_minus scaleR_one)
then obtain e where "0 < e" "(z + e *\<^sub>R (z - w)) \<in> rel_frontier S"
using \<open>w \<noteq> z\<close> \<open>w \<in> S\<close> by (metis assms ray_to_rel_frontier right_minus_eq z)
moreover obtain d where "0 < d" "(z + d *\<^sub>R (w - z)) \<in> rel_frontier S"
using ray_to_rel_frontier [OF \<open>bounded S\<close> z, of "1 *\<^sub>R (w - z)"] \<open>w \<noteq> z\<close> \<open>w \<in> S\<close>
by (metis add.commute add.right_neutral diff_add_cancel hull_inc scaleR_one)
ultimately have "d *\<^sub>R (w - z) = e *\<^sub>R (z - w)"
using \<open>rel_frontier S = {a}\<close> by force
moreover have "e \<noteq> -d "
using \<open>0 < e\<close> \<open>0 < d\<close> by force
ultimately show False
by (metis (no_types, lifting) \<open>w \<noteq> z\<close> eq_iff_diff_eq_0 minus_diff_eq real_vector.scale_cancel_right real_vector.scale_minus_right scaleR_left.minus)
qed
qed
qed
qed
subsection\<^marker>\<open>tag unimportant\<close> \<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 (sum S I) = sum (\<lambda>i. rel_interior (S i)) I"
apply (subst sum_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)) = sum S I"
(is "?lhs = ?rhs")
proof -
{
fix x
assume "x \<in> ?lhs"
then obtain c xs where
x: "x = sum (\<lambda>i. c i *\<^sub>R xs i) I \<and> (\<forall>i\<in>I. c i \<ge> 0) \<and> sum 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 = sum s I" using x s_def by auto
ultimately have "x \<in> ?rhs"
using set_sum_alt[of I S] assms by auto
}
moreover
{
fix x
assume "x \<in> ?rhs"
then obtain s where x: "x = sum s I \<and> (\<forall>i\<in>I. s i \<in> S i)"
using set_sum_alt[of I S] assms by auto
define xs where "xs i = of_nat(card I) *\<^sub>R s i" for i
then have "x = sum (\<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 "sum (\<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) = sum A I"
apply (subst convex_hull_finite_union_cones[of I A])
using assms A_def I_def
apply auto
done
moreover have "sum 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))) =
{sum (\<lambda>i. c i *\<^sub>R s i) I | c s. (\<forall>i\<in>I. c i > 0) \<and> sum 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 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> = sum 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 = sum K I" by auto
then have *: "rel_interior K0 = sum (\<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) = sum k I \<and> (\<forall>i\<in>I. k i \<in> rel_interior (K i))"
using \<open>finite I\<close> * set_sum_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> sum c I = 1"
using k by (simp add: sum_prod)
then have "x \<in> ?rhs"
using k cs by auto
}
moreover
{
fix x
assume "x \<in> ?rhs"
then obtain c s where cs: "x = sum (\<lambda>i. c i *\<^sub>R s i) I \<and>
(\<forall>i\<in>I. c i > 0) \<and> sum 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 \<in> 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_sum_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: sum_prod)
done
then have "x \<in> ?lhs"
using K0_def C0_def rel_interior_convex_cone_aux[of C0 1 x]
by auto
}
ultimately show ?thesis by blast
qed
lemma convex_le_Inf_differential:
fixes f :: "real \<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\<^marker>\<open>tag unimportant\<close>\<open>Explicit formulas for interior and relative interior of convex hull\<close>
lemma at_within_cbox_finite:
assumes "x \<in> box a b" "x \<notin> S" "finite S"
shows "(at x within cbox a b - S) = at x"
proof -
have "interior (cbox a b - S) = box a b - S"
using \<open>finite S\<close> by (simp add: interior_diff finite_imp_closed)
then show ?thesis
using at_within_interior assms by fastforce
qed
lemma affine_independent_convex_affine_hull:
fixes s :: "'a::euclidean_space set"
assumes "\<not> 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: "sum u t = 1" "\<forall>x\<in>s. 0 \<le> v x" "sum 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)"
"sum v t + sum v (s - t) = 1"
using uv fin s
by (auto simp: sum.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 sum.union_disjoint, auto simp: algebra_simps sum_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 "\<not> 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: "\<not> 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)
by (auto simp: Suc_le_lessD card_image_le dual_order.trans intro!: dim_le_card'[THEN le_less_trans])
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> sum u s = 1 \<and> sum (\<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> sum u s = 1 \<and> sum (\<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 "\<not> affine_dependent s"
shows "rel_interior(convex hull s) =
{y. \<exists>u. (\<forall>x \<in> s. 0 < u x) \<and> sum u s = 1 \<and> sum (\<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 sum.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" "sum 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> sum 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"
"sum 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 sum_subtractf sum.distrib)
done
} note * = this
have "y \<notin> rel_interior (convex hull s)"
using y
apply (simp add: mem_rel_interior)
apply (auto simp: convex_hull_finite [OF fs])
apply (drule_tac x=u in spec)
apply (auto intro: *)
done
} with rel_interior_subset show "?lhs \<le> ?rhs"
by blast
qed
qed
lemma interior_convex_hull_explicit_minimal:
fixes s :: "'a::euclidean_space set"
shows
"\<not> 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> sum 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_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 "\<not> 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> sum 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" "sum 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> sum u {a,b}"
using a b by simp
also have "... \<le> sum u s"
apply (rule Groups_Big.sum_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\<^marker>\<open>tag unimportant\<close>\<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 "\<not> 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> sum u s = 1 \<and> sum (\<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: sum_subtractf scaleR_diff_left)
done
qed
lemma frontier_convex_hull_explicit:
fixes s :: "'a::euclidean_space set"
assumes "\<not> 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>
sum u s = 1 \<and> sum (\<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 "\<not> 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> sum u s = 1 \<Longrightarrow>
\<exists>x v. x \<in> s \<and>
(\<forall>x\<in>s - {x}. 0 \<le> v x) \<and>
sum 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.sum_diff1 fs)
done }
moreover
{ fix a u
have "a \<in> s \<Longrightarrow> \<forall>x\<in>s - {a}. 0 \<le> u x \<Longrightarrow> sum u (s - {a}) = 1 \<Longrightarrow>
\<exists>v. (\<forall>x\<in>s. 0 \<le> v x) \<and>
(\<exists>x\<in>s. v x = 0) \<and> sum 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: sum.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 "\<not> 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 "\<not> 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)
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\<^marker>\<open>tag important\<close> 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 collinear_between_cases:
fixes c :: "'a::euclidean_space"
shows "collinear {a,b,c} \<longleftrightarrow> between (b,c) a \<or> between (c,a) b \<or> between (a,b) c"
(is "?lhs = ?rhs")
proof
assume ?lhs
then obtain u v where uv: "\<And>x. x \<in> {a, b, c} \<Longrightarrow> \<exists>c. x = u + c *\<^sub>R v"
by (auto simp: collinear_alt)
show ?rhs
using uv [of a] uv [of b] uv [of c] by (auto simp: between_1)
next
assume ?rhs
then show ?lhs
unfolding between_mem_convex_hull
by (metis (no_types, hide_lams) collinear_closed_segment collinear_subset hull_redundant hull_subset insert_commute segment_convex_hull)
qed
lemma subset_continuous_image_segment_1:
fixes f :: "'a::euclidean_space \<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 continuous_injective_image_segment_1:
fixes f :: "'a::euclidean_space \<Rightarrow> real"
assumes contf: "continuous_on (closed_segment a b) f"
and injf: "inj_on f (closed_segment a b)"
shows "f ` (closed_segment a b) = closed_segment (f a) (f b)"
proof
show "closed_segment (f a) (f b) \<subseteq> f ` closed_segment a b"
by (metis subset_continuous_image_segment_1 contf)
show "f ` closed_segment a b \<subseteq> closed_segment (f a) (f b)"
proof (cases "a = b")
case True
then show ?thesis by auto
next
case False
then have fnot: "f a \<noteq> f b"
using inj_onD injf by fastforce
moreover
have "f a \<notin> open_segment (f c) (f b)" if c: "c \<in> closed_segment a b" for c
proof (clarsimp simp add: open_segment_def)
assume fa: "f a \<in> closed_segment (f c) (f b)"
moreover have "closed_segment (f c) (f b) \<subseteq> f ` closed_segment c b"
by (meson closed_segment_subset contf continuous_on_subset convex_closed_segment ends_in_segment(2) subset_continuous_image_segment_1 that)
ultimately have "f a \<in> f ` closed_segment c b"
by blast
then have a: "a \<in> closed_segment c b"
by (meson ends_in_segment inj_on_image_mem_iff_alt injf subset_closed_segment that)
have cb: "closed_segment c b \<subseteq> closed_segment a b"
by (simp add: closed_segment_subset that)
show "f a = f c"
proof (rule between_antisym)
show "between (f c, f b) (f a)"
by (simp add: between_mem_segment fa)
show "between (f a, f b) (f c)"
by (metis a cb between_antisym between_mem_segment between_triv1 subset_iff)
qed
qed
moreover
have "f b \<notin> open_segment (f a) (f c)" if c: "c \<in> closed_segment a b" for c
proof (clarsimp simp add: open_segment_def fnot eq_commute)
assume fb: "f b \<in> closed_segment (f a) (f c)"
moreover have "closed_segment (f a) (f c) \<subseteq> f ` closed_segment a c"
by (meson contf continuous_on_subset ends_in_segment(1) subset_closed_segment subset_continuous_image_segment_1 that)
ultimately have "f b \<in> f ` closed_segment a c"
by blast
then have b: "b \<in> closed_segment a c"
by (meson ends_in_segment inj_on_image_mem_iff_alt injf subset_closed_segment that)
have ca: "closed_segment a c \<subseteq> closed_segment a b"
by (simp add: closed_segment_subset that)
show "f b = f c"
proof (rule between_antisym)
show "between (f c, f a) (f b)"
by (simp add: between_commute between_mem_segment fb)
show "between (f b, f a) (f c)"
by (metis b between_antisym between_commute between_mem_segment between_triv2 that)
qed
qed
ultimately show ?thesis
by (force simp: closed_segment_eq_real_ivl open_segment_eq_real_ivl split: if_split_asm)
qed
qed
lemma continuous_injective_image_open_segment_1:
fixes f :: "'a::euclidean_space \<Rightarrow> real"
assumes contf: "continuous_on (closed_segment a b) f"
and injf: "inj_on f (closed_segment a b)"
shows "f ` (open_segment a b) = open_segment (f a) (f b)"
proof -
have "f ` (open_segment a b) = f ` (closed_segment a b) - {f a, f b}"
by (metis (no_types, hide_lams) empty_subsetI ends_in_segment image_insert image_is_empty inj_on_image_set_diff injf insert_subset open_segment_def segment_open_subset_closed)
also have "... = open_segment (f a) (f b)"
using continuous_injective_image_segment_1 [OF assms]
by (simp add: open_segment_def inj_on_image_set_diff [OF injf])
finally show ?thesis .
qed
lemma collinear_imp_coplanar:
"collinear s ==> coplanar s"
by (metis collinear_affine_hull coplanar_def insert_absorb2)
lemma collinear_small:
assumes "finite s" "card s \<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 numeral_2_eq_2 by (force simp: card_1_singleton_iff)
qed
lemma coplanar_small:
assumes "finite s" "card s \<le> 3"
shows "coplanar s"
proof -
consider "card s \<le> 2" | "card s = Suc (Suc (Suc 0))"
using assms by linarith
then show ?thesis
proof cases
case 1
then show ?thesis
by (simp add: \<open>finite s\<close> collinear_imp_coplanar collinear_small)
next
case 2
then show ?thesis
using hull_subset [of "{_,_,_}"]
by (fastforce simp: coplanar_def dest!: card_eq_SucD)
qed
qed
lemma coplanar_empty: "coplanar {}"
by (simp add: coplanar_small)
lemma coplanar_sing: "coplanar {a}"
by (simp add: coplanar_small)
lemma coplanar_2: "coplanar {a,b}"
by (auto simp: card_insert_if coplanar_small)
lemma coplanar_3: "coplanar {a,b,c}"
by (auto simp: card_insert_if coplanar_small)
lemma collinear_affine_hull_collinear: "collinear(affine hull s) \<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)
lemma collinear_3_expand:
"collinear{a,b,c} \<longleftrightarrow> a = c \<or> (\<exists>u. b = u *\<^sub>R a + (1 - u) *\<^sub>R c)"
proof -
have "collinear{a,b,c} = collinear{a,c,b}"
by (simp add: insert_commute)
also have "... = collinear {0, a - c, b - c}"
by (simp add: collinear_3)
also have "... \<longleftrightarrow> (a = c \<or> b = c \<or> (\<exists>ca. b - c = ca *\<^sub>R (a - c)))"
by (simp add: collinear_lemma)
also have "... \<longleftrightarrow> a = c \<or> (\<exists>u. b = u *\<^sub>R a + (1 - u) *\<^sub>R c)"
by (cases "a = c \<or> b = c") (auto simp: algebra_simps)
finally show ?thesis .
qed
lemma collinear_aff_dim: "collinear S \<longleftrightarrow> aff_dim S \<le> 1"
proof
assume "collinear S"
then obtain u and v :: "'a" where "aff_dim S \<le> aff_dim {u,v}"
by (metis \<open>collinear S\<close> aff_dim_affine_hull aff_dim_subset collinear_affine_hull)
then show "aff_dim S \<le> 1"
using order_trans by fastforce
next
assume "aff_dim S \<le> 1"
then have le1: "aff_dim (affine hull S) \<le> 1"
by simp
obtain B where "B \<subseteq> S" and B: "\<not> affine_dependent B" "affine hull S = affine hull B"
using affine_basis_exists [of S] by auto
then have "finite B" "card B \<le> 2"
using B le1 by (auto simp: affine_independent_iff_card)
then have "collinear B"
by (rule collinear_small)
then show "collinear S"
by (metis \<open>affine hull S = affine hull B\<close> collinear_affine_hull_collinear)
qed
lemma collinear_midpoint: "collinear{a,midpoint a b,b}"
apply (auto simp: collinear_3 collinear_lemma)
apply (drule_tac x="-1" in spec)
apply (simp add: algebra_simps)
done
lemma midpoint_collinear:
fixes a b c :: "'a::real_normed_vector"
assumes "a \<noteq> c"
shows "b = midpoint a c \<longleftrightarrow> collinear{a,b,c} \<and> dist a b = dist b c"
proof -
have *: "a - (u *\<^sub>R a + (1 - u) *\<^sub>R c) = (1 - u) *\<^sub>R (a - c)"
"u *\<^sub>R a + (1 - u) *\<^sub>R c - c = u *\<^sub>R (a - c)"
"\<bar>1 - u\<bar> = \<bar>u\<bar> \<longleftrightarrow> u = 1/2" for u::real
by (auto simp: algebra_simps)
have "b = midpoint a c \<Longrightarrow> collinear{a,b,c} "
using collinear_midpoint by blast
moreover have "collinear{a,b,c} \<Longrightarrow> b = midpoint a c \<longleftrightarrow> dist a b = dist b c"
apply (auto simp: collinear_3_expand assms dist_midpoint)
apply (simp add: dist_norm * assms midpoint_def del: divide_const_simps)
apply (simp add: algebra_simps)
done
ultimately show ?thesis by blast
qed
lemma between_imp_collinear:
fixes x :: "'a :: euclidean_space"
assumes "between (a,b) x"
shows "collinear {a,x,b}"
proof (cases "x = a \<or> x = b \<or> a = b")
case True with assms show ?thesis
by (auto simp: dist_commute)
next
case False with assms show ?thesis
apply (auto simp: collinear_3 collinear_lemma between_norm)
apply (drule_tac x="-(norm(b - x) / norm(x - a))" in spec)
apply (simp add: vector_add_divide_simps real_vector.scale_minus_right [symmetric])
done
qed
lemma midpoint_between:
fixes a b :: "'a::euclidean_space"
shows "b = midpoint a c \<longleftrightarrow> between (a,c) b \<and> dist a b = dist b c"
proof (cases "a = c")
case True then show ?thesis
by (auto simp: dist_commute)
next
case False
show ?thesis
apply (rule iffI)
apply (simp add: between_midpoint(1) dist_midpoint)
using False between_imp_collinear midpoint_collinear by blast
qed
lemma collinear_triples:
assumes "a \<noteq> b"
shows "collinear(insert a (insert b S)) \<longleftrightarrow> (\<forall>x \<in> S. collinear{a,b,x})"
(is "?lhs = ?rhs")
proof safe
fix x
assume ?lhs and "x \<in> S"
then show "collinear {a, b, x}"
using collinear_subset by force
next
assume ?rhs
then have "\<forall>x \<in> S. collinear{a,x,b}"
by (simp add: insert_commute)
then have *: "\<exists>u. x = u *\<^sub>R a + (1 - u) *\<^sub>R b" if "x \<in> (insert a (insert b S))" for x
using that assms collinear_3_expand by fastforce+
show ?lhs
unfolding collinear_def
apply (rule_tac x="b-a" in exI)
apply (clarify dest!: *)
by (metis (no_types, hide_lams) add.commute diff_add_cancel diff_diff_eq2 real_vector.scale_right_diff_distrib scaleR_left.diff)
qed
lemma collinear_4_3:
assumes "a \<noteq> b"
shows "collinear {a,b,c,d} \<longleftrightarrow> collinear{a,b,c} \<and> collinear{a,b,d}"
using collinear_triples [OF assms, of "{c,d}"] by (force simp:)
lemma collinear_3_trans:
assumes "collinear{a,b,c}" "collinear{b,c,d}" "b \<noteq> c"
shows "collinear{a,b,d}"
proof -
have "collinear{b,c,a,d}"
by (metis (full_types) assms collinear_4_3 insert_commute)
then show ?thesis
by (simp add: collinear_subset)
qed
lemma affine_hull_2_alt:
fixes a b :: "'a::real_vector"
shows "affine hull {a,b} = range (\<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>\<not> 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\<^marker>\<open>tag unimportant\<close>\<open>Basic lemmas about hyperplanes and halfspaces\<close>
lemma halfspace_Int_eq:
"{x. a \<bullet> x \<le> b} \<inter> {x. b \<le> a \<bullet> x} = {x. a \<bullet> x = b}"
"{x. b \<le> a \<bullet> x} \<inter> {x. a \<bullet> x \<le> b} = {x. a \<bullet> x = b}"
by auto
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 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\<^marker>\<open>tag unimportant\<close>\<open>Use set distance for an easy proof of separation properties\<close>
proposition\<^marker>\<open>tag unimportant\<close> 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)
show "open {x. f x < 0}"
by (simp add: open_Collect_less contf)
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 (top_of_set U) S"
and UT: "closedin (top_of_set U) T"
and "S \<inter> T = {}"
obtains S' T' where "openin (top_of_set U) S'"
"openin (top_of_set U) T'"
"S \<subseteq> S'" "T \<subseteq> T'" "S' \<inter> T' = {}"
proof (cases "S = {} \<or> T = {}")
case True with that show ?thesis
using UT US by (blast dest: closedin_subset)
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' = "(U \<inter> f -` {0<..})" and T' = "(U \<inter> f -` {..<0})" in that)
show "(U \<inter> f -` {0<..}) \<inter> (U \<inter> f -` {..<0}) = {}"
by auto
show "openin (top_of_set U) (U \<inter> f -` {0<..})"
by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf)
next
show "openin (top_of_set U) (U \<inter> f -` {..<0})"
by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf)
next
have "S \<subseteq> U" "T \<subseteq> U"
using closedin_imp_subset assms by blast+
then show "S \<subseteq> U \<inter> f -` {0<..}" "T \<subseteq> U \<inter> f -` {..<0}"
using assms False by (force simp add: f_def setdist_sing_in_set intro!: setdist_gt_0_closedin)+
qed
qed
lemma separation_normal_compact:
fixes S :: "'a::euclidean_space set"
assumes "compact S" "closed T" "S \<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>Connectedness of the intersection of a chain\<close>
proposition connected_chain:
fixes \<F> :: "'a :: euclidean_space set set"
assumes cc: "\<And>S. S \<in> \<F> \<Longrightarrow> compact S \<and> connected S"
and linear: "\<And>S T. S \<in> \<F> \<and> T \<in> \<F> \<Longrightarrow> S \<subseteq> T \<or> T \<subseteq> S"
shows "connected(\<Inter>\<F>)"
proof (cases "\<F> = {}")
case True then show ?thesis
by auto
next
case False
then have cf: "compact(\<Inter>\<F>)"
by (simp add: cc compact_Inter)
have False if AB: "closed A" "closed B" "A \<inter> B = {}"
and ABeq: "A \<union> B = \<Inter>\<F>" and "A \<noteq> {}" "B \<noteq> {}" for A B
proof -
obtain U V where "open U" "open V" "A \<subseteq> U" "B \<subseteq> V" "U \<inter> V = {}"
using separation_normal [OF AB] by metis
obtain K where "K \<in> \<F>" "compact K"
using cc False by blast
then obtain N where "open N" and "K \<subseteq> N"
by blast
let ?\<C> = "insert (U \<union> V) ((\<lambda>S. N - S) ` \<F>)"
obtain \<D> where "\<D> \<subseteq> ?\<C>" "finite \<D>" "K \<subseteq> \<Union>\<D>"
proof (rule compactE [OF \<open>compact K\<close>])
show "K \<subseteq> \<Union>(insert (U \<union> V) ((-) N ` \<F>))"
using \<open>K \<subseteq> N\<close> ABeq \<open>A \<subseteq> U\<close> \<open>B \<subseteq> V\<close> by auto
show "\<And>B. B \<in> insert (U \<union> V) ((-) N ` \<F>) \<Longrightarrow> open B"
by (auto simp: \<open>open U\<close> \<open>open V\<close> open_Un \<open>open N\<close> cc compact_imp_closed open_Diff)
qed
then have "finite(\<D> - {U \<union> V})"
by blast
moreover have "\<D> - {U \<union> V} \<subseteq> (\<lambda>S. N - S) ` \<F>"
using \<open>\<D> \<subseteq> ?\<C>\<close> by blast
ultimately obtain \<G> where "\<G> \<subseteq> \<F>" "finite \<G>" and Deq: "\<D> - {U \<union> V} = (\<lambda>S. N-S) ` \<G>"
using finite_subset_image by metis
obtain J where "J \<in> \<F>" and J: "(\<Union>S\<in>\<G>. N - S) \<subseteq> N - J"
proof (cases "\<G> = {}")
case True
with \<open>\<F> \<noteq> {}\<close> that show ?thesis
by auto
next
case False
have "\<And>S T. \<lbrakk>S \<in> \<G>; T \<in> \<G>\<rbrakk> \<Longrightarrow> S \<subseteq> T \<or> T \<subseteq> S"
by (meson \<open>\<G> \<subseteq> \<F>\<close> in_mono local.linear)
with \<open>finite \<G>\<close> \<open>\<G> \<noteq> {}\<close>
have "\<exists>J \<in> \<G>. (\<Union>S\<in>\<G>. N - S) \<subseteq> N - J"
proof induction
case (insert X \<H>)
show ?case
proof (cases "\<H> = {}")
case True then show ?thesis by auto
next
case False
then have "\<And>S T. \<lbrakk>S \<in> \<H>; T \<in> \<H>\<rbrakk> \<Longrightarrow> S \<subseteq> T \<or> T \<subseteq> S"
by (simp add: insert.prems)
with insert.IH False obtain J where "J \<in> \<H>" and J: "(\<Union>Y\<in>\<H>. N - Y) \<subseteq> N - J"
by metis
have "N - J \<subseteq> N - X \<or> N - X \<subseteq> N - J"
by (meson Diff_mono \<open>J \<in> \<H>\<close> insert.prems(2) insert_iff order_refl)
then show ?thesis
proof
assume "N - J \<subseteq> N - X" with J show ?thesis
by auto
next
assume "N - X \<subseteq> N - J"
with J have "N - X \<union> \<Union> ((-) N ` \<H>) \<subseteq> N - J"
by auto
with \<open>J \<in> \<H>\<close> show ?thesis
by blast
qed
qed
qed simp
with \<open>\<G> \<subseteq> \<F>\<close> show ?thesis by (blast intro: that)
qed
have "K \<subseteq> \<Union>(insert (U \<union> V) (\<D> - {U \<union> V}))"
using \<open>K \<subseteq> \<Union>\<D>\<close> by auto
also have "... \<subseteq> (U \<union> V) \<union> (N - J)"
by (metis (no_types, hide_lams) Deq Un_subset_iff Un_upper2 J Union_insert order_trans sup_ge1)
finally have "J \<inter> K \<subseteq> U \<union> V"
by blast
moreover have "connected(J \<inter> K)"
by (metis Int_absorb1 \<open>J \<in> \<F>\<close> \<open>K \<in> \<F>\<close> cc inf.orderE local.linear)
moreover have "U \<inter> (J \<inter> K) \<noteq> {}"
using ABeq \<open>J \<in> \<F>\<close> \<open>K \<in> \<F>\<close> \<open>A \<noteq> {}\<close> \<open>A \<subseteq> U\<close> by blast
moreover have "V \<inter> (J \<inter> K) \<noteq> {}"
using ABeq \<open>J \<in> \<F>\<close> \<open>K \<in> \<F>\<close> \<open>B \<noteq> {}\<close> \<open>B \<subseteq> V\<close> by blast
ultimately show False
using connectedD [of "J \<inter> K" U V] \<open>open U\<close> \<open>open V\<close> \<open>U \<inter> V = {}\<close> by auto
qed
with cf show ?thesis
by (auto simp: connected_closed_set compact_imp_closed)
qed
lemma connected_chain_gen:
fixes \<F> :: "'a :: euclidean_space set set"
assumes X: "X \<in> \<F>" "compact X"
and cc: "\<And>T. T \<in> \<F> \<Longrightarrow> closed T \<and> connected T"
and linear: "\<And>S T. S \<in> \<F> \<and> T \<in> \<F> \<Longrightarrow> S \<subseteq> T \<or> T \<subseteq> S"
shows "connected(\<Inter>\<F>)"
proof -
have "\<Inter>\<F> = (\<Inter>T\<in>\<F>. X \<inter> T)"
using X by blast
moreover have "connected (\<Inter>T\<in>\<F>. X \<inter> T)"
proof (rule connected_chain)
show "\<And>T. T \<in> (\<inter>) X ` \<F> \<Longrightarrow> compact T \<and> connected T"
using cc X by auto (metis inf.absorb2 inf.orderE local.linear)
show "\<And>S T. S \<in> (\<inter>) X ` \<F> \<and> T \<in> (\<inter>) X ` \<F> \<Longrightarrow> S \<subseteq> T \<or> T \<subseteq> S"
using local.linear by blast
qed
ultimately show ?thesis
by metis
qed
lemma connected_nest:
fixes S :: "'a::linorder \<Rightarrow> 'b::euclidean_space set"
assumes S: "\<And>n. compact(S n)" "\<And>n. connected(S n)"
and nest: "\<And>m n. m \<le> n \<Longrightarrow> S n \<subseteq> S m"
shows "connected(\<Inter> (range S))"
apply (rule connected_chain)
using S apply blast
by (metis image_iff le_cases nest)
lemma connected_nest_gen:
fixes S :: "'a::linorder \<Rightarrow> 'b::euclidean_space set"
assumes S: "\<And>n. closed(S n)" "\<And>n. connected(S n)" "compact(S k)"
and nest: "\<And>m n. m \<le> n \<Longrightarrow> S n \<subseteq> S m"
shows "connected(\<Inter> (range S))"
apply (rule connected_chain_gen [of "S k"])
using S apply auto
by (meson le_cases nest subsetCE)
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 (top_of_set S) K"
and com: "\<And>U. \<lbrakk>U \<subseteq> T; compact U\<rbrakk> \<Longrightarrow> compact (S \<inter> f -` U)"
and "f ` S \<subseteq> T"
shows "closedin (top_of_set 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> (S \<inter> f -` 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 *: "(S \<inter> f -` 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 (f ` S \<inter> g -` C)" if C: "C \<subseteq> S" "compact C" for C
proof -
obtain h where "h C \<in> C \<and> h C \<notin> S \<or> compact (f ` C)"
by (force simp: C RHS)
moreover have "f ` C = (f ` S \<inter> g -` C)"
using C gf by auto
ultimately show ?thesis
using C by auto
qed
show ?lhs
using proper_map [OF _ _ gfim] **
by (simp add: continuous_on_closed * closedin_imp_subset)
qed
subsection\<^marker>\<open>tag unimportant\<close>\<open>Trivial fact: convexity equals connectedness for collinear sets\<close>
lemma convex_connected_collinear:
fixes S :: "'a::euclidean_space set"
assumes "collinear S"
shows "convex S \<longleftrightarrow> connected S"
proof
assume "convex S"
then show "connected S"
using convex_connected by blast
next
assume S: "connected S"
show "convex S"
proof (cases "S = {}")
case True
then show ?thesis by simp
next
case False
then obtain a where "a \<in> S" by auto
have "collinear (affine hull S)"
by (simp add: assms collinear_affine_hull_collinear)
then obtain z where "z \<noteq> 0" "\<And>x. x \<in> affine hull S \<Longrightarrow> \<exists>c. x - a = c *\<^sub>R z"
by (meson \<open>a \<in> S\<close> collinear hull_inc)
then obtain f where f: "\<And>x. x \<in> affine hull S \<Longrightarrow> x - a = f x *\<^sub>R z"
by metis
then have inj_f: "inj_on f (affine hull S)"
by (metis diff_add_cancel inj_onI)
have diff: "x - y = (f x - f y) *\<^sub>R z" if x: "x \<in> affine hull S" and y: "y \<in> affine hull S" for x y
proof -
have "f x *\<^sub>R z = x - a"
by (simp add: f hull_inc x)
moreover have "f y *\<^sub>R z = y - a"
by (simp add: f hull_inc y)
ultimately show ?thesis
by (simp add: scaleR_left.diff)
qed
have cont_f: "continuous_on (affine hull S) f"
apply (clarsimp simp: dist_norm continuous_on_iff diff)
by (metis \<open>z \<noteq> 0\<close> mult.commute mult_less_cancel_left_pos norm_minus_commute real_norm_def zero_less_mult_iff zero_less_norm_iff)
then have conn_fS: "connected (f ` S)"
by (meson S connected_continuous_image continuous_on_subset hull_subset)
show ?thesis
proof (clarsimp simp: convex_contains_segment)
fix x y z
assume "x \<in> S" "y \<in> S" "z \<in> closed_segment x y"
have False if "z \<notin> S"
proof -
have "f ` (closed_segment x y) = closed_segment (f x) (f y)"
apply (rule continuous_injective_image_segment_1)
apply (meson \<open>x \<in> S\<close> \<open>y \<in> S\<close> convex_affine_hull convex_contains_segment hull_inc continuous_on_subset [OF cont_f])
by (meson \<open>x \<in> S\<close> \<open>y \<in> S\<close> convex_affine_hull convex_contains_segment hull_inc inj_on_subset [OF inj_f])
then have fz: "f z \<in> closed_segment (f x) (f y)"
using \<open>z \<in> closed_segment x y\<close> by blast
have "z \<in> affine hull S"
by (meson \<open>x \<in> S\<close> \<open>y \<in> S\<close> \<open>z \<in> closed_segment x y\<close> convex_affine_hull convex_contains_segment hull_inc subset_eq)
then have fz_notin: "f z \<notin> f ` S"
using hull_subset inj_f inj_onD that by fastforce
moreover have "{..<f z} \<inter> f ` S \<noteq> {}" "{f z<..} \<inter> f ` S \<noteq> {}"
proof -
have "{..<f z} \<inter> f ` {x,y} \<noteq> {}" "{f z<..} \<inter> f ` {x,y} \<noteq> {}"
using fz fz_notin \<open>x \<in> S\<close> \<open>y \<in> S\<close>
apply (auto simp: closed_segment_eq_real_ivl split: if_split_asm)
apply (metis image_eqI less_eq_real_def)+
done
then show "{..<f z} \<inter> f ` S \<noteq> {}" "{f z<..} \<inter> f ` S \<noteq> {}"
using \<open>x \<in> S\<close> \<open>y \<in> S\<close> by blast+
qed
ultimately show False
using connectedD [OF conn_fS, of "{..<f z}" "{f z<..}"] by force
qed
then show "z \<in> S" by meson
qed
qed
qed
lemma compact_convex_collinear_segment_alt:
fixes S :: "'a::euclidean_space set"
assumes "S \<noteq> {}" "compact S" "connected S" "collinear S"
obtains a b where "S = closed_segment a b"
proof -
obtain \<xi> where "\<xi> \<in> S" using \<open>S \<noteq> {}\<close> by auto
have "collinear (affine hull S)"
by (simp add: assms collinear_affine_hull_collinear)
then obtain z where "z \<noteq> 0" "\<And>x. x \<in> affine hull S \<Longrightarrow> \<exists>c. x - \<xi> = c *\<^sub>R z"
by (meson \<open>\<xi> \<in> S\<close> collinear hull_inc)
then obtain f where f: "\<And>x. x \<in> affine hull S \<Longrightarrow> x - \<xi> = f x *\<^sub>R z"
by metis
let ?g = "\<lambda>r. r *\<^sub>R z + \<xi>"
have gf: "?g (f x) = x" if "x \<in> affine hull S" for x
by (metis diff_add_cancel f that)
then have inj_f: "inj_on f (affine hull S)"
by (metis inj_onI)
have diff: "x - y = (f x - f y) *\<^sub>R z" if x: "x \<in> affine hull S" and y: "y \<in> affine hull S" for x y
proof -
have "f x *\<^sub>R z = x - \<xi>"
by (simp add: f hull_inc x)
moreover have "f y *\<^sub>R z = y - \<xi>"
by (simp add: f hull_inc y)
ultimately show ?thesis
by (simp add: scaleR_left.diff)
qed
have cont_f: "continuous_on (affine hull S) f"
apply (clarsimp simp: dist_norm continuous_on_iff diff)
by (metis \<open>z \<noteq> 0\<close> mult.commute mult_less_cancel_left_pos norm_minus_commute real_norm_def zero_less_mult_iff zero_less_norm_iff)
then have "connected (f ` S)"
by (meson \<open>connected S\<close> connected_continuous_image continuous_on_subset hull_subset)
moreover have "compact (f ` S)"
by (meson \<open>compact S\<close> compact_continuous_image_eq cont_f hull_subset inj_f)
ultimately obtain x y where "f ` S = {x..y}"
by (meson connected_compact_interval_1)
then have fS_eq: "f ` S = closed_segment x y"
using \<open>S \<noteq> {}\<close> closed_segment_eq_real_ivl by auto
obtain a b where "a \<in> S" "f a = x" "b \<in> S" "f b = y"
by (metis (full_types) ends_in_segment fS_eq imageE)
have "f ` (closed_segment a b) = closed_segment (f a) (f b)"
apply (rule continuous_injective_image_segment_1)
apply (meson \<open>a \<in> S\<close> \<open>b \<in> S\<close> convex_affine_hull convex_contains_segment hull_inc continuous_on_subset [OF cont_f])
by (meson \<open>a \<in> S\<close> \<open>b \<in> S\<close> convex_affine_hull convex_contains_segment hull_inc inj_on_subset [OF inj_f])
then have "f ` (closed_segment a b) = f ` S"
by (simp add: \<open>f a = x\<close> \<open>f b = y\<close> fS_eq)
then have "?g ` f ` (closed_segment a b) = ?g ` f ` S"
by simp
moreover have "(\<lambda>x. f x *\<^sub>R z + \<xi>) ` closed_segment a b = closed_segment a b"
apply safe
apply (metis (mono_tags, hide_lams) \<open>a \<in> S\<close> \<open>b \<in> S\<close> convex_affine_hull convex_contains_segment gf hull_inc subsetCE)
by (metis (mono_tags, lifting) \<open>a \<in> S\<close> \<open>b \<in> S\<close> convex_affine_hull convex_contains_segment gf hull_subset image_iff subsetCE)
ultimately have "closed_segment a b = S"
using gf by (simp add: image_comp o_def hull_inc cong: image_cong)
then show ?thesis
using that by blast
qed
lemma compact_convex_collinear_segment:
fixes S :: "'a::euclidean_space set"
assumes "S \<noteq> {}" "compact S" "convex S" "collinear S"
obtains a b where "S = closed_segment a b"
using assms convex_connected_collinear compact_convex_collinear_segment_alt by blast
lemma proper_map_from_compact:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes contf: "continuous_on S f" and imf: "f ` S \<subseteq> T" and "compact S"
"closedin (top_of_set T) K"
shows "compact (S \<inter> f -` 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 (S \<times> T \<inter> fst -` K)"
proof -
have "(S \<times> T \<inter> fst -` 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 (top_of_set (S \<times> T)) c"
shows "closedin (top_of_set 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 (S \<times> T \<inter> snd -` K)"
proof -
have "(S \<times> T \<inter> snd -` 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 (top_of_set (S \<times> T)) c"
shows "closedin (top_of_set 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 (top_of_set (S \<times> T)) U"
shows "closedin (top_of_set 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 (top_of_set T) (snd ` U)"
by (rule closed_map_snd [OF assms])
ultimately show ?thesis
by simp
qed
lemma closed_compact_projection:
fixes S :: "'a::euclidean_space set"
and T :: "('a * 'b::euclidean_space) set"
assumes "compact S" and clo: "closed T"
shows "closed {y. \<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\<^marker>\<open>tag unimportant\<close>\<open>Representing affine hull as a finite intersection of hyperplanes\<close>
proposition\<^marker>\<open>tag unimportant\<close> 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> (+) 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) (simp add: 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_base)
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 (top_of_set (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_openin:
fixes S :: "'a::real_normed_vector set"
assumes "openin (top_of_set (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_Diff:
fixes S:: "'a::real_normed_vector set"
assumes ope: "openin (top_of_set (affine hull S)) S" and "finite F" "F \<subset> S"
shows "affine hull (S - F) = affine hull S"
proof -
have clo: "closedin (top_of_set S) F"
using assms finite_imp_closedin by auto
moreover have "S - F \<noteq> {}"
using assms by auto
ultimately show ?thesis
by (metis ope closedin_def topspace_euclidean_subtopology affine_hull_openin openin_trans)
qed
lemma affine_hull_halfspace_lt:
fixes a :: "'a::euclidean_space"
shows "affine hull {x. a \<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)
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)
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
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
using span_zero [of S]
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 (auto simp add: \<open>c = 0\<close>)
done
next
case False
have xc_im: "x \<in> (+) 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> (+) c ` {y. a \<bullet> y = 0}"
by blast
qed
have 2: "span ((\<lambda>x. x - c) ` S) = {x. a \<bullet> x = 0}"
if "(+) 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 "(+) 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 cong: image_cong_simp, 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\<^marker>\<open>tag unimportant\<close>\<open>Some stepping theorems\<close>
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)
by (force simp add: span_zero insert_commute [of a] hull_inc aff_dim_eq_dim [of x] dim_insert
cong: image_cong_simp)
qed
lemma affine_dependent_choose:
fixes a :: "'a :: euclidean_space"
assumes "\<not>(affine_dependent S)"
shows "affine_dependent(insert a S) \<longleftrightarrow> a \<notin> S \<and> a \<in> affine hull S"
(is "?lhs = ?rhs")
proof safe
assume "affine_dependent (insert a S)" and "a \<in> S"
then show "False"
using \<open>a \<in> S\<close> assms insert_absorb by fastforce
next
assume lhs: "affine_dependent (insert a S)"
then have "a \<notin> S"
by (metis (no_types) assms insert_absorb)
moreover have "finite S"
using affine_independent_iff_card assms by blast
moreover have "aff_dim (insert a S) \<noteq> int (card S)"
using \<open>finite S\<close> affine_independent_iff_card \<open>a \<notin> S\<close> lhs by fastforce
ultimately show "a \<in> affine hull S"
by (metis aff_dim_affine_independent aff_dim_insert assms)
next
assume "a \<notin> S" and "a \<in> affine hull S"
show "affine_dependent (insert a S)"
by (simp add: \<open>a \<in> affine hull S\<close> \<open>a \<notin> S\<close> affine_dependent_def)
qed
lemma affine_independent_insert:
fixes a :: "'a :: euclidean_space"
shows "\<lbrakk>\<not> affine_dependent S; a \<notin> affine hull S\<rbrakk> \<Longrightarrow> \<not> affine_dependent(insert a S)"
by (simp add: affine_dependent_choose)
lemma subspace_bounded_eq_trivial:
fixes S :: "'a::real_normed_vector set"
assumes "subspace S"
shows "bounded S \<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)
next
assume "DIM('a) = 1"
then show "bounded {x. a \<bullet> x = 0}"
by (simp add: affine_bounded_eq_lowdim affine_hyperplane assms)
qed
lemma bounded_hyperplane_eq_trivial:
fixes a :: "'a::euclidean_space"
shows "bounded {x. a \<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\<^marker>\<open>tag unimportant\<close>\<open>General case without assuming closure and getting non-strict separation\<close>
proposition\<^marker>\<open>tag unimportant\<close> 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\<^marker>\<open>tag unimportant\<close> 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_scale span_base)
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_scale span_base)
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: field_split_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 ((+) (- z) ` S)"
using \<open>convex S\<close> by simp
moreover have "(+) (- z) ` S \<noteq> {}"
by (simp add: \<open>S \<noteq> {}\<close>)
moreover have "0 \<notin> (+) (- z) ` S"
using zno by auto
ultimately obtain a where "a \<in> span ((+) (- z) ` S)" "a \<noteq> 0"
and a: "\<And>x. x \<in> ((+) (- 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 ((+) (- z) ` S)\<close> affine_hull_insert_span_gen apply blast
apply (simp_all add: \<open>a \<noteq> 0\<close> szx)
done
qed
proposition\<^marker>\<open>tag unimportant\<close> 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)) ((\<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\<^marker>\<open>tag unimportant\<close>\<open> Some results on decomposing convex hulls: intersections, simplicial subdivision\<close>
lemma
fixes s :: "'a::euclidean_space set"
assumes "\<not> 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 "sum 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]: "sum u s = 1"
"sum 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 "sum f (s \<union> t) = 0"
apply (simp add: f_def sum_Un sum_subtractf)
apply (simp add: sum.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 sum_Un scaleR_left_diff_distrib sum_subtractf)
apply (simp add: if_smult sum.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 "sum u (s \<inter> t) = sum u s"
by (simp add: sum.inter_restrict)
then have "sum 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 sum.inter_restrict intro: sum.cong)
ultimately show ?thesis
by force
qed
then show ?A ?C
by (auto simp: convex_hull_finite affine_hull_finite)
qed
proposition\<^marker>\<open>tag unimportant\<close> affine_hull_Int:
fixes s :: "'a::euclidean_space set"
assumes "\<not> 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\<^marker>\<open>tag unimportant\<close> convex_hull_Int:
fixes s :: "'a::euclidean_space set"
assumes "\<not> 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\<^marker>\<open>tag unimportant\<close>
fixes s :: "'a::euclidean_space set set"
assumes "\<not> 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
proposition\<^marker>\<open>tag unimportant\<close> in_convex_hull_exchange_unique:
fixes S :: "'a::euclidean_space set"
assumes naff: "\<not> affine_dependent S" and a: "a \<in> convex hull S"
and S: "T \<subseteq> S" "T' \<subseteq> S"
and x: "x \<in> convex hull (insert a T)"
and x': "x \<in> convex hull (insert a T')"
shows "x \<in> convex hull (insert a (T \<inter> T'))"
proof (cases "a \<in> S")
case True
then have "\<not> affine_dependent (insert a T \<union> insert a T')"
using affine_dependent_subset assms by auto
then have "x \<in> convex hull (insert a T \<inter> insert a T')"
by (metis IntI convex_hull_Int x x')
then show ?thesis
by simp
next
case False
then have anot: "a \<notin> T" "a \<notin> T'"
using assms by auto
have [simp]: "finite S"
by (simp add: aff_independent_finite assms)
then obtain b where b0: "\<And>s. s \<in> S \<Longrightarrow> 0 \<le> b s"
and b1: "sum b S = 1" and aeq: "a = (\<Sum>s\<in>S. b s *\<^sub>R s)"
using a by (auto simp: convex_hull_finite)
have fin [simp]: "finite T" "finite T'"
using assms infinite_super \<open>finite S\<close> by blast+
then obtain c c' where c0: "\<And>t. t \<in> insert a T \<Longrightarrow> 0 \<le> c t"
and c1: "sum c (insert a T) = 1"
and xeq: "x = (\<Sum>t \<in> insert a T. c t *\<^sub>R t)"
and c'0: "\<And>t. t \<in> insert a T' \<Longrightarrow> 0 \<le> c' t"
and c'1: "sum c' (insert a T') = 1"
and x'eq: "x = (\<Sum>t \<in> insert a T'. c' t *\<^sub>R t)"
using x x' by (auto simp: convex_hull_finite)
with fin anot
have sumTT': "sum c T = 1 - c a" "sum c' T' = 1 - c' a"
and wsumT: "(\<Sum>t \<in> T. c t *\<^sub>R t) = x - c a *\<^sub>R a"
by simp_all
have wsumT': "(\<Sum>t \<in> T'. c' t *\<^sub>R t) = x - c' a *\<^sub>R a"
using x'eq fin anot by simp
define cc where "cc \<equiv> \<lambda>x. if x \<in> T then c x else 0"
define cc' where "cc' \<equiv> \<lambda>x. if x \<in> T' then c' x else 0"
define dd where "dd \<equiv> \<lambda>x. cc x - cc' x + (c a - c' a) * b x"
have sumSS': "sum cc S = 1 - c a" "sum cc' S = 1 - c' a"
unfolding cc_def cc'_def using S
by (simp_all add: Int_absorb1 Int_absorb2 sum_subtractf sum.inter_restrict [symmetric] sumTT')
have wsumSS: "(\<Sum>t \<in> S. cc t *\<^sub>R t) = x - c a *\<^sub>R a" "(\<Sum>t \<in> S. cc' t *\<^sub>R t) = x - c' a *\<^sub>R a"
unfolding cc_def cc'_def using S
by (simp_all add: Int_absorb1 Int_absorb2 if_smult sum.inter_restrict [symmetric] wsumT wsumT' cong: if_cong)
have sum_dd0: "sum dd S = 0"
unfolding dd_def using S
by (simp add: sumSS' comm_monoid_add_class.sum.distrib sum_subtractf
algebra_simps sum_distrib_right [symmetric] b1)
have "(\<Sum>v\<in>S. (b v * x) *\<^sub>R v) = x *\<^sub>R (\<Sum>v\<in>S. b v *\<^sub>R v)" for x
by (simp add: pth_5 real_vector.scale_sum_right mult.commute)
then have *: "(\<Sum>v\<in>S. (b v * x) *\<^sub>R v) = x *\<^sub>R a" for x
using aeq by blast
have "(\<Sum>v \<in> S. dd v *\<^sub>R v) = 0"
unfolding dd_def using S
by (simp add: * wsumSS sum.distrib sum_subtractf algebra_simps)
then have dd0: "dd v = 0" if "v \<in> S" for v
using naff that \<open>finite S\<close> sum_dd0 unfolding affine_dependent_explicit
apply (simp only: not_ex)
apply (drule_tac x=S in spec)
apply (drule_tac x=dd in spec, simp)
done
consider "c' a \<le> c a" | "c a \<le> c' a" by linarith
then show ?thesis
proof cases
case 1
then have "sum cc S \<le> sum cc' S"
by (simp add: sumSS')
then have le: "cc x \<le> cc' x" if "x \<in> S" for x
using dd0 [OF that] 1 b0 mult_left_mono that
by (fastforce simp add: dd_def algebra_simps)
have cc0: "cc x = 0" if "x \<in> S" "x \<notin> T \<inter> T'" for x
using le [OF \<open>x \<in> S\<close>] that c0
by (force simp: cc_def cc'_def split: if_split_asm)
show ?thesis
proof (simp add: convex_hull_finite, intro exI conjI)
show "\<forall>x\<in>T \<inter> T'. 0 \<le> (cc(a := c a)) x"
by (simp add: c0 cc_def)
show "0 \<le> (cc(a := c a)) a"
by (simp add: c0)
have "sum (cc(a := c a)) (insert a (T \<inter> T')) = c a + sum (cc(a := c a)) (T \<inter> T')"
by (simp add: anot)
also have "... = c a + sum (cc(a := c a)) S"
apply simp
apply (rule sum.mono_neutral_left)
using \<open>T \<subseteq> S\<close> apply (auto simp: \<open>a \<notin> S\<close> cc0)
done
also have "... = c a + (1 - c a)"
- by (metis \<open>a \<notin> S\<close> fun_upd_other sum.cong sumSS')
+ by (metis \<open>a \<notin> S\<close> fun_upd_other sum.cong sumSS'(1))
finally show "sum (cc(a := c a)) (insert a (T \<inter> T')) = 1"
by simp
have "(\<Sum>x\<in>insert a (T \<inter> T'). (cc(a := c a)) x *\<^sub>R x) = c a *\<^sub>R a + (\<Sum>x \<in> T \<inter> T'. (cc(a := c a)) x *\<^sub>R x)"
by (simp add: anot)
also have "... = c a *\<^sub>R a + (\<Sum>x \<in> S. (cc(a := c a)) x *\<^sub>R x)"
apply simp
apply (rule sum.mono_neutral_left)
using \<open>T \<subseteq> S\<close> apply (auto simp: \<open>a \<notin> S\<close> cc0)
done
also have "... = c a *\<^sub>R a + x - c a *\<^sub>R a"
by (simp add: wsumSS \<open>a \<notin> S\<close> if_smult sum_delta_notmem)
finally show "(\<Sum>x\<in>insert a (T \<inter> T'). (cc(a := c a)) x *\<^sub>R x) = x"
by simp
qed
next
case 2
then have "sum cc' S \<le> sum cc S"
by (simp add: sumSS')
then have le: "cc' x \<le> cc x" if "x \<in> S" for x
using dd0 [OF that] 2 b0 mult_left_mono that
by (fastforce simp add: dd_def algebra_simps)
have cc0: "cc' x = 0" if "x \<in> S" "x \<notin> T \<inter> T'" for x
using le [OF \<open>x \<in> S\<close>] that c'0
by (force simp: cc_def cc'_def split: if_split_asm)
show ?thesis
proof (simp add: convex_hull_finite, intro exI conjI)
show "\<forall>x\<in>T \<inter> T'. 0 \<le> (cc'(a := c' a)) x"
by (simp add: c'0 cc'_def)
show "0 \<le> (cc'(a := c' a)) a"
by (simp add: c'0)
have "sum (cc'(a := c' a)) (insert a (T \<inter> T')) = c' a + sum (cc'(a := c' a)) (T \<inter> T')"
by (simp add: anot)
also have "... = c' a + sum (cc'(a := c' a)) S"
apply simp
apply (rule sum.mono_neutral_left)
using \<open>T \<subseteq> S\<close> apply (auto simp: \<open>a \<notin> S\<close> cc0)
done
also have "... = c' a + (1 - c' a)"
by (metis \<open>a \<notin> S\<close> fun_upd_other sum.cong sumSS')
finally show "sum (cc'(a := c' a)) (insert a (T \<inter> T')) = 1"
by simp
have "(\<Sum>x\<in>insert a (T \<inter> T'). (cc'(a := c' a)) x *\<^sub>R x) = c' a *\<^sub>R a + (\<Sum>x \<in> T \<inter> T'. (cc'(a := c' a)) x *\<^sub>R x)"
by (simp add: anot)
also have "... = c' a *\<^sub>R a + (\<Sum>x \<in> S. (cc'(a := c' a)) x *\<^sub>R x)"
apply simp
apply (rule sum.mono_neutral_left)
using \<open>T \<subseteq> S\<close> apply (auto simp: \<open>a \<notin> S\<close> cc0)
done
also have "... = c a *\<^sub>R a + x - c a *\<^sub>R a"
by (simp add: wsumSS \<open>a \<notin> S\<close> if_smult sum_delta_notmem)
finally show "(\<Sum>x\<in>insert a (T \<inter> T'). (cc'(a := c' a)) x *\<^sub>R x) = x"
by simp
qed
qed
qed
corollary\<^marker>\<open>tag unimportant\<close> convex_hull_exchange_Int:
fixes a :: "'a::euclidean_space"
assumes "\<not> affine_dependent S" "a \<in> convex hull S" "T \<subseteq> S" "T' \<subseteq> S"
shows "(convex hull (insert a T)) \<inter> (convex hull (insert a T')) =
convex hull (insert a (T \<inter> T'))"
apply (rule subset_antisym)
using in_convex_hull_exchange_unique assms apply blast
by (metis hull_mono inf_le1 inf_le2 insert_inter_insert le_inf_iff)
lemma Int_closed_segment:
fixes b :: "'a::euclidean_space"
assumes "b \<in> closed_segment a c \<or> \<not> collinear{a,b,c}"
shows "closed_segment a b \<inter> closed_segment b c = {b}"
proof (cases "c = a")
case True
then show ?thesis
using assms collinear_3_eq_affine_dependent by fastforce
next
case False
from assms show ?thesis
proof
assume "b \<in> closed_segment a c"
moreover have "\<not> affine_dependent {a, c}"
by (simp)
ultimately show ?thesis
using False convex_hull_exchange_Int [of "{a,c}" b "{a}" "{c}"]
by (simp add: segment_convex_hull insert_commute)
next
assume ncoll: "\<not> collinear {a, b, c}"
have False if "closed_segment a b \<inter> closed_segment b c \<noteq> {b}"
proof -
have "b \<in> closed_segment a b" and "b \<in> closed_segment b c"
by auto
with that obtain d where "b \<noteq> d" "d \<in> closed_segment a b" "d \<in> closed_segment b c"
by force
then have d: "collinear {a, d, b}" "collinear {b, d, c}"
by (auto simp: between_mem_segment between_imp_collinear)
have "collinear {a, b, c}"
apply (rule collinear_3_trans [OF _ _ \<open>b \<noteq> d\<close>])
using d by (auto simp: insert_commute)
with ncoll show False ..
qed
then show ?thesis
by blast
qed
qed
lemma affine_hull_finite_intersection_hyperplanes:
fixes s :: "'a::euclidean_space set"
obtains 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 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
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_base)
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_iff subs)
qed
proposition\<^marker>\<open>tag unimportant\<close> 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 ((+) (- c) ` S)" by blast
then have aff: "affine ((+) (- c) ` S)"
by (simp add: subspace_imp_affine)
have 0: "0 \<in> (+) (- 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> (+) (- c) ` S"
using assms by auto
then have adc: "a \<bullet> (d - c) \<noteq> 0"
by (simp add: c inner_diff_right)
let ?U = "(+) (c+c) ` {x + y |x y. x \<in> (+) (- c) ` S \<and> a \<bullet> y = 0}"
have "u + v \<in> (+) (c + c) ` {x + v |x v. x \<in> (+) (- 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 "... = range ((+) (c + c))"
by (simp only: affine_hyperplane_sums_eq_UNIV_0 [OF aff 0 dc adc])
also have "... = UNIV"
by simp
finally show ?thesis .
qed
lemma aff_dim_sums_Int_0:
assumes "affine S"
and "affine T"
and "0 \<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 ((+) (-a) ` S)" "affine ((+) (-a) ` T)"
using affine_translation [symmetric, of "- a"] assms by (simp_all cong: image_cong_simp)
have zero: "0 \<in> ((+) (-a) ` S)" "0 \<in> ((+) (-a) ` T)"
using a assms by auto
have "{x + y |x y. x \<in> (+) (- a) ` S \<and> y \<in> (+) (- a) ` T} =
(+) (- 2 *\<^sub>R a) ` {x + y| x y. x \<in> S \<and> y \<in> T}"
by (force simp: algebra_simps scaleR_2)
moreover have "(+) (- a) ` S \<inter> (+) (- a) ` T = (+) (- a) ` (S \<inter> T)"
by auto
ultimately show ?thesis
using aff_dim_sums_Int_0 [OF aff zero] aff_dim_translation_eq
by (metis (lifting))
qed
lemma aff_dim_affine_Int_hyperplane:
fixes a :: "'a::euclidean_space"
assumes "affine S"
shows "aff_dim(S \<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)
lemma aff_dim_openin:
fixes S :: "'a::euclidean_space set"
assumes ope: "openin (top_of_set T) S" and "affine T" "S \<noteq> {}"
shows "aff_dim S = aff_dim T"
proof -
show ?thesis
proof (rule order_antisym)
show "aff_dim S \<le> aff_dim T"
by (blast intro: aff_dim_subset [OF openin_imp_subset] ope)
next
obtain a where "a \<in> S"
using \<open>S \<noteq> {}\<close> by blast
have "S \<subseteq> T"
using ope openin_imp_subset by auto
then have "a \<in> T"
using \<open>a \<in> S\<close> by auto
then have subT': "subspace ((\<lambda>x. - a + x) ` T)"
using affine_diffs_subspace \<open>affine T\<close> by auto
then obtain B where Bsub: "B \<subseteq> ((\<lambda>x. - a + x) ` T)" and po: "pairwise orthogonal B"
and eq1: "\<And>x. x \<in> B \<Longrightarrow> norm x = 1" and "independent B"
and cardB: "card B = dim ((\<lambda>x. - a + x) ` T)"
and spanB: "span B = ((\<lambda>x. - a + x) ` T)"
by (rule orthonormal_basis_subspace) auto
obtain e where "0 < e" and e: "cball a e \<inter> T \<subseteq> S"
by (meson \<open>a \<in> S\<close> openin_contains_cball ope)
have "aff_dim T = aff_dim ((\<lambda>x. - a + x) ` T)"
by (metis aff_dim_translation_eq)
also have "... = dim ((\<lambda>x. - a + x) ` T)"
using aff_dim_subspace subT' by blast
also have "... = card B"
by (simp add: cardB)
also have "... = card ((\<lambda>x. e *\<^sub>R x) ` B)"
using \<open>0 < e\<close> by (force simp: inj_on_def card_image)
also have "... \<le> dim ((\<lambda>x. - a + x) ` S)"
proof (simp, rule independent_card_le_dim)
have e': "cball 0 e \<inter> (\<lambda>x. x - a) ` T \<subseteq> (\<lambda>x. x - a) ` S"
using e by (auto simp: dist_norm norm_minus_commute subset_eq)
have "(\<lambda>x. e *\<^sub>R x) ` B \<subseteq> cball 0 e \<inter> (\<lambda>x. x - a) ` T"
using Bsub \<open>0 < e\<close> eq1 subT' \<open>a \<in> T\<close> by (auto simp: subspace_def)
then show "(\<lambda>x. e *\<^sub>R x) ` B \<subseteq> (\<lambda>x. x - a) ` S"
using e' by blast
show "independent ((\<lambda>x. e *\<^sub>R x) ` B)"
using linear_scale_self \<open>independent B\<close>
apply (rule linear_independent_injective_image)
using \<open>0 < e\<close> inj_on_def by fastforce
qed
also have "... = aff_dim S"
using \<open>a \<in> S\<close> aff_dim_eq_dim hull_inc by (force cong: image_cong_simp)
finally show "aff_dim T \<le> aff_dim S" .
qed
qed
lemma dim_openin:
fixes S :: "'a::euclidean_space set"
assumes ope: "openin (top_of_set T) S" and "subspace T" "S \<noteq> {}"
shows "dim S = dim T"
proof (rule order_antisym)
show "dim S \<le> dim T"
by (metis ope dim_subset openin_subset topspace_euclidean_subtopology)
next
have "dim T = aff_dim S"
using aff_dim_openin
by (metis aff_dim_subspace \<open>subspace T\<close> \<open>S \<noteq> {}\<close> ope subspace_affine)
also have "... \<le> dim S"
by (metis aff_dim_subset aff_dim_subspace dim_span span_superset
subspace_span)
finally show "dim T \<le> dim S" by simp
qed
subsection\<open>Lower-dimensional affine subsets are nowhere dense\<close>
proposition dense_complement_subspace:
fixes S :: "'a :: euclidean_space set"
assumes dim_less: "dim T < dim S" and "subspace S" shows "closure(S - T) = S"
proof -
have "closure(S - U) = S" if "dim U < dim S" "U \<subseteq> S" for U
proof -
have "span U \<subset> span S"
by (metis neq_iff psubsetI span_eq_dim span_mono that)
then obtain a where "a \<noteq> 0" "a \<in> span S" and a: "\<And>y. y \<in> span U \<Longrightarrow> orthogonal a y"
using orthogonal_to_subspace_exists_gen by metis
show ?thesis
proof
have "closed S"
by (simp add: \<open>subspace S\<close> closed_subspace)
then show "closure (S - U) \<subseteq> S"
by (simp add: closure_minimal)
show "S \<subseteq> closure (S - U)"
proof (clarsimp simp: closure_approachable)
fix x and e::real
assume "x \<in> S" "0 < e"
show "\<exists>y\<in>S - U. dist y x < e"
proof (cases "x \<in> U")
case True
let ?y = "x + (e/2 / norm a) *\<^sub>R a"
show ?thesis
proof
show "dist ?y x < e"
using \<open>0 < e\<close> by (simp add: dist_norm)
next
have "?y \<in> S"
by (metis \<open>a \<in> span S\<close> \<open>x \<in> S\<close> assms(2) span_eq_iff subspace_add subspace_scale)
moreover have "?y \<notin> U"
proof -
have "e/2 / norm a \<noteq> 0"
using \<open>0 < e\<close> \<open>a \<noteq> 0\<close> by auto
then show ?thesis
by (metis True \<open>a \<noteq> 0\<close> a orthogonal_scaleR orthogonal_self real_vector.scale_eq_0_iff span_add_eq span_base)
qed
ultimately show "?y \<in> S - U" by blast
qed
next
case False
with \<open>0 < e\<close> \<open>x \<in> S\<close> show ?thesis by force
qed
qed
qed
qed
moreover have "S - S \<inter> T = S-T"
by blast
moreover have "dim (S \<inter> T) < dim S"
by (metis dim_less dim_subset inf.cobounded2 inf.orderE inf.strict_boundedE not_le)
ultimately show ?thesis
by force
qed
corollary\<^marker>\<open>tag unimportant\<close> dense_complement_affine:
fixes S :: "'a :: euclidean_space set"
assumes less: "aff_dim T < aff_dim S" and "affine S" shows "closure(S - T) = S"
proof (cases "S \<inter> T = {}")
case True
then show ?thesis
by (metis Diff_triv affine_hull_eq \<open>affine S\<close> closure_same_affine_hull closure_subset hull_subset subset_antisym)
next
case False
then obtain z where z: "z \<in> S \<inter> T" by blast
then have "subspace ((+) (- z) ` S)"
by (meson IntD1 affine_diffs_subspace \<open>affine S\<close>)
moreover have "int (dim ((+) (- z) ` T)) < int (dim ((+) (- z) ` S))"
thm aff_dim_eq_dim
using z less by (simp add: aff_dim_eq_dim_subtract [of z] hull_inc cong: image_cong_simp)
ultimately have "closure(((+) (- z) ` S) - ((+) (- z) ` T)) = ((+) (- z) ` S)"
by (simp add: dense_complement_subspace)
then show ?thesis
by (metis closure_translation translation_diff translation_invert)
qed
corollary\<^marker>\<open>tag unimportant\<close> dense_complement_openin_affine_hull:
fixes S :: "'a :: euclidean_space set"
assumes less: "aff_dim T < aff_dim S"
and ope: "openin (top_of_set (affine hull S)) S"
shows "closure(S - T) = closure S"
proof -
have "affine hull S - T \<subseteq> affine hull S"
by blast
then have "closure (S \<inter> closure (affine hull S - T)) = closure (S \<inter> (affine hull S - T))"
by (rule closure_openin_Int_closure [OF ope])
then show ?thesis
by (metis Int_Diff aff_dim_affine_hull affine_affine_hull dense_complement_affine hull_subset inf.orderE less)
qed
corollary\<^marker>\<open>tag unimportant\<close> dense_complement_convex:
fixes S :: "'a :: euclidean_space set"
assumes "aff_dim T < aff_dim S" "convex S"
shows "closure(S - T) = closure S"
proof
show "closure (S - T) \<subseteq> closure S"
by (simp add: closure_mono)
have "closure (rel_interior S - T) = closure (rel_interior S)"
apply (rule dense_complement_openin_affine_hull)
apply (simp add: assms rel_interior_aff_dim)
using \<open>convex S\<close> rel_interior_rel_open rel_open by blast
then show "closure S \<subseteq> closure (S - T)"
by (metis Diff_mono \<open>convex S\<close> closure_mono convex_closure_rel_interior order_refl rel_interior_subset)
qed
corollary\<^marker>\<open>tag unimportant\<close> dense_complement_convex_closed:
fixes S :: "'a :: euclidean_space set"
assumes "aff_dim T < aff_dim S" "convex S" "closed S"
shows "closure(S - T) = S"
by (simp add: assms dense_complement_convex)
subsection\<^marker>\<open>tag unimportant\<close>\<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\<^marker>\<open>tag unimportant\<close> affine_parallel_slice:
fixes S :: "'a :: euclidean_space set"
assumes "affine S"
and "S \<inter> {x. a \<bullet> x \<le> b} \<noteq> {}"
and "\<not> (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 ((+) (- z) ` S)" by blast
then have aff: "affine ((+) (- z) ` S)" and span: "span ((+) (- z) ` S) = ((+) (- z) ` S)"
by (auto simp: subspace_imp_affine)
obtain a' a'' where a': "a' \<in> span ((+) (- z) ` S)" and a: "a = a' + a''"
and "\<And>w. w \<in> span ((+) (- z) ` S) \<Longrightarrow> orthogonal a'' w"
using orthogonal_subspace_decomp_exists [of "(+) (- z) ` S" "a"] by metis
then have "\<And>w. w \<in> S \<Longrightarrow> a'' \<bullet> (w-z) = 0"
by (simp add: span_base orthogonal_def)
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> (+) (- z) ` S \<Longrightarrow> (w + a') \<in> (+) (- z) ` S"
by (metis subspace_add a' span_eq_iff 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 cong: image_cong_simp)
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 assms(1) assms(2) linear_injective_left_inverse 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
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 ((+) (- a) ` S)"
by (simp add: affine_diffs_subspace_subtract cong: image_cong_simp)
have "nat d \<le> dim ((+) (- a) ` S)"
by (metis aff_dim_subspace aff_dim_translation_eq dle nat_int nat_mono ss)
then obtain T where "subspace T" and Tsb: "T \<subseteq> span ((+) (- a) ` S)"
and Tdim: "dim T = nat d"
using choose_subspace_of_subspace [of "nat d" "(+) (- a) ` S"] by blast
then have "affine T"
using subspace_affine by blast
then have "affine ((+) a ` T)"
by (metis affine_hull_eq affine_hull_translation)
moreover have "(+) a ` T \<subseteq> S"
proof -
have "T \<subseteq> (+) (- a) ` S"
by (metis (no_types) span_eq_iff Tsb ss)
then show "(+) a ` T \<subseteq> S"
using add_ac by auto
qed
moreover have "aff_dim ((+) 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>Paracompactness\<close>
proposition paracompact:
fixes S :: "'a :: {metric_space,second_countable_topology} 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)
using closed_cball closure_minimal by blast
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>(G ` S)"
using Lindelof [of "G ` S"] by (metis image_iff)
then obtain K where K: "K \<subseteq> S" "countable K" and eq: "\<Union>(G ` K) = \<Union>(G ` S)"
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 :: {metric_space,second_countable_topology} set"
assumes cin: "closedin (top_of_set U) S"
and oin: "\<And>T. T \<in> \<C> \<Longrightarrow> openin (top_of_set U) T"
and "S \<subseteq> \<Union>\<C>"
obtains \<C>' where "S \<subseteq> \<Union> \<C>'"
and "\<And>V. V \<in> \<C>' \<Longrightarrow> openin (top_of_set U) V \<and> (\<exists>T. T \<in> \<C> \<and> V \<subseteq> T)"
and "\<And>x. x \<in> U
\<Longrightarrow> \<exists>V. openin (top_of_set 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> {}}"
by (blast intro: paracompact [OF 1 2])
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 (top_of_set 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 (top_of_set 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\<^marker>\<open>tag unimportant\<close> paracompact_closed:
fixes S :: "'a :: {metric_space,second_countable_topology} 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> {})}"
by (rule paracompact_closedin [of UNIV S \<C>]) (auto simp: assms)
subsection\<^marker>\<open>tag unimportant\<close>\<open>Closed-graph characterization of continuity\<close>
lemma continuous_closed_graph_gen:
fixes T :: "'b::real_normed_vector set"
assumes contf: "continuous_on S f" and fim: "f ` S \<subseteq> T"
shows "closedin (top_of_set (S \<times> T)) ((\<lambda>x. Pair x (f x)) ` S)"
proof -
have eq: "((\<lambda>x. Pair x (f x)) ` S) =(S \<times> T \<inter> (\<lambda>z. (f \<circ> fst)z - snd z) -` {0})"
using fim by auto
show ?thesis
apply (subst eq)
apply (intro continuous_intros continuous_closedin_preimage continuous_on_subset [OF contf])
by auto
qed
lemma continuous_closed_graph_eq:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "compact T" and fim: "f ` S \<subseteq> T"
shows "continuous_on S f \<longleftrightarrow>
closedin (top_of_set (S \<times> T)) ((\<lambda>x. Pair x (f x)) ` S)"
(is "?lhs = ?rhs")
proof -
have "?lhs" if ?rhs
proof (clarsimp simp add: continuous_on_closed_gen [OF fim])
fix U
assume U: "closedin (top_of_set T) U"
have eq: "(S \<inter> f -` U) = fst ` (((\<lambda>x. Pair x (f x)) ` S) \<inter> (S \<times> U))"
by (force simp: image_iff)
show "closedin (top_of_set S) (S \<inter> f -` U)"
by (simp add: U closedin_Int closedin_Times closed_map_fst [OF \<open>compact T\<close>] that eq)
qed
with continuous_closed_graph_gen assms show ?thesis by blast
qed
lemma continuous_closed_graph:
fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
assumes "closed S" and contf: "continuous_on S f"
shows "closed ((\<lambda>x. Pair x (f x)) ` S)"
apply (rule closedin_closed_trans)
apply (rule continuous_closed_graph_gen [OF contf subset_UNIV])
by (simp add: \<open>closed S\<close> closed_Times)
lemma continuous_from_closed_graph:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "compact T" and fim: "f ` S \<subseteq> T" and clo: "closed ((\<lambda>x. Pair x (f x)) ` S)"
shows "continuous_on S f"
using fim clo
by (auto intro: closed_subset simp: continuous_closed_graph_eq [OF \<open>compact T\<close> fim])
lemma continuous_on_Un_local_open:
assumes opS: "openin (top_of_set (S \<union> T)) S"
and opT: "openin (top_of_set (S \<union> T)) T"
and contf: "continuous_on S f" and contg: "continuous_on T f"
shows "continuous_on (S \<union> T) f"
using pasting_lemma [of "{S,T}" "top_of_set (S \<union> T)" id euclidean "\<lambda>i. f" f] contf contg opS opT
by (simp add: subtopology_subtopology) (metis inf.absorb2 openin_imp_subset)
lemma continuous_on_cases_local_open:
assumes opS: "openin (top_of_set (S \<union> T)) S"
and opT: "openin (top_of_set (S \<union> T)) T"
and contf: "continuous_on S f" and contg: "continuous_on T g"
and fg: "\<And>x. x \<in> S \<and> \<not>P x \<or> x \<in> T \<and> P x \<Longrightarrow> f x = g x"
shows "continuous_on (S \<union> T) (\<lambda>x. if P x then f x else g x)"
proof -
have "\<And>x. x \<in> S \<Longrightarrow> (if P x then f x else g x) = f x" "\<And>x. x \<in> T \<Longrightarrow> (if P x then f x else g x) = g x"
by (simp_all add: fg)
then have "continuous_on S (\<lambda>x. if P x then f x else g x)" "continuous_on T (\<lambda>x. if P x then f x else g x)"
by (simp_all add: contf contg cong: continuous_on_cong)
then show ?thesis
by (rule continuous_on_Un_local_open [OF opS opT])
qed
subsection\<^marker>\<open>tag unimportant\<close>\<open>The union of two collinear segments is another segment\<close>
proposition\<^marker>\<open>tag unimportant\<close> in_convex_hull_exchange:
fixes a :: "'a::euclidean_space"
assumes a: "a \<in> convex hull S" and xS: "x \<in> convex hull S"
obtains b where "b \<in> S" "x \<in> convex hull (insert a (S - {b}))"
proof (cases "a \<in> S")
case True
with xS insert_Diff that show ?thesis by fastforce
next
case False
show ?thesis
proof (cases "finite S \<and> card S \<le> Suc (DIM('a))")
case True
then obtain u where u0: "\<And>i. i \<in> S \<Longrightarrow> 0 \<le> u i" and u1: "sum u S = 1"
and ua: "(\<Sum>i\<in>S. u i *\<^sub>R i) = a"
using a by (auto simp: convex_hull_finite)
obtain v where v0: "\<And>i. i \<in> S \<Longrightarrow> 0 \<le> v i" and v1: "sum v S = 1"
and vx: "(\<Sum>i\<in>S. v i *\<^sub>R i) = x"
using True xS by (auto simp: convex_hull_finite)
show ?thesis
proof (cases "\<exists>b. b \<in> S \<and> v b = 0")
case True
then obtain b where b: "b \<in> S" "v b = 0"
by blast
show ?thesis
proof
have fin: "finite (insert a (S - {b}))"
using sum.infinite v1 by fastforce
show "x \<in> convex hull insert a (S - {b})"
unfolding convex_hull_finite [OF fin] mem_Collect_eq
proof (intro conjI exI ballI)
have "(\<Sum>x \<in> insert a (S - {b}). if x = a then 0 else v x) =
(\<Sum>x \<in> S - {b}. if x = a then 0 else v x)"
apply (rule sum.mono_neutral_right)
using fin by auto
also have "... = (\<Sum>x \<in> S - {b}. v x)"
using b False by (auto intro!: sum.cong split: if_split_asm)
also have "... = (\<Sum>x\<in>S. v x)"
by (metis \<open>v b = 0\<close> diff_zero sum.infinite sum_diff1 u1 zero_neq_one)
finally show "(\<Sum>x\<in>insert a (S - {b}). if x = a then 0 else v x) = 1"
by (simp add: v1)
show "\<And>x. x \<in> insert a (S - {b}) \<Longrightarrow> 0 \<le> (if x = a then 0 else v x)"
by (auto simp: v0)
have "(\<Sum>x \<in> insert a (S - {b}). (if x = a then 0 else v x) *\<^sub>R x) =
(\<Sum>x \<in> S - {b}. (if x = a then 0 else v x) *\<^sub>R x)"
apply (rule sum.mono_neutral_right)
using fin by auto
also have "... = (\<Sum>x \<in> S - {b}. v x *\<^sub>R x)"
using b False by (auto intro!: sum.cong split: if_split_asm)
also have "... = (\<Sum>x\<in>S. v x *\<^sub>R x)"
by (metis (no_types, lifting) b(2) diff_zero fin finite.emptyI finite_Diff2 finite_insert scale_eq_0_iff sum_diff1)
finally show "(\<Sum>x\<in>insert a (S - {b}). (if x = a then 0 else v x) *\<^sub>R x) = x"
by (simp add: vx)
qed
qed (rule \<open>b \<in> S\<close>)
next
case False
have le_Max: "u i / v i \<le> Max ((\<lambda>i. u i / v i) ` S)" if "i \<in> S" for i
by (simp add: True that)
have "Max ((\<lambda>i. u i / v i) ` S) \<in> (\<lambda>i. u i / v i) ` S"
using True v1 by (auto intro: Max_in)
then obtain b where "b \<in> S" and beq: "Max ((\<lambda>b. u b / v b) ` S) = u b / v b"
by blast
then have "0 \<noteq> u b / v b"
using le_Max beq divide_le_0_iff le_numeral_extra(2) sum_nonpos u1
by (metis False eq_iff v0)
then have "0 < u b" "0 < v b"
using False \<open>b \<in> S\<close> u0 v0 by force+
have fin: "finite (insert a (S - {b}))"
using sum.infinite v1 by fastforce
show ?thesis
proof
show "x \<in> convex hull insert a (S - {b})"
unfolding convex_hull_finite [OF fin] mem_Collect_eq
proof (intro conjI exI ballI)
have "(\<Sum>x \<in> insert a (S - {b}). if x=a then v b / u b else v x - (v b / u b) * u x) =
v b / u b + (\<Sum>x \<in> S - {b}. v x - (v b / u b) * u x)"
using \<open>a \<notin> S\<close> \<open>b \<in> S\<close> True apply simp
apply (rule sum.cong, auto)
done
also have "... = v b / u b + (\<Sum>x \<in> S - {b}. v x) - (v b / u b) * (\<Sum>x \<in> S - {b}. u x)"
by (simp add: Groups_Big.sum_subtractf sum_distrib_left)
also have "... = (\<Sum>x\<in>S. v x)"
using \<open>0 < u b\<close> True by (simp add: Groups_Big.sum_diff1 u1 field_simps)
finally show "sum (\<lambda>x. if x=a then v b / u b else v x - (v b / u b) * u x) (insert a (S - {b})) = 1"
by (simp add: v1)
show "0 \<le> (if i = a then v b / u b else v i - v b / u b * u i)"
if "i \<in> insert a (S - {b})" for i
using \<open>0 < u b\<close> \<open>0 < v b\<close> v0 [of i] le_Max [of i] beq that False
by (auto simp: field_simps split: if_split_asm)
have "(\<Sum>x\<in>insert a (S - {b}). (if x=a then v b / u b else v x - v b / u b * u x) *\<^sub>R x) =
(v b / u b) *\<^sub>R a + (\<Sum>x\<in>S - {b}. (v x - v b / u b * u x) *\<^sub>R x)"
using \<open>a \<notin> S\<close> \<open>b \<in> S\<close> True apply simp
apply (rule sum.cong, auto)
done
also have "... = (v b / u b) *\<^sub>R a + (\<Sum>x \<in> S - {b}. v x *\<^sub>R x) - (v b / u b) *\<^sub>R (\<Sum>x \<in> S - {b}. u x *\<^sub>R x)"
by (simp add: Groups_Big.sum_subtractf scaleR_left_diff_distrib sum_distrib_left scale_sum_right)
also have "... = (\<Sum>x\<in>S. v x *\<^sub>R x)"
using \<open>0 < u b\<close> True by (simp add: ua vx Groups_Big.sum_diff1 algebra_simps)
finally
show "(\<Sum>x\<in>insert a (S - {b}). (if x=a then v b / u b else v x - v b / u b * u x) *\<^sub>R x) = x"
by (simp add: vx)
qed
qed (rule \<open>b \<in> S\<close>)
qed
next
case False
obtain T where "finite T" "T \<subseteq> S" and caT: "card T \<le> Suc (DIM('a))" and xT: "x \<in> convex hull T"
using xS by (auto simp: caratheodory [of S])
with False obtain b where b: "b \<in> S" "b \<notin> T"
by (metis antisym subsetI)
show ?thesis
proof
show "x \<in> convex hull insert a (S - {b})"
using \<open>T \<subseteq> S\<close> b by (blast intro: subsetD [OF hull_mono xT])
qed (rule \<open>b \<in> S\<close>)
qed
qed
lemma convex_hull_exchange_Union:
fixes a :: "'a::euclidean_space"
assumes "a \<in> convex hull S"
shows "convex hull S = (\<Union>b \<in> S. convex hull (insert a (S - {b})))" (is "?lhs = ?rhs")
proof
show "?lhs \<subseteq> ?rhs"
by (blast intro: in_convex_hull_exchange [OF assms])
show "?rhs \<subseteq> ?lhs"
proof clarify
fix x b
assume"b \<in> S" "x \<in> convex hull insert a (S - {b})"
then show "x \<in> convex hull S" if "b \<in> S"
by (metis (no_types) that assms order_refl hull_mono hull_redundant insert_Diff_single insert_subset subsetCE)
qed
qed
lemma Un_closed_segment:
fixes a :: "'a::euclidean_space"
assumes "b \<in> closed_segment a c"
shows "closed_segment a b \<union> closed_segment b c = closed_segment a c"
proof (cases "c = a")
case True
with assms show ?thesis by simp
next
case False
with assms have "convex hull {a, b} \<union> convex hull {b, c} = (\<Union>ba\<in>{a, c}. convex hull insert b ({a, c} - {ba}))"
by (auto simp: insert_Diff_if insert_commute)
then show ?thesis
using convex_hull_exchange_Union
by (metis assms segment_convex_hull)
qed
lemma Un_open_segment:
fixes a :: "'a::euclidean_space"
assumes "b \<in> open_segment a c"
shows "open_segment a b \<union> {b} \<union> open_segment b c = open_segment a c"
proof -
have b: "b \<in> closed_segment a c"
by (simp add: assms open_closed_segment)
have *: "open_segment a c \<subseteq> insert b (open_segment a b \<union> open_segment b c)"
if "{b,c,a} \<union> open_segment a b \<union> open_segment b c = {c,a} \<union> open_segment a c"
proof -
have "insert a (insert c (insert b (open_segment a b \<union> open_segment b c))) = insert a (insert c (open_segment a c))"
using that by (simp add: insert_commute)
then show ?thesis
by (metis (no_types) Diff_cancel Diff_eq_empty_iff Diff_insert2 open_segment_def)
qed
show ?thesis
using Un_closed_segment [OF b]
apply (simp add: closed_segment_eq_open)
apply (rule equalityI)
using assms
apply (simp add: b subset_open_segment)
using * by (simp add: insert_commute)
qed
subsection\<open>Covering an open set by a countable chain of compact sets\<close>
proposition open_Union_compact_subsets:
fixes S :: "'a::euclidean_space set"
assumes "open S"
obtains C where "\<And>n. compact(C n)" "\<And>n. C n \<subseteq> S"
"\<And>n. C n \<subseteq> interior(C(Suc n))"
"\<Union>(range C) = S"
"\<And>K. \<lbrakk>compact K; K \<subseteq> S\<rbrakk> \<Longrightarrow> \<exists>N. \<forall>n\<ge>N. K \<subseteq> (C n)"
proof (cases "S = {}")
case True
then show ?thesis
by (rule_tac C = "\<lambda>n. {}" in that) auto
next
case False
then obtain a where "a \<in> S"
by auto
let ?C = "\<lambda>n. cball a (real n) - (\<Union>x \<in> -S. \<Union>e \<in> ball 0 (1 / real(Suc n)). {x + e})"
have "\<exists>N. \<forall>n\<ge>N. K \<subseteq> (f n)"
if "\<And>n. compact(f n)" and sub_int: "\<And>n. f n \<subseteq> interior (f(Suc n))"
and eq: "\<Union>(range f) = S" and "compact K" "K \<subseteq> S" for f K
proof -
have *: "\<forall>n. f n \<subseteq> (\<Union>n. interior (f n))"
by (meson Sup_upper2 UNIV_I \<open>\<And>n. f n \<subseteq> interior (f (Suc n))\<close> image_iff)
have mono: "\<And>m n. m \<le> n \<Longrightarrow>f m \<subseteq> f n"
by (meson dual_order.trans interior_subset lift_Suc_mono_le sub_int)
obtain I where "finite I" and I: "K \<subseteq> (\<Union>i\<in>I. interior (f i))"
proof (rule compactE_image [OF \<open>compact K\<close>])
show "K \<subseteq> (\<Union>n. interior (f n))"
using \<open>K \<subseteq> S\<close> \<open>\<Union>(f ` UNIV) = S\<close> * by blast
qed auto
{ fix n
assume n: "Max I \<le> n"
have "(\<Union>i\<in>I. interior (f i)) \<subseteq> f n"
by (rule UN_least) (meson dual_order.trans interior_subset mono I Max_ge [OF \<open>finite I\<close>] n)
then have "K \<subseteq> f n"
using I by auto
}
then show ?thesis
by blast
qed
moreover have "\<exists>f. (\<forall>n. compact(f n)) \<and> (\<forall>n. (f n) \<subseteq> S) \<and> (\<forall>n. (f n) \<subseteq> interior(f(Suc n))) \<and>
((\<Union>(range f) = S))"
proof (intro exI conjI allI)
show "\<And>n. compact (?C n)"
by (auto simp: compact_diff open_sums)
show "\<And>n. ?C n \<subseteq> S"
by auto
show "?C n \<subseteq> interior (?C (Suc n))" for n
proof (simp add: interior_diff, rule Diff_mono)
show "cball a (real n) \<subseteq> ball a (1 + real n)"
by (simp add: cball_subset_ball_iff)
have cl: "closed (\<Union>x\<in>- S. \<Union>e\<in>cball 0 (1 / (2 + real n)). {x + e})"
using assms by (auto intro: closed_compact_sums)
have "closure (\<Union>x\<in>- S. \<Union>y\<in>ball 0 (1 / (2 + real n)). {x + y})
\<subseteq> (\<Union>x \<in> -S. \<Union>e \<in> cball 0 (1 / (2 + real n)). {x + e})"
by (intro closure_minimal UN_mono ball_subset_cball order_refl cl)
also have "... \<subseteq> (\<Union>x \<in> -S. \<Union>y\<in>ball 0 (1 / (1 + real n)). {x + y})"
apply (intro UN_mono order_refl)
apply (simp add: cball_subset_ball_iff field_split_simps)
done
finally show "closure (\<Union>x\<in>- S. \<Union>y\<in>ball 0 (1 / (2 + real n)). {x + y})
\<subseteq> (\<Union>x \<in> -S. \<Union>y\<in>ball 0 (1 / (1 + real n)). {x + y})" .
qed
have "S \<subseteq> \<Union> (range ?C)"
proof
fix x
assume x: "x \<in> S"
then obtain e where "e > 0" and e: "ball x e \<subseteq> S"
using assms open_contains_ball by blast
then obtain N1 where "N1 > 0" and N1: "real N1 > 1/e"
using reals_Archimedean2
by (metis divide_less_0_iff less_eq_real_def neq0_conv not_le of_nat_0 of_nat_1 of_nat_less_0_iff)
obtain N2 where N2: "norm(x - a) \<le> real N2"
by (meson real_arch_simple)
have N12: "inverse((N1 + N2) + 1) \<le> inverse(N1)"
using \<open>N1 > 0\<close> by (auto simp: field_split_simps)
have "x \<noteq> y + z" if "y \<notin> S" "norm z < 1 / (1 + (real N1 + real N2))" for y z
proof -
have "e * real N1 < e * (1 + (real N1 + real N2))"
by (simp add: \<open>0 < e\<close>)
then have "1 / (1 + (real N1 + real N2)) < e"
using N1 \<open>e > 0\<close>
by (metis divide_less_eq less_trans mult.commute of_nat_add of_nat_less_0_iff of_nat_Suc)
then have "x - z \<in> ball x e"
using that by simp
then have "x - z \<in> S"
using e by blast
with that show ?thesis
by auto
qed
with N2 show "x \<in> \<Union> (range ?C)"
by (rule_tac a = "N1+N2" in UN_I) (auto simp: dist_norm norm_minus_commute)
qed
then show "\<Union> (range ?C) = S" by auto
qed
ultimately show ?thesis
using that by metis
qed
subsection\<open>Orthogonal complement\<close>
definition\<^marker>\<open>tag important\<close> orthogonal_comp ("_\<^sup>\<bottom>" [80] 80)
where "orthogonal_comp W \<equiv> {x. \<forall>y \<in> W. orthogonal y x}"
proposition subspace_orthogonal_comp: "subspace (W\<^sup>\<bottom>)"
unfolding subspace_def orthogonal_comp_def orthogonal_def
by (auto simp: inner_right_distrib)
lemma orthogonal_comp_anti_mono:
assumes "A \<subseteq> B"
shows "B\<^sup>\<bottom> \<subseteq> A\<^sup>\<bottom>"
proof
fix x assume x: "x \<in> B\<^sup>\<bottom>"
show "x \<in> orthogonal_comp A" using x unfolding orthogonal_comp_def
by (simp add: orthogonal_def, metis assms in_mono)
qed
lemma orthogonal_comp_null [simp]: "{0}\<^sup>\<bottom> = UNIV"
by (auto simp: orthogonal_comp_def orthogonal_def)
lemma orthogonal_comp_UNIV [simp]: "UNIV\<^sup>\<bottom> = {0}"
unfolding orthogonal_comp_def orthogonal_def
by auto (use inner_eq_zero_iff in blast)
lemma orthogonal_comp_subset: "U \<subseteq> U\<^sup>\<bottom>\<^sup>\<bottom>"
by (auto simp: orthogonal_comp_def orthogonal_def inner_commute)
lemma subspace_sum_minimal:
assumes "S \<subseteq> U" "T \<subseteq> U" "subspace U"
shows "S + T \<subseteq> U"
proof
fix x
assume "x \<in> S + T"
then obtain xs xt where "xs \<in> S" "xt \<in> T" "x = xs+xt"
by (meson set_plus_elim)
then show "x \<in> U"
by (meson assms subsetCE subspace_add)
qed
proposition subspace_sum_orthogonal_comp:
fixes U :: "'a :: euclidean_space set"
assumes "subspace U"
shows "U + U\<^sup>\<bottom> = UNIV"
proof -
obtain B where "B \<subseteq> U"
and ortho: "pairwise orthogonal B" "\<And>x. x \<in> B \<Longrightarrow> norm x = 1"
and "independent B" "card B = dim U" "span B = U"
using orthonormal_basis_subspace [OF assms] by metis
then have "finite B"
by (simp add: indep_card_eq_dim_span)
have *: "\<forall>x\<in>B. \<forall>y\<in>B. x \<bullet> y = (if x=y then 1 else 0)"
using ortho norm_eq_1 by (auto simp: orthogonal_def pairwise_def)
{ fix v
let ?u = "\<Sum>b\<in>B. (v \<bullet> b) *\<^sub>R b"
have "v = ?u + (v - ?u)"
by simp
moreover have "?u \<in> U"
by (metis (no_types, lifting) \<open>span B = U\<close> assms subspace_sum span_base span_mul)
moreover have "(v - ?u) \<in> U\<^sup>\<bottom>"
proof (clarsimp simp: orthogonal_comp_def orthogonal_def)
fix y
assume "y \<in> U"
with \<open>span B = U\<close> span_finite [OF \<open>finite B\<close>]
obtain u where u: "y = (\<Sum>b\<in>B. u b *\<^sub>R b)"
by auto
have "b \<bullet> (v - ?u) = 0" if "b \<in> B" for b
using that \<open>finite B\<close>
by (simp add: * algebra_simps inner_sum_right if_distrib [of "(*)v" for v] inner_commute cong: if_cong)
then show "y \<bullet> (v - ?u) = 0"
by (simp add: u inner_sum_left)
qed
ultimately have "v \<in> U + U\<^sup>\<bottom>"
using set_plus_intro by fastforce
} then show ?thesis
by auto
qed
lemma orthogonal_Int_0:
assumes "subspace U"
shows "U \<inter> U\<^sup>\<bottom> = {0}"
using orthogonal_comp_def orthogonal_self
by (force simp: assms subspace_0 subspace_orthogonal_comp)
lemma orthogonal_comp_self:
fixes U :: "'a :: euclidean_space set"
assumes "subspace U"
shows "U\<^sup>\<bottom>\<^sup>\<bottom> = U"
proof
have ssU': "subspace (U\<^sup>\<bottom>)"
by (simp add: subspace_orthogonal_comp)
have "u \<in> U" if "u \<in> U\<^sup>\<bottom>\<^sup>\<bottom>" for u
proof -
obtain v w where "u = v+w" "v \<in> U" "w \<in> U\<^sup>\<bottom>"
using subspace_sum_orthogonal_comp [OF assms] set_plus_elim by blast
then have "u-v \<in> U\<^sup>\<bottom>"
by simp
moreover have "v \<in> U\<^sup>\<bottom>\<^sup>\<bottom>"
using \<open>v \<in> U\<close> orthogonal_comp_subset by blast
then have "u-v \<in> U\<^sup>\<bottom>\<^sup>\<bottom>"
by (simp add: subspace_diff subspace_orthogonal_comp that)
ultimately have "u-v = 0"
using orthogonal_Int_0 ssU' by blast
with \<open>v \<in> U\<close> show ?thesis
by auto
qed
then show "U\<^sup>\<bottom>\<^sup>\<bottom> \<subseteq> U"
by auto
qed (use orthogonal_comp_subset in auto)
lemma ker_orthogonal_comp_adjoint:
fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
assumes "linear f"
shows "f -` {0} = (range (adjoint f))\<^sup>\<bottom>"
apply (auto simp: orthogonal_comp_def orthogonal_def)
apply (simp add: adjoint_works assms(1) inner_commute)
by (metis adjoint_works all_zero_iff assms(1) inner_commute)
subsection\<^marker>\<open>tag unimportant\<close> \<open>A non-injective linear function maps into a hyperplane.\<close>
lemma linear_surj_adj_imp_inj:
fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
assumes "linear f" "surj (adjoint f)"
shows "inj f"
proof -
have "\<exists>x. y = adjoint f x" for y
using assms by (simp add: surjD)
then show "inj f"
using assms unfolding inj_on_def image_def
by (metis (no_types) adjoint_works euclidean_eqI)
qed
\<comment> \<open>\<^url>\<open>https://mathonline.wikidot.com/injectivity-and-surjectivity-of-the-adjoint-of-a-linear-map\<close>\<close>
lemma surj_adjoint_iff_inj [simp]:
fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
assumes "linear f"
shows "surj (adjoint f) \<longleftrightarrow> inj f"
proof
assume "surj (adjoint f)"
then show "inj f"
by (simp add: assms linear_surj_adj_imp_inj)
next
assume "inj f"
have "f -` {0} = {0}"
using assms \<open>inj f\<close> linear_0 linear_injective_0 by fastforce
moreover have "f -` {0} = range (adjoint f)\<^sup>\<bottom>"
by (intro ker_orthogonal_comp_adjoint assms)
ultimately have "range (adjoint f)\<^sup>\<bottom>\<^sup>\<bottom> = UNIV"
by (metis orthogonal_comp_null)
then show "surj (adjoint f)"
using adjoint_linear \<open>linear f\<close>
by (subst (asm) orthogonal_comp_self)
(simp add: adjoint_linear linear_subspace_image)
qed
lemma inj_adjoint_iff_surj [simp]:
fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
assumes "linear f"
shows "inj (adjoint f) \<longleftrightarrow> surj f"
proof
assume "inj (adjoint f)"
have "(adjoint f) -` {0} = {0}"
by (metis \<open>inj (adjoint f)\<close> adjoint_linear assms surj_adjoint_iff_inj ker_orthogonal_comp_adjoint orthogonal_comp_UNIV)
then have "(range(f))\<^sup>\<bottom> = {0}"
by (metis (no_types, hide_lams) adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint set_zero)
then show "surj f"
by (metis \<open>inj (adjoint f)\<close> adjoint_adjoint adjoint_linear assms surj_adjoint_iff_inj)
next
assume "surj f"
then have "range f = (adjoint f -` {0})\<^sup>\<bottom>"
by (simp add: adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint)
then have "{0} = adjoint f -` {0}"
using \<open>surj f\<close> adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint by force
then show "inj (adjoint f)"
by (simp add: \<open>surj f\<close> adjoint_adjoint adjoint_linear assms linear_surj_adj_imp_inj)
qed
lemma linear_singular_into_hyperplane:
fixes f :: "'n::euclidean_space \<Rightarrow> 'n"
assumes "linear f"
shows "\<not> inj f \<longleftrightarrow> (\<exists>a. a \<noteq> 0 \<and> (\<forall>x. a \<bullet> f x = 0))" (is "_ = ?rhs")
proof
assume "\<not>inj f"
then show ?rhs
using all_zero_iff
by (metis (no_types, hide_lams) adjoint_clauses(2) adjoint_linear assms
linear_injective_0 linear_injective_imp_surjective linear_surj_adj_imp_inj)
next
assume ?rhs
then show "\<not>inj f"
by (metis assms linear_injective_isomorphism all_zero_iff)
qed
lemma linear_singular_image_hyperplane:
fixes f :: "'n::euclidean_space \<Rightarrow> 'n"
assumes "linear f" "\<not>inj f"
obtains a where "a \<noteq> 0" "\<And>S. f ` S \<subseteq> {x. a \<bullet> x = 0}"
using assms by (fastforce simp add: linear_singular_into_hyperplane)
end
diff --git a/src/HOL/Analysis/Summation_Tests.thy b/src/HOL/Analysis/Summation_Tests.thy
--- a/src/HOL/Analysis/Summation_Tests.thy
+++ b/src/HOL/Analysis/Summation_Tests.thy
@@ -1,901 +1,901 @@
(* Title: HOL/Analysis/Summation_Tests.thy
Author: Manuel Eberl, TU München
*)
section \<open>Radius of Convergence and Summation Tests\<close>
theory Summation_Tests
imports
Complex_Main
"HOL-Library.Discrete"
"HOL-Library.Extended_Real"
"HOL-Library.Liminf_Limsup"
Extended_Real_Limits
begin
text \<open>
The definition of the radius of convergence of a power series,
various summability tests, lemmas to compute the radius of convergence etc.
\<close>
subsection \<open>Convergence tests for infinite sums\<close>
subsubsection \<open>Root test\<close>
lemma limsup_root_powser:
fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
shows "limsup (\<lambda>n. ereal (root n (norm (f n * z ^ n)))) =
limsup (\<lambda>n. ereal (root n (norm (f n)))) * ereal (norm z)"
proof -
have A: "(\<lambda>n. ereal (root n (norm (f n * z ^ n)))) =
(\<lambda>n. ereal (root n (norm (f n))) * ereal (norm z))" (is "?g = ?h")
proof
fix n show "?g n = ?h n"
by (cases "n = 0") (simp_all add: norm_mult real_root_mult real_root_pos2 norm_power)
qed
show ?thesis by (subst A, subst limsup_ereal_mult_right) simp_all
qed
lemma limsup_root_limit:
assumes "(\<lambda>n. ereal (root n (norm (f n)))) \<longlonglongrightarrow> l" (is "?g \<longlonglongrightarrow> _")
shows "limsup (\<lambda>n. ereal (root n (norm (f n)))) = l"
proof -
from assms have "convergent ?g" "lim ?g = l"
unfolding convergent_def by (blast intro: limI)+
with convergent_limsup_cl show ?thesis by force
qed
lemma limsup_root_limit':
assumes "(\<lambda>n. root n (norm (f n))) \<longlonglongrightarrow> l"
shows "limsup (\<lambda>n. ereal (root n (norm (f n)))) = ereal l"
by (intro limsup_root_limit tendsto_ereal assms)
theorem root_test_convergence':
fixes f :: "nat \<Rightarrow> 'a :: banach"
defines "l \<equiv> limsup (\<lambda>n. ereal (root n (norm (f n))))"
assumes l: "l < 1"
shows "summable f"
proof -
have "0 = limsup (\<lambda>n. 0)" by (simp add: Limsup_const)
also have "... \<le> l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero)
finally have "l \<ge> 0" by simp
with l obtain l' where l': "l = ereal l'" by (cases l) simp_all
define c where "c = (1 - l') / 2"
from l and \<open>l \<ge> 0\<close> have c: "l + c > l" "l' + c \<ge> 0" "l' + c < 1" unfolding c_def
by (simp_all add: field_simps l')
have "\<forall>C>l. eventually (\<lambda>n. ereal (root n (norm (f n))) < C) sequentially"
by (subst Limsup_le_iff[symmetric]) (simp add: l_def)
with c have "eventually (\<lambda>n. ereal (root n (norm (f n))) < l + ereal c) sequentially" by simp
with eventually_gt_at_top[of "0::nat"]
have "eventually (\<lambda>n. norm (f n) \<le> (l' + c) ^ n) sequentially"
proof eventually_elim
fix n :: nat assume n: "n > 0"
assume "ereal (root n (norm (f n))) < l + ereal c"
hence "root n (norm (f n)) \<le> l' + c" by (simp add: l')
with c n have "root n (norm (f n)) ^ n \<le> (l' + c) ^ n"
by (intro power_mono) (simp_all add: real_root_ge_zero)
also from n have "root n (norm (f n)) ^ n = norm (f n)" by simp
finally show "norm (f n) \<le> (l' + c) ^ n" by simp
qed
thus ?thesis
by (rule summable_comparison_test_ev[OF _ summable_geometric]) (simp add: c)
qed
theorem root_test_divergence:
fixes f :: "nat \<Rightarrow> 'a :: banach"
defines "l \<equiv> limsup (\<lambda>n. ereal (root n (norm (f n))))"
assumes l: "l > 1"
shows "\<not>summable f"
proof
assume "summable f"
hence bounded: "Bseq f" by (simp add: summable_imp_Bseq)
have "0 = limsup (\<lambda>n. 0)" by (simp add: Limsup_const)
also have "... \<le> l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero)
finally have l_nonneg: "l \<ge> 0" by simp
define c where "c = (if l = \<infinity> then 2 else 1 + (real_of_ereal l - 1) / 2)"
from l l_nonneg consider "l = \<infinity>" | "\<exists>l'. l = ereal l'" by (cases l) simp_all
hence c: "c > 1 \<and> ereal c < l" by cases (insert l, auto simp: c_def field_simps)
have unbounded: "\<not>bdd_above {n. root n (norm (f n)) > c}"
proof
assume "bdd_above {n. root n (norm (f n)) > c}"
then obtain N where "\<forall>n. root n (norm (f n)) > c \<longrightarrow> n \<le> N" unfolding bdd_above_def by blast
hence "\<exists>N. \<forall>n\<ge>N. root n (norm (f n)) \<le> c"
by (intro exI[of _ "N + 1"]) (force simp: not_less_eq_eq[symmetric])
hence "eventually (\<lambda>n. root n (norm (f n)) \<le> c) sequentially"
by (auto simp: eventually_at_top_linorder)
hence "l \<le> c" unfolding l_def by (intro Limsup_bounded) simp_all
with c show False by auto
qed
from bounded obtain K where K: "K > 0" "\<And>n. norm (f n) \<le> K" using BseqE by blast
define n where "n = nat \<lceil>log c K\<rceil>"
from unbounded have "\<exists>m>n. c < root m (norm (f m))" unfolding bdd_above_def
by (auto simp: not_le)
then guess m by (elim exE conjE) note m = this
from c K have "K = c powr log c K" by (simp add: powr_def log_def)
also from c have "c powr log c K \<le> c powr real n" unfolding n_def
by (intro powr_mono, linarith, simp)
finally have "K \<le> c ^ n" using c by (simp add: powr_realpow)
also from c m have "c ^ n < c ^ m" by simp
also from c m have "c ^ m < root m (norm (f m)) ^ m" by (intro power_strict_mono) simp_all
also from m have "... = norm (f m)" by simp
finally show False using K(2)[of m] by simp
qed
subsubsection \<open>Cauchy's condensation test\<close>
context
fixes f :: "nat \<Rightarrow> real"
begin
private lemma condensation_inequality:
assumes mono: "\<And>m n. 0 < m \<Longrightarrow> m \<le> n \<Longrightarrow> f n \<le> f m"
shows "(\<Sum>k=1..<n. f k) \<ge> (\<Sum>k=1..<n. f (2 * 2 ^ Discrete.log k))" (is "?thesis1")
"(\<Sum>k=1..<n. f k) \<le> (\<Sum>k=1..<n. f (2 ^ Discrete.log k))" (is "?thesis2")
by (intro sum_mono mono Discrete.log_exp2_ge Discrete.log_exp2_le, simp, simp)+
private lemma condensation_condense1: "(\<Sum>k=1..<2^n. f (2 ^ Discrete.log k)) = (\<Sum>k<n. 2^k * f (2 ^ k))"
proof (induction n)
case (Suc n)
have "{1..<2^Suc n} = {1..<2^n} \<union> {2^n..<(2^Suc n :: nat)}" by auto
also have "(\<Sum>k\<in>\<dots>. f (2 ^ Discrete.log k)) =
(\<Sum>k<n. 2^k * f (2^k)) + (\<Sum>k = 2^n..<2^Suc n. f (2^Discrete.log k))"
by (subst sum.union_disjoint) (insert Suc, auto)
also have "Discrete.log k = n" if "k \<in> {2^n..<2^Suc n}" for k using that by (intro Discrete.log_eqI) simp_all
hence "(\<Sum>k = 2^n..<2^Suc n. f (2^Discrete.log k)) = (\<Sum>(_::nat) = 2^n..<2^Suc n. f (2^n))"
by (intro sum.cong) simp_all
- also have "\<dots> = 2^n * f (2^n)" by (simp add: of_nat_power)
+ also have "\<dots> = 2^n * f (2^n)" by (simp)
finally show ?case by simp
qed simp
private lemma condensation_condense2: "(\<Sum>k=1..<2^n. f (2 * 2 ^ Discrete.log k)) = (\<Sum>k<n. 2^k * f (2 ^ Suc k))"
proof (induction n)
case (Suc n)
have "{1..<2^Suc n} = {1..<2^n} \<union> {2^n..<(2^Suc n :: nat)}" by auto
also have "(\<Sum>k\<in>\<dots>. f (2 * 2 ^ Discrete.log k)) =
(\<Sum>k<n. 2^k * f (2^Suc k)) + (\<Sum>k = 2^n..<2^Suc n. f (2 * 2^Discrete.log k))"
by (subst sum.union_disjoint) (insert Suc, auto)
also have "Discrete.log k = n" if "k \<in> {2^n..<2^Suc n}" for k using that by (intro Discrete.log_eqI) simp_all
hence "(\<Sum>k = 2^n..<2^Suc n. f (2*2^Discrete.log k)) = (\<Sum>(_::nat) = 2^n..<2^Suc n. f (2^Suc n))"
by (intro sum.cong) simp_all
- also have "\<dots> = 2^n * f (2^Suc n)" by (simp add: of_nat_power)
+ also have "\<dots> = 2^n * f (2^Suc n)" by (simp)
finally show ?case by simp
qed simp
theorem condensation_test:
assumes mono: "\<And>m. 0 < m \<Longrightarrow> f (Suc m) \<le> f m"
assumes nonneg: "\<And>n. f n \<ge> 0"
shows "summable f \<longleftrightarrow> summable (\<lambda>n. 2^n * f (2^n))"
proof -
define f' where "f' n = (if n = 0 then 0 else f n)" for n
from mono have mono': "decseq (\<lambda>n. f (Suc n))" by (intro decseq_SucI) simp
hence mono': "f n \<le> f m" if "m \<le> n" "m > 0" for m n
using that decseqD[OF mono', of "m - 1" "n - 1"] by simp
have "(\<lambda>n. f (Suc n)) = (\<lambda>n. f' (Suc n))" by (intro ext) (simp add: f'_def)
hence "summable f \<longleftrightarrow> summable f'"
by (subst (1 2) summable_Suc_iff [symmetric]) (simp only:)
also have "\<dots> \<longleftrightarrow> convergent (\<lambda>n. \<Sum>k<n. f' k)" unfolding summable_iff_convergent ..
also have "monoseq (\<lambda>n. \<Sum>k<n. f' k)" unfolding f'_def
by (intro mono_SucI1) (auto intro!: mult_nonneg_nonneg nonneg)
hence "convergent (\<lambda>n. \<Sum>k<n. f' k) \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k<n. f' k)"
by (rule monoseq_imp_convergent_iff_Bseq)
also have "\<dots> \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k=1..<n. f' k)" unfolding One_nat_def
by (subst sum_shift_lb_Suc0_0_upt) (simp_all add: f'_def atLeast0LessThan)
also have "\<dots> \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k=1..<n. f k)" unfolding f'_def by simp
also have "\<dots> \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k=1..<2^n. f k)"
by (rule nonneg_incseq_Bseq_subseq_iff[symmetric])
(auto intro!: sum_nonneg incseq_SucI nonneg simp: strict_mono_def)
also have "\<dots> \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k<n. 2^k * f (2^k))"
proof (intro iffI)
assume A: "Bseq (\<lambda>n. \<Sum>k=1..<2^n. f k)"
have "eventually (\<lambda>n. norm (\<Sum>k<n. 2^k * f (2^Suc k)) \<le> norm (\<Sum>k=1..<2^n. f k)) sequentially"
proof (intro always_eventually allI)
fix n :: nat
have "norm (\<Sum>k<n. 2^k * f (2^Suc k)) = (\<Sum>k<n. 2^k * f (2^Suc k))" unfolding real_norm_def
by (intro abs_of_nonneg sum_nonneg ballI mult_nonneg_nonneg nonneg) simp_all
also have "\<dots> \<le> (\<Sum>k=1..<2^n. f k)"
by (subst condensation_condense2 [symmetric]) (intro condensation_inequality mono')
also have "\<dots> = norm \<dots>" unfolding real_norm_def
by (intro abs_of_nonneg[symmetric] sum_nonneg ballI mult_nonneg_nonneg nonneg)
finally show "norm (\<Sum>k<n. 2 ^ k * f (2 ^ Suc k)) \<le> norm (\<Sum>k=1..<2^n. f k)" .
qed
from this and A have "Bseq (\<lambda>n. \<Sum>k<n. 2^k * f (2^Suc k))" by (rule Bseq_eventually_mono)
from Bseq_mult[OF Bfun_const[of 2] this] have "Bseq (\<lambda>n. \<Sum>k<n. 2^Suc k * f (2^Suc k))"
by (simp add: sum_distrib_left sum_distrib_right mult_ac)
hence "Bseq (\<lambda>n. (\<Sum>k=Suc 0..<Suc n. 2^k * f (2^k)) + f 1)"
by (intro Bseq_add, subst sum.shift_bounds_Suc_ivl) (simp add: atLeast0LessThan)
hence "Bseq (\<lambda>n. (\<Sum>k=0..<Suc n. 2^k * f (2^k)))"
by (subst sum.atLeast_Suc_lessThan) (simp_all add: add_ac)
thus "Bseq (\<lambda>n. (\<Sum>k<n. 2^k * f (2^k)))"
by (subst (asm) Bseq_Suc_iff) (simp add: atLeast0LessThan)
next
assume A: "Bseq (\<lambda>n. (\<Sum>k<n. 2^k * f (2^k)))"
have "eventually (\<lambda>n. norm (\<Sum>k=1..<2^n. f k) \<le> norm (\<Sum>k<n. 2^k * f (2^k))) sequentially"
proof (intro always_eventually allI)
fix n :: nat
have "norm (\<Sum>k=1..<2^n. f k) = (\<Sum>k=1..<2^n. f k)" unfolding real_norm_def
by (intro abs_of_nonneg sum_nonneg ballI mult_nonneg_nonneg nonneg)
also have "\<dots> \<le> (\<Sum>k<n. 2^k * f (2^k))"
by (subst condensation_condense1 [symmetric]) (intro condensation_inequality mono')
also have "\<dots> = norm \<dots>" unfolding real_norm_def
by (intro abs_of_nonneg [symmetric] sum_nonneg ballI mult_nonneg_nonneg nonneg) simp_all
finally show "norm (\<Sum>k=1..<2^n. f k) \<le> norm (\<Sum>k<n. 2^k * f (2^k))" .
qed
from this and A show "Bseq (\<lambda>n. \<Sum>k=1..<2^n. f k)" by (rule Bseq_eventually_mono)
qed
also have "monoseq (\<lambda>n. (\<Sum>k<n. 2^k * f (2^k)))"
by (intro mono_SucI1) (auto intro!: mult_nonneg_nonneg nonneg)
hence "Bseq (\<lambda>n. (\<Sum>k<n. 2^k * f (2^k))) \<longleftrightarrow> convergent (\<lambda>n. (\<Sum>k<n. 2^k * f (2^k)))"
by (rule monoseq_imp_convergent_iff_Bseq [symmetric])
also have "\<dots> \<longleftrightarrow> summable (\<lambda>k. 2^k * f (2^k))" by (simp only: summable_iff_convergent)
finally show ?thesis .
qed
end
subsubsection \<open>Summability of powers\<close>
lemma abs_summable_complex_powr_iff:
"summable (\<lambda>n. norm (exp (of_real (ln (of_nat n)) * s))) \<longleftrightarrow> Re s < -1"
proof (cases "Re s \<le> 0")
let ?l = "\<lambda>n. complex_of_real (ln (of_nat n))"
case False
have "eventually (\<lambda>n. norm (1 :: real) \<le> norm (exp (?l n * s))) sequentially"
apply (rule eventually_mono [OF eventually_gt_at_top[of "0::nat"]])
using False ge_one_powr_ge_zero by auto
from summable_comparison_test_ev[OF this] False show ?thesis by (auto simp: summable_const_iff)
next
let ?l = "\<lambda>n. complex_of_real (ln (of_nat n))"
case True
hence "summable (\<lambda>n. norm (exp (?l n * s))) \<longleftrightarrow> summable (\<lambda>n. 2^n * norm (exp (?l (2^n) * s)))"
by (intro condensation_test) (auto intro!: mult_right_mono_neg)
also have "(\<lambda>n. 2^n * norm (exp (?l (2^n) * s))) = (\<lambda>n. (2 powr (Re s + 1)) ^ n)"
proof
fix n :: nat
have "2^n * norm (exp (?l (2^n) * s)) = exp (real n * ln 2) * exp (real n * ln 2 * Re s)"
using True by (subst exp_of_nat_mult) (simp add: ln_realpow algebra_simps)
also have "\<dots> = exp (real n * (ln 2 * (Re s + 1)))"
by (simp add: algebra_simps exp_add)
also have "\<dots> = exp (ln 2 * (Re s + 1)) ^ n" by (subst exp_of_nat_mult) simp
also have "exp (ln 2 * (Re s + 1)) = 2 powr (Re s + 1)" by (simp add: powr_def)
finally show "2^n * norm (exp (?l (2^n) * s)) = (2 powr (Re s + 1)) ^ n" .
qed
also have "summable \<dots> \<longleftrightarrow> 2 powr (Re s + 1) < 2 powr 0"
by (subst summable_geometric_iff) simp
also have "\<dots> \<longleftrightarrow> Re s < -1" by (subst powr_less_cancel_iff) (simp, linarith)
finally show ?thesis .
qed
theorem summable_complex_powr_iff:
assumes "Re s < -1"
shows "summable (\<lambda>n. exp (of_real (ln (of_nat n)) * s))"
by (rule summable_norm_cancel, subst abs_summable_complex_powr_iff) fact
lemma summable_real_powr_iff: "summable (\<lambda>n. of_nat n powr s :: real) \<longleftrightarrow> s < -1"
proof -
from eventually_gt_at_top[of "0::nat"]
have "summable (\<lambda>n. of_nat n powr s) \<longleftrightarrow> summable (\<lambda>n. exp (ln (of_nat n) * s))"
by (intro summable_cong) (auto elim!: eventually_mono simp: powr_def)
also have "\<dots> \<longleftrightarrow> s < -1" using abs_summable_complex_powr_iff[of "of_real s"] by simp
finally show ?thesis .
qed
lemma inverse_power_summable:
assumes s: "s \<ge> 2"
shows "summable (\<lambda>n. inverse (of_nat n ^ s :: 'a :: {real_normed_div_algebra,banach}))"
proof (rule summable_norm_cancel, subst summable_cong)
from eventually_gt_at_top[of "0::nat"]
show "eventually (\<lambda>n. norm (inverse (of_nat n ^ s:: 'a)) = real_of_nat n powr (-real s)) at_top"
by eventually_elim (simp add: norm_inverse norm_power powr_minus powr_realpow)
qed (insert s summable_real_powr_iff[of "-s"], simp_all)
lemma not_summable_harmonic: "\<not>summable (\<lambda>n. inverse (of_nat n) :: 'a :: real_normed_field)"
proof
assume "summable (\<lambda>n. inverse (of_nat n) :: 'a)"
hence "convergent (\<lambda>n. norm (of_real (\<Sum>k<n. inverse (of_nat k)) :: 'a))"
by (simp add: summable_iff_convergent convergent_norm)
hence "convergent (\<lambda>n. abs (\<Sum>k<n. inverse (of_nat k)) :: real)" by (simp only: norm_of_real)
also have "(\<lambda>n. abs (\<Sum>k<n. inverse (of_nat k)) :: real) = (\<lambda>n. \<Sum>k<n. inverse (of_nat k))"
by (intro ext abs_of_nonneg sum_nonneg) auto
also have "convergent \<dots> \<longleftrightarrow> summable (\<lambda>k. inverse (of_nat k) :: real)"
by (simp add: summable_iff_convergent)
finally show False using summable_real_powr_iff[of "-1"] by (simp add: powr_minus)
qed
subsubsection \<open>Kummer's test\<close>
theorem kummers_test_convergence:
fixes f p :: "nat \<Rightarrow> real"
assumes pos_f: "eventually (\<lambda>n. f n > 0) sequentially"
assumes nonneg_p: "eventually (\<lambda>n. p n \<ge> 0) sequentially"
defines "l \<equiv> liminf (\<lambda>n. ereal (p n * f n / f (Suc n) - p (Suc n)))"
assumes l: "l > 0"
shows "summable f"
unfolding summable_iff_convergent'
proof -
define r where "r = (if l = \<infinity> then 1 else real_of_ereal l / 2)"
from l have "r > 0 \<and> of_real r < l" by (cases l) (simp_all add: r_def)
hence r: "r > 0" "of_real r < l" by simp_all
hence "eventually (\<lambda>n. p n * f n / f (Suc n) - p (Suc n) > r) sequentially"
unfolding l_def by (force dest: less_LiminfD)
moreover from pos_f have "eventually (\<lambda>n. f (Suc n) > 0) sequentially"
by (subst eventually_sequentially_Suc)
ultimately have "eventually (\<lambda>n. p n * f n - p (Suc n) * f (Suc n) > r * f (Suc n)) sequentially"
by eventually_elim (simp add: field_simps)
from eventually_conj[OF pos_f eventually_conj[OF nonneg_p this]]
obtain m where m: "\<And>n. n \<ge> m \<Longrightarrow> f n > 0" "\<And>n. n \<ge> m \<Longrightarrow> p n \<ge> 0"
"\<And>n. n \<ge> m \<Longrightarrow> p n * f n - p (Suc n) * f (Suc n) > r * f (Suc n)"
unfolding eventually_at_top_linorder by blast
let ?c = "(norm (\<Sum>k\<le>m. r * f k) + p m * f m) / r"
have "Bseq (\<lambda>n. (\<Sum>k\<le>n + Suc m. f k))"
proof (rule BseqI')
fix k :: nat
define n where "n = k + Suc m"
have n: "n > m" by (simp add: n_def)
from r have "r * norm (\<Sum>k\<le>n. f k) = norm (\<Sum>k\<le>n. r * f k)"
by (simp add: sum_distrib_left[symmetric] abs_mult)
also from n have "{..n} = {..m} \<union> {Suc m..n}" by auto
hence "(\<Sum>k\<le>n. r * f k) = (\<Sum>k\<in>{..m} \<union> {Suc m..n}. r * f k)" by (simp only:)
also have "\<dots> = (\<Sum>k\<le>m. r * f k) + (\<Sum>k=Suc m..n. r * f k)"
by (subst sum.union_disjoint) auto
also have "norm \<dots> \<le> norm (\<Sum>k\<le>m. r * f k) + norm (\<Sum>k=Suc m..n. r * f k)"
by (rule norm_triangle_ineq)
also from r less_imp_le[OF m(1)] have "(\<Sum>k=Suc m..n. r * f k) \<ge> 0"
by (intro sum_nonneg) auto
hence "norm (\<Sum>k=Suc m..n. r * f k) = (\<Sum>k=Suc m..n. r * f k)" by simp
also have "(\<Sum>k=Suc m..n. r * f k) = (\<Sum>k=m..<n. r * f (Suc k))"
by (subst sum.shift_bounds_Suc_ivl [symmetric])
(simp only: atLeastLessThanSuc_atLeastAtMost)
also from m have "\<dots> \<le> (\<Sum>k=m..<n. p k * f k - p (Suc k) * f (Suc k))"
by (intro sum_mono[OF less_imp_le]) simp_all
also have "\<dots> = -(\<Sum>k=m..<n. p (Suc k) * f (Suc k) - p k * f k)"
by (simp add: sum_negf [symmetric] algebra_simps)
also from n have "\<dots> = p m * f m - p n * f n"
by (cases n, simp, simp only: atLeastLessThanSuc_atLeastAtMost, subst sum_Suc_diff) simp_all
also from less_imp_le[OF m(1)] m(2) n have "\<dots> \<le> p m * f m" by simp
finally show "norm (\<Sum>k\<le>n. f k) \<le> (norm (\<Sum>k\<le>m. r * f k) + p m * f m) / r" using r
by (subst pos_le_divide_eq[OF r(1)]) (simp only: mult_ac)
qed
moreover have "(\<Sum>k\<le>n. f k) \<le> (\<Sum>k\<le>n'. f k)" if "Suc m \<le> n" "n \<le> n'" for n n'
using less_imp_le[OF m(1)] that by (intro sum_mono2) auto
ultimately show "convergent (\<lambda>n. \<Sum>k\<le>n. f k)" by (rule Bseq_monoseq_convergent'_inc)
qed
theorem kummers_test_divergence:
fixes f p :: "nat \<Rightarrow> real"
assumes pos_f: "eventually (\<lambda>n. f n > 0) sequentially"
assumes pos_p: "eventually (\<lambda>n. p n > 0) sequentially"
assumes divergent_p: "\<not>summable (\<lambda>n. inverse (p n))"
defines "l \<equiv> limsup (\<lambda>n. ereal (p n * f n / f (Suc n) - p (Suc n)))"
assumes l: "l < 0"
shows "\<not>summable f"
proof
assume "summable f"
from eventually_conj[OF pos_f eventually_conj[OF pos_p Limsup_lessD[OF l[unfolded l_def]]]]
obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> p n > 0" "\<And>n. n \<ge> N \<Longrightarrow> f n > 0"
"\<And>n. n \<ge> N \<Longrightarrow> p n * f n / f (Suc n) - p (Suc n) < 0"
by (auto simp: eventually_at_top_linorder)
hence A: "p n * f n < p (Suc n) * f (Suc n)" if "n \<ge> N" for n using that N[of n] N[of "Suc n"]
by (simp add: field_simps)
have B: "p n * f n \<ge> p N * f N" if "n \<ge> N" for n using that and A
by (induction n rule: dec_induct) (auto intro!: less_imp_le elim!: order.trans)
have "eventually (\<lambda>n. norm (p N * f N * inverse (p n)) \<le> f n) sequentially"
apply (rule eventually_mono [OF eventually_ge_at_top[of N]])
using B N by (auto simp: field_simps abs_of_pos)
from this and \<open>summable f\<close> have "summable (\<lambda>n. p N * f N * inverse (p n))"
by (rule summable_comparison_test_ev)
from summable_mult[OF this, of "inverse (p N * f N)"] N[OF le_refl]
have "summable (\<lambda>n. inverse (p n))" by (simp add: field_split_simps)
with divergent_p show False by contradiction
qed
subsubsection \<open>Ratio test\<close>
theorem ratio_test_convergence:
fixes f :: "nat \<Rightarrow> real"
assumes pos_f: "eventually (\<lambda>n. f n > 0) sequentially"
defines "l \<equiv> liminf (\<lambda>n. ereal (f n / f (Suc n)))"
assumes l: "l > 1"
shows "summable f"
proof (rule kummers_test_convergence[OF pos_f])
note l
also have "l = liminf (\<lambda>n. ereal (f n / f (Suc n) - 1)) + 1"
by (subst Liminf_add_ereal_right[symmetric]) (simp_all add: minus_ereal_def l_def one_ereal_def)
finally show "liminf (\<lambda>n. ereal (1 * f n / f (Suc n) - 1)) > 0"
by (cases "liminf (\<lambda>n. ereal (1 * f n / f (Suc n) - 1))") simp_all
qed simp
theorem ratio_test_divergence:
fixes f :: "nat \<Rightarrow> real"
assumes pos_f: "eventually (\<lambda>n. f n > 0) sequentially"
defines "l \<equiv> limsup (\<lambda>n. ereal (f n / f (Suc n)))"
assumes l: "l < 1"
shows "\<not>summable f"
proof (rule kummers_test_divergence[OF pos_f])
have "limsup (\<lambda>n. ereal (f n / f (Suc n) - 1)) + 1 = l"
by (subst Limsup_add_ereal_right[symmetric]) (simp_all add: minus_ereal_def l_def one_ereal_def)
also note l
finally show "limsup (\<lambda>n. ereal (1 * f n / f (Suc n) - 1)) < 0"
by (cases "limsup (\<lambda>n. ereal (1 * f n / f (Suc n) - 1))") simp_all
qed (simp_all add: summable_const_iff)
subsubsection \<open>Raabe's test\<close>
theorem raabes_test_convergence:
fixes f :: "nat \<Rightarrow> real"
assumes pos: "eventually (\<lambda>n. f n > 0) sequentially"
defines "l \<equiv> liminf (\<lambda>n. ereal (of_nat n * (f n / f (Suc n) - 1)))"
assumes l: "l > 1"
shows "summable f"
proof (rule kummers_test_convergence)
let ?l' = "liminf (\<lambda>n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)))"
have "1 < l" by fact
also have "l = liminf (\<lambda>n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)) + 1)"
by (simp add: l_def algebra_simps)
also have "\<dots> = ?l' + 1" by (subst Liminf_add_ereal_right) simp_all
finally show "?l' > 0" by (cases ?l') (simp_all add: algebra_simps)
qed (simp_all add: pos)
theorem raabes_test_divergence:
fixes f :: "nat \<Rightarrow> real"
assumes pos: "eventually (\<lambda>n. f n > 0) sequentially"
defines "l \<equiv> limsup (\<lambda>n. ereal (of_nat n * (f n / f (Suc n) - 1)))"
assumes l: "l < 1"
shows "\<not>summable f"
proof (rule kummers_test_divergence)
let ?l' = "limsup (\<lambda>n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)))"
note l
also have "l = limsup (\<lambda>n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)) + 1)"
by (simp add: l_def algebra_simps)
also have "\<dots> = ?l' + 1" by (subst Limsup_add_ereal_right) simp_all
finally show "?l' < 0" by (cases ?l') (simp_all add: algebra_simps)
qed (insert pos eventually_gt_at_top[of "0::nat"] not_summable_harmonic, simp_all)
subsection \<open>Radius of convergence\<close>
text \<open>
The radius of convergence of a power series. This value always exists, ranges from
\<^term>\<open>0::ereal\<close> to \<^term>\<open>\<infinity>::ereal\<close>, and the power series is guaranteed to converge for
all inputs with a norm that is smaller than that radius and to diverge for all inputs with a
norm that is greater.
\<close>
definition\<^marker>\<open>tag important\<close> conv_radius :: "(nat \<Rightarrow> 'a :: banach) \<Rightarrow> ereal" where
"conv_radius f = inverse (limsup (\<lambda>n. ereal (root n (norm (f n)))))"
lemma conv_radius_cong_weak [cong]: "(\<And>n. f n = g n) \<Longrightarrow> conv_radius f = conv_radius g"
by (drule ext) simp_all
lemma conv_radius_nonneg: "conv_radius f \<ge> 0"
proof -
have "0 = limsup (\<lambda>n. 0)" by (subst Limsup_const) simp_all
also have "\<dots> \<le> limsup (\<lambda>n. ereal (root n (norm (f n))))"
by (intro Limsup_mono) (simp_all add: real_root_ge_zero)
finally show ?thesis
unfolding conv_radius_def by (auto simp: ereal_inverse_nonneg_iff)
qed
lemma conv_radius_zero [simp]: "conv_radius (\<lambda>_. 0) = \<infinity>"
by (auto simp: conv_radius_def zero_ereal_def [symmetric] Limsup_const)
lemma conv_radius_altdef:
"conv_radius f = liminf (\<lambda>n. inverse (ereal (root n (norm (f n)))))"
by (subst Liminf_inverse_ereal) (simp_all add: real_root_ge_zero conv_radius_def)
lemma conv_radius_cong':
assumes "eventually (\<lambda>x. f x = g x) sequentially"
shows "conv_radius f = conv_radius g"
unfolding conv_radius_altdef by (intro Liminf_eq eventually_mono [OF assms]) auto
lemma conv_radius_cong:
assumes "eventually (\<lambda>x. norm (f x) = norm (g x)) sequentially"
shows "conv_radius f = conv_radius g"
unfolding conv_radius_altdef by (intro Liminf_eq eventually_mono [OF assms]) auto
theorem abs_summable_in_conv_radius:
fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
assumes "ereal (norm z) < conv_radius f"
shows "summable (\<lambda>n. norm (f n * z ^ n))"
proof (rule root_test_convergence')
define l where "l = limsup (\<lambda>n. ereal (root n (norm (f n))))"
have "0 = limsup (\<lambda>n. 0)" by (simp add: Limsup_const)
also have "... \<le> l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero)
finally have l_nonneg: "l \<ge> 0" .
have "limsup (\<lambda>n. root n (norm (f n * z^n))) = l * ereal (norm z)" unfolding l_def
by (rule limsup_root_powser)
also from l_nonneg consider "l = 0" | "l = \<infinity>" | "\<exists>l'. l = ereal l' \<and> l' > 0"
by (cases "l") (auto simp: less_le)
hence "l * ereal (norm z) < 1"
proof cases
assume "l = \<infinity>"
hence "conv_radius f = 0" unfolding conv_radius_def l_def by simp
with assms show ?thesis by simp
next
assume "\<exists>l'. l = ereal l' \<and> l' > 0"
then guess l' by (elim exE conjE) note l' = this
hence "l \<noteq> \<infinity>" by auto
have "l * ereal (norm z) < l * conv_radius f"
by (intro ereal_mult_strict_left_mono) (simp_all add: l' assms)
also have "conv_radius f = inverse l" by (simp add: conv_radius_def l_def)
also from l' have "l * inverse l = 1" by simp
finally show ?thesis .
qed simp_all
finally show "limsup (\<lambda>n. ereal (root n (norm (norm (f n * z ^ n))))) < 1" by simp
qed
lemma summable_in_conv_radius:
fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
assumes "ereal (norm z) < conv_radius f"
shows "summable (\<lambda>n. f n * z ^ n)"
by (rule summable_norm_cancel, rule abs_summable_in_conv_radius) fact+
theorem not_summable_outside_conv_radius:
fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
assumes "ereal (norm z) > conv_radius f"
shows "\<not>summable (\<lambda>n. f n * z ^ n)"
proof (rule root_test_divergence)
define l where "l = limsup (\<lambda>n. ereal (root n (norm (f n))))"
have "0 = limsup (\<lambda>n. 0)" by (simp add: Limsup_const)
also have "... \<le> l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero)
finally have l_nonneg: "l \<ge> 0" .
from assms have l_nz: "l \<noteq> 0" unfolding conv_radius_def l_def by auto
have "limsup (\<lambda>n. ereal (root n (norm (f n * z^n)))) = l * ereal (norm z)"
unfolding l_def by (rule limsup_root_powser)
also have "... > 1"
proof (cases l)
assume "l = \<infinity>"
with assms conv_radius_nonneg[of f] show ?thesis
by (auto simp: zero_ereal_def[symmetric])
next
fix l' assume l': "l = ereal l'"
from l_nonneg l_nz have "1 = l * inverse l" by (auto simp: l' field_simps)
also from l_nz have "inverse l = conv_radius f"
unfolding l_def conv_radius_def by auto
also from l' l_nz l_nonneg assms have "l * \<dots> < l * ereal (norm z)"
by (intro ereal_mult_strict_left_mono) (auto simp: l')
finally show ?thesis .
qed (insert l_nonneg, simp_all)
finally show "limsup (\<lambda>n. ereal (root n (norm (f n * z^n)))) > 1" .
qed
lemma conv_radius_geI:
assumes "summable (\<lambda>n. f n * z ^ n :: 'a :: {banach, real_normed_div_algebra})"
shows "conv_radius f \<ge> norm z"
using not_summable_outside_conv_radius[of f z] assms by (force simp: not_le[symmetric])
lemma conv_radius_leI:
assumes "\<not>summable (\<lambda>n. norm (f n * z ^ n :: 'a :: {banach, real_normed_div_algebra}))"
shows "conv_radius f \<le> norm z"
using abs_summable_in_conv_radius[of z f] assms by (force simp: not_le[symmetric])
lemma conv_radius_leI':
assumes "\<not>summable (\<lambda>n. f n * z ^ n :: 'a :: {banach, real_normed_div_algebra})"
shows "conv_radius f \<le> norm z"
using summable_in_conv_radius[of z f] assms by (force simp: not_le[symmetric])
lemma conv_radius_geI_ex:
fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
assumes "\<And>r. 0 < r \<Longrightarrow> ereal r < R \<Longrightarrow> \<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z^n)"
shows "conv_radius f \<ge> R"
proof (rule linorder_cases[of "conv_radius f" R])
assume R: "conv_radius f < R"
with conv_radius_nonneg[of f] obtain conv_radius'
where [simp]: "conv_radius f = ereal conv_radius'"
by (cases "conv_radius f") simp_all
define r where "r = (if R = \<infinity> then conv_radius' + 1 else (real_of_ereal R + conv_radius') / 2)"
from R conv_radius_nonneg[of f] have "0 < r \<and> ereal r < R \<and> ereal r > conv_radius f"
unfolding r_def by (cases R) (auto simp: r_def field_simps)
with assms(1)[of r] obtain z where "norm z > conv_radius f" "summable (\<lambda>n. f n * z^n)" by auto
with not_summable_outside_conv_radius[of f z] show ?thesis by simp
qed simp_all
lemma conv_radius_geI_ex':
fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
assumes "\<And>r. 0 < r \<Longrightarrow> ereal r < R \<Longrightarrow> summable (\<lambda>n. f n * of_real r^n)"
shows "conv_radius f \<ge> R"
proof (rule conv_radius_geI_ex)
fix r assume "0 < r" "ereal r < R"
with assms[of r] show "\<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z ^ n)"
by (intro exI[of _ "of_real r :: 'a"]) auto
qed
lemma conv_radius_leI_ex:
fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
assumes "R \<ge> 0"
assumes "\<And>r. 0 < r \<Longrightarrow> ereal r > R \<Longrightarrow> \<exists>z. norm z = r \<and> \<not>summable (\<lambda>n. norm (f n * z^n))"
shows "conv_radius f \<le> R"
proof (rule linorder_cases[of "conv_radius f" R])
assume R: "conv_radius f > R"
from R assms(1) obtain R' where R': "R = ereal R'" by (cases R) simp_all
define r where
"r = (if conv_radius f = \<infinity> then R' + 1 else (R' + real_of_ereal (conv_radius f)) / 2)"
from R conv_radius_nonneg[of f] have "r > R \<and> r < conv_radius f" unfolding r_def
by (cases "conv_radius f") (auto simp: r_def field_simps R')
with assms(1) assms(2)[of r] R'
obtain z where "norm z < conv_radius f" "\<not>summable (\<lambda>n. norm (f n * z^n))" by auto
with abs_summable_in_conv_radius[of z f] show ?thesis by auto
qed simp_all
lemma conv_radius_leI_ex':
fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
assumes "R \<ge> 0"
assumes "\<And>r. 0 < r \<Longrightarrow> ereal r > R \<Longrightarrow> \<not>summable (\<lambda>n. f n * of_real r^n)"
shows "conv_radius f \<le> R"
proof (rule conv_radius_leI_ex)
fix r assume "0 < r" "ereal r > R"
with assms(2)[of r] show "\<exists>z. norm z = r \<and> \<not>summable (\<lambda>n. norm (f n * z ^ n))"
by (intro exI[of _ "of_real r :: 'a"]) (auto dest: summable_norm_cancel)
qed fact+
lemma conv_radius_eqI:
fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
assumes "R \<ge> 0"
assumes "\<And>r. 0 < r \<Longrightarrow> ereal r < R \<Longrightarrow> \<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z^n)"
assumes "\<And>r. 0 < r \<Longrightarrow> ereal r > R \<Longrightarrow> \<exists>z. norm z = r \<and> \<not>summable (\<lambda>n. norm (f n * z^n))"
shows "conv_radius f = R"
by (intro antisym conv_radius_geI_ex conv_radius_leI_ex assms)
lemma conv_radius_eqI':
fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
assumes "R \<ge> 0"
assumes "\<And>r. 0 < r \<Longrightarrow> ereal r < R \<Longrightarrow> summable (\<lambda>n. f n * (of_real r)^n)"
assumes "\<And>r. 0 < r \<Longrightarrow> ereal r > R \<Longrightarrow> \<not>summable (\<lambda>n. norm (f n * (of_real r)^n))"
shows "conv_radius f = R"
proof (intro conv_radius_eqI[OF assms(1)])
fix r assume "0 < r" "ereal r < R" with assms(2)[OF this]
show "\<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z ^ n)" by force
next
fix r assume "0 < r" "ereal r > R" with assms(3)[OF this]
show "\<exists>z. norm z = r \<and> \<not>summable (\<lambda>n. norm (f n * z ^ n))" by force
qed
lemma conv_radius_zeroI:
fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
assumes "\<And>z. z \<noteq> 0 \<Longrightarrow> \<not>summable (\<lambda>n. f n * z^n)"
shows "conv_radius f = 0"
proof (rule ccontr)
assume "conv_radius f \<noteq> 0"
with conv_radius_nonneg[of f] have pos: "conv_radius f > 0" by simp
define r where "r = (if conv_radius f = \<infinity> then 1 else real_of_ereal (conv_radius f) / 2)"
from pos have r: "ereal r > 0 \<and> ereal r < conv_radius f"
by (cases "conv_radius f") (simp_all add: r_def)
hence "summable (\<lambda>n. f n * of_real r ^ n)" by (intro summable_in_conv_radius) simp
moreover from r and assms[of "of_real r"] have "\<not>summable (\<lambda>n. f n * of_real r ^ n)" by simp
ultimately show False by contradiction
qed
lemma conv_radius_inftyI':
fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
assumes "\<And>r. r > c \<Longrightarrow> \<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z^n)"
shows "conv_radius f = \<infinity>"
proof -
{
fix r :: real
have "max r (c + 1) > c" by (auto simp: max_def)
from assms[OF this] obtain z where "norm z = max r (c + 1)" "summable (\<lambda>n. f n * z^n)" by blast
from conv_radius_geI[OF this(2)] this(1) have "conv_radius f \<ge> r" by simp
}
from this[of "real_of_ereal (conv_radius f + 1)"] show "conv_radius f = \<infinity>"
by (cases "conv_radius f") simp_all
qed
lemma conv_radius_inftyI:
fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
assumes "\<And>r. \<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z^n)"
shows "conv_radius f = \<infinity>"
using assms by (rule conv_radius_inftyI')
lemma conv_radius_inftyI'':
fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
assumes "\<And>z. summable (\<lambda>n. f n * z^n)"
shows "conv_radius f = \<infinity>"
proof (rule conv_radius_inftyI')
fix r :: real assume "r > 0"
with assms show "\<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z^n)"
by (intro exI[of _ "of_real r"]) simp
qed
lemma conv_radius_conv_Sup:
fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
shows "conv_radius f = Sup {r. \<forall>z. ereal (norm z) < r \<longrightarrow> summable (\<lambda>n. f n * z ^ n)}"
proof (rule Sup_eqI [symmetric], goal_cases)
case (1 r)
thus ?case
by (intro conv_radius_geI_ex') auto
next
case (2 r)
from 2[of 0] have r: "r \<ge> 0" by auto
show ?case
proof (intro conv_radius_leI_ex' r)
fix R assume R: "R > 0" "ereal R > r"
with r obtain r' where [simp]: "r = ereal r'" by (cases r) auto
show "\<not>summable (\<lambda>n. f n * of_real R ^ n)"
proof
assume *: "summable (\<lambda>n. f n * of_real R ^ n)"
define R' where "R' = (R + r') / 2"
from R have R': "R' > r'" "R' < R" by (simp_all add: R'_def)
hence "\<forall>z. norm z < R' \<longrightarrow> summable (\<lambda>n. f n * z ^ n)"
using powser_inside[OF *] by auto
from 2[of R'] and this have "R' \<le> r'" by auto
with \<open>R' > r'\<close> show False by simp
qed
qed
qed
lemma conv_radius_shift:
fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
shows "conv_radius (\<lambda>n. f (n + m)) = conv_radius f"
unfolding conv_radius_conv_Sup summable_powser_ignore_initial_segment ..
lemma conv_radius_norm [simp]: "conv_radius (\<lambda>x. norm (f x)) = conv_radius f"
by (simp add: conv_radius_def)
lemma conv_radius_ratio_limit_ereal:
fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
assumes nz: "eventually (\<lambda>n. f n \<noteq> 0) sequentially"
assumes lim: "(\<lambda>n. ereal (norm (f n) / norm (f (Suc n)))) \<longlonglongrightarrow> c"
shows "conv_radius f = c"
proof (rule conv_radius_eqI')
show "c \<ge> 0" by (intro Lim_bounded2[OF lim]) simp_all
next
fix r assume r: "0 < r" "ereal r < c"
let ?l = "liminf (\<lambda>n. ereal (norm (f n * of_real r ^ n) / norm (f (Suc n) * of_real r ^ Suc n)))"
have "?l = liminf (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n)))) * ereal (inverse r))"
using r by (simp add: norm_mult norm_power field_split_simps)
also from r have "\<dots> = liminf (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n))))) * ereal (inverse r)"
by (intro Liminf_ereal_mult_right) simp_all
also have "liminf (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n))))) = c"
by (intro lim_imp_Liminf lim) simp
finally have l: "?l = c * ereal (inverse r)" by simp
from r have l': "c * ereal (inverse r) > 1" by (cases c) (simp_all add: field_simps)
show "summable (\<lambda>n. f n * of_real r^n)"
by (rule summable_norm_cancel, rule ratio_test_convergence)
(insert r nz l l', auto elim!: eventually_mono)
next
fix r assume r: "0 < r" "ereal r > c"
let ?l = "limsup (\<lambda>n. ereal (norm (f n * of_real r ^ n) / norm (f (Suc n) * of_real r ^ Suc n)))"
have "?l = limsup (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n)))) * ereal (inverse r))"
using r by (simp add: norm_mult norm_power field_split_simps)
also from r have "\<dots> = limsup (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n))))) * ereal (inverse r)"
by (intro Limsup_ereal_mult_right) simp_all
also have "limsup (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n))))) = c"
by (intro lim_imp_Limsup lim) simp
finally have l: "?l = c * ereal (inverse r)" by simp
from r have l': "c * ereal (inverse r) < 1" by (cases c) (simp_all add: field_simps)
show "\<not>summable (\<lambda>n. norm (f n * of_real r^n))"
by (rule ratio_test_divergence) (insert r nz l l', auto elim!: eventually_mono)
qed
lemma conv_radius_ratio_limit_ereal_nonzero:
fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
assumes nz: "c \<noteq> 0"
assumes lim: "(\<lambda>n. ereal (norm (f n) / norm (f (Suc n)))) \<longlonglongrightarrow> c"
shows "conv_radius f = c"
proof (rule conv_radius_ratio_limit_ereal[OF _ lim], rule ccontr)
assume "\<not>eventually (\<lambda>n. f n \<noteq> 0) sequentially"
hence "frequently (\<lambda>n. f n = 0) sequentially" by (simp add: frequently_def)
hence "frequently (\<lambda>n. ereal (norm (f n) / norm (f (Suc n))) = 0) sequentially"
by (force elim!: frequently_elim1)
hence "c = 0" by (intro limit_frequently_eq[OF _ _ lim]) auto
with nz show False by contradiction
qed
lemma conv_radius_ratio_limit:
fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
assumes "c' = ereal c"
assumes nz: "eventually (\<lambda>n. f n \<noteq> 0) sequentially"
assumes lim: "(\<lambda>n. norm (f n) / norm (f (Suc n))) \<longlonglongrightarrow> c"
shows "conv_radius f = c'"
using assms by (intro conv_radius_ratio_limit_ereal) simp_all
lemma conv_radius_ratio_limit_nonzero:
fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
assumes "c' = ereal c"
assumes nz: "c \<noteq> 0"
assumes lim: "(\<lambda>n. norm (f n) / norm (f (Suc n))) \<longlonglongrightarrow> c"
shows "conv_radius f = c'"
using assms by (intro conv_radius_ratio_limit_ereal_nonzero) simp_all
lemma conv_radius_cmult_left:
assumes "c \<noteq> (0 :: 'a :: {banach, real_normed_div_algebra})"
shows "conv_radius (\<lambda>n. c * f n) = conv_radius f"
proof -
have "conv_radius (\<lambda>n. c * f n) =
inverse (limsup (\<lambda>n. ereal (root n (norm (c * f n)))))"
unfolding conv_radius_def ..
also have "(\<lambda>n. ereal (root n (norm (c * f n)))) =
(\<lambda>n. ereal (root n (norm c)) * ereal (root n (norm (f n))))"
by (rule ext) (auto simp: norm_mult real_root_mult)
also have "limsup \<dots> = ereal 1 * limsup (\<lambda>n. ereal (root n (norm (f n))))"
using assms by (intro ereal_limsup_lim_mult tendsto_ereal LIMSEQ_root_const) auto
also have "inverse \<dots> = conv_radius f" by (simp add: conv_radius_def)
finally show ?thesis .
qed
lemma conv_radius_cmult_right:
assumes "c \<noteq> (0 :: 'a :: {banach, real_normed_div_algebra})"
shows "conv_radius (\<lambda>n. f n * c) = conv_radius f"
proof -
have "conv_radius (\<lambda>n. f n * c) = conv_radius (\<lambda>n. c * f n)"
by (simp add: conv_radius_def norm_mult mult.commute)
with conv_radius_cmult_left[OF assms, of f] show ?thesis by simp
qed
lemma conv_radius_mult_power:
assumes "c \<noteq> (0 :: 'a :: {real_normed_div_algebra,banach})"
shows "conv_radius (\<lambda>n. c ^ n * f n) = conv_radius f / ereal (norm c)"
proof -
have "limsup (\<lambda>n. ereal (root n (norm (c ^ n * f n)))) =
limsup (\<lambda>n. ereal (norm c) * ereal (root n (norm (f n))))"
by (intro Limsup_eq eventually_mono [OF eventually_gt_at_top[of "0::nat"]])
(auto simp: norm_mult norm_power real_root_mult real_root_power)
also have "\<dots> = ereal (norm c) * limsup (\<lambda>n. ereal (root n (norm (f n))))"
using assms by (subst Limsup_ereal_mult_left[symmetric]) simp_all
finally have A: "limsup (\<lambda>n. ereal (root n (norm (c ^ n * f n)))) =
ereal (norm c) * limsup (\<lambda>n. ereal (root n (norm (f n))))" .
show ?thesis using assms
apply (cases "limsup (\<lambda>n. ereal (root n (norm (f n)))) = 0")
apply (simp add: A conv_radius_def)
apply (unfold conv_radius_def A divide_ereal_def, simp add: mult.commute ereal_inverse_mult)
done
qed
lemma conv_radius_mult_power_right:
assumes "c \<noteq> (0 :: 'a :: {real_normed_div_algebra,banach})"
shows "conv_radius (\<lambda>n. f n * c ^ n) = conv_radius f / ereal (norm c)"
using conv_radius_mult_power[OF assms, of f]
unfolding conv_radius_def by (simp add: mult.commute norm_mult)
lemma conv_radius_divide_power:
assumes "c \<noteq> (0 :: 'a :: {real_normed_div_algebra,banach})"
shows "conv_radius (\<lambda>n. f n / c^n) = conv_radius f * ereal (norm c)"
proof -
from assms have "inverse c \<noteq> 0" by simp
from conv_radius_mult_power_right[OF this, of f] show ?thesis
by (simp add: divide_inverse divide_ereal_def assms norm_inverse power_inverse)
qed
lemma conv_radius_add_ge:
"min (conv_radius f) (conv_radius g) \<le>
conv_radius (\<lambda>x. f x + g x :: 'a :: {banach,real_normed_div_algebra})"
by (rule conv_radius_geI_ex')
(auto simp: algebra_simps intro!: summable_add summable_in_conv_radius)
lemma conv_radius_mult_ge:
fixes f g :: "nat \<Rightarrow> ('a :: {banach,real_normed_div_algebra})"
shows "conv_radius (\<lambda>x. \<Sum>i\<le>x. f i * g (x - i)) \<ge> min (conv_radius f) (conv_radius g)"
proof (rule conv_radius_geI_ex')
fix r assume r: "r > 0" "ereal r < min (conv_radius f) (conv_radius g)"
from r have "summable (\<lambda>n. (\<Sum>i\<le>n. (f i * of_real r^i) * (g (n - i) * of_real r^(n - i))))"
by (intro summable_Cauchy_product abs_summable_in_conv_radius) simp_all
thus "summable (\<lambda>n. (\<Sum>i\<le>n. f i * g (n - i)) * of_real r ^ n)"
by (simp add: algebra_simps of_real_def power_add [symmetric] scaleR_sum_right)
qed
lemma le_conv_radius_iff:
fixes a :: "nat \<Rightarrow> 'a::{real_normed_div_algebra,banach}"
shows "r \<le> conv_radius a \<longleftrightarrow> (\<forall>x. norm (x-\<xi>) < r \<longrightarrow> summable (\<lambda>i. a i * (x - \<xi>) ^ i))"
apply (intro iffI allI impI summable_in_conv_radius conv_radius_geI_ex)
apply (meson less_ereal.simps(1) not_le order_trans)
apply (rule_tac x="of_real ra" in exI, simp)
apply (metis abs_of_nonneg add_diff_cancel_left' less_eq_real_def norm_of_real)
done
end
diff --git a/src/HOL/Analysis/Tagged_Division.thy b/src/HOL/Analysis/Tagged_Division.thy
--- a/src/HOL/Analysis/Tagged_Division.thy
+++ b/src/HOL/Analysis/Tagged_Division.thy
@@ -1,2602 +1,2602 @@
(* Title: HOL/Analysis/Tagged_Division.thy
Author: John Harrison
Author: Robert Himmelmann, TU Muenchen (Translation from HOL light); proofs reworked by LCP
*)
section \<open>Tagged Divisions for Henstock-Kurzweil Integration\<close>
theory Tagged_Division
imports Topology_Euclidean_Space
begin
lemma sum_Sigma_product:
assumes "finite S"
and "\<And>i. i \<in> S \<Longrightarrow> finite (T i)"
shows "(\<Sum>i\<in>S. sum (x i) (T i)) = (\<Sum>(i, j)\<in>Sigma S T. x i j)"
using assms
proof induct
case empty
then show ?case
by simp
next
case (insert a S)
show ?case
by (simp add: insert.hyps(1) insert.prems sum.Sigma)
qed
lemmas scaleR_simps = scaleR_zero_left scaleR_minus_left scaleR_left_diff_distrib
scaleR_zero_right scaleR_minus_right scaleR_right_diff_distrib scaleR_eq_0_iff
scaleR_cancel_left scaleR_cancel_right scaleR_add_right scaleR_add_left real_vector_class.scaleR_one
subsection\<^marker>\<open>tag unimportant\<close> \<open>Sundries\<close>
text\<open>A transitive relation is well-founded if all initial segments are finite.\<close>
lemma wf_finite_segments:
assumes "irrefl r" and "trans r" and "\<And>x. finite {y. (y, x) \<in> r}"
shows "wf (r)"
apply (simp add: trans_wf_iff wf_iff_acyclic_if_finite converse_def assms)
using acyclic_def assms irrefl_def trans_Restr by fastforce
text\<open>For creating values between \<^term>\<open>u\<close> and \<^term>\<open>v\<close>.\<close>
lemma scaling_mono:
fixes u::"'a::linordered_field"
assumes "u \<le> v" "0 \<le> r" "r \<le> s"
shows "u + r * (v - u) / s \<le> v"
proof -
have "r/s \<le> 1" using assms
using divide_le_eq_1 by fastforce
then have "(r/s) * (v - u) \<le> 1 * (v - u)"
by (meson diff_ge_0_iff_ge mult_right_mono \<open>u \<le> v\<close>)
then show ?thesis
by (simp add: field_simps)
qed
subsection \<open>Some useful lemmas about intervals\<close>
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 Int_interval_mixed_eq_empty[of c d a b] assms
unfolding 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 by auto
ultimately
show ?thesis
unfolding assms interior_cbox
by auto (metis IntI UnE empty_iff interior_maximal open_box subsetCE subsetI)
qed
lemma interior_Union_subset_cbox:
assumes "finite f"
assumes f: "\<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
qed
finally show ?thesis .
qed
lemma Int_interior_Union_intervals:
"\<lbrakk>finite \<F>; open S; \<And>T. T\<in>\<F> \<Longrightarrow> \<exists>a b. T = cbox a b; \<And>T. T\<in>\<F> \<Longrightarrow> S \<inter> (interior T) = {}\<rbrakk>
\<Longrightarrow> S \<inter> interior (\<Union>\<F>) = {}"
using interior_Union_subset_cbox[of \<F> "UNIV - S"] by auto
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"
using assms by (rule_tac set_eqI; auto simp: mem_box)+
lemma interval_not_empty: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow> cbox a b \<noteq> {}"
by (simp add: box_ne_empty)
subsection \<open>Bounds on intervals where they exist\<close>
definition\<^marker>\<open>tag important\<close> interval_upperbound :: "('a::euclidean_space) set \<Rightarrow> 'a"
where "interval_upperbound s = (\<Sum>i\<in>Basis. (SUP x\<in>s. x\<bullet>i) *\<^sub>R i)"
definition\<^marker>\<open>tag important\<close> interval_lowerbound :: "('a::euclidean_space) set \<Rightarrow> 'a"
where "interval_lowerbound s = (\<Sum>i\<in>Basis. (INF x\<in>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_sum 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_sum cbox_def
by (safe intro!: cInf_eq) auto
lemmas interval_bounds = interval_upperbound interval_lowerbound
lemma
fixes X::"real set"
shows interval_upperbound_real[simp]: "interval_upperbound X = Sup X"
and interval_lowerbound_real[simp]: "interval_lowerbound X = Inf X"
by (auto simp: interval_upperbound_def interval_lowerbound_def)
lemma interval_bounds'[simp]:
assumes "cbox a b \<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\<in>A \<times> B. x \<bullet> (i, 0)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (SUP x\<in>A. x \<bullet> i) *\<^sub>R i)"
by (subst (2) fst_image_times') (simp del: fst_image_times add: image_comp inner_Pair_0)
moreover from assms have snd_image_times': "B = snd ` (A \<times> B)" by simp
have "(\<Sum>i\<in>Basis. (SUP x\<in>A \<times> B. x \<bullet> (0, i)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (SUP x\<in>B. x \<bullet> i) *\<^sub>R i)"
by (subst (2) snd_image_times') (simp del: snd_image_times add: image_comp inner_Pair_0)
ultimately show ?thesis unfolding interval_upperbound_def
by (subst sum_Basis_prod_eq) (auto simp add: sum_prod)
qed
lemma interval_lowerbound_Times:
assumes "A \<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\<in>A \<times> B. x \<bullet> (i, 0)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (INF x\<in>A. x \<bullet> i) *\<^sub>R i)"
by (subst (2) fst_image_times') (simp del: fst_image_times add: image_comp inner_Pair_0)
moreover from assms have snd_image_times': "B = snd ` (A \<times> B)" by simp
have "(\<Sum>i\<in>Basis. (INF x\<in>A \<times> B. x \<bullet> (0, i)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (INF x\<in>B. x \<bullet> i) *\<^sub>R i)"
by (subst (2) snd_image_times') (simp del: snd_image_times add: image_comp inner_Pair_0)
ultimately show ?thesis unfolding interval_lowerbound_def
by (subst sum_Basis_prod_eq) (auto simp add: sum_prod)
qed
subsection \<open>The notion of a gauge --- simply an open set containing the point\<close>
definition\<^marker>\<open>tag important\<close> "gauge \<gamma> \<longleftrightarrow> (\<forall>x. x \<in> \<gamma> x \<and> open (\<gamma> x))"
lemma gaugeI:
assumes "\<And>x. x \<in> \<gamma> x"
and "\<And>x. open (\<gamma> x)"
shows "gauge \<gamma>"
using assms unfolding gauge_def by auto
lemma gaugeD[dest]:
assumes "gauge \<gamma>"
shows "x \<in> \<gamma> x"
and "open (\<gamma> 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_Int[intro]: "gauge \<gamma>1 \<Longrightarrow> gauge \<gamma>2 \<Longrightarrow> gauge (\<lambda>x. \<gamma>1 x \<inter> \<gamma>2 x)"
unfolding gauge_def by auto
lemma gauge_reflect:
fixes \<gamma> :: "'a::euclidean_space \<Rightarrow> 'a set"
shows "gauge \<gamma> \<Longrightarrow> gauge (\<lambda>x. uminus ` \<gamma> (- x))"
using equation_minus_iff
by (auto simp: gauge_def surj_def intro!: open_surjective_linear_image linear_uminus)
lemma gauge_Inter:
assumes "finite S"
and "\<And>u. u\<in>S \<Longrightarrow> gauge (f u)"
shows "gauge (\<lambda>x. \<Inter>{f u x | u. u \<in> S})"
proof -
have *: "\<And>x. {f u x |u. u \<in> S} = (\<lambda>u. f u x) ` S"
by auto
show ?thesis
unfolding gauge_def unfolding *
using assms unfolding Ball_def Inter_iff mem_Collect_eq gauge_def by auto
qed
lemma gauge_existence_lemma:
"(\<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>Attempt a systematic general set of "offset" results for components\<close>
(*FIXME Restructure, the subsection consists only of 1 lemma *)
lemma gauge_modify:
assumes "(\<forall>S. open S \<longrightarrow> open {x. f(x) \<in> S})" "gauge d"
shows "gauge (\<lambda>x. {y. f y \<in> d (f x)})"
using assms unfolding gauge_def by force
subsection \<open>Divisions\<close>
definition\<^marker>\<open>tag important\<close> 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"
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
done
}
ultimately show ?l
unfolding division_of_def cbox_sing by auto
next
assume ?l
have "x = {a}" if "x \<in> s" for x
by (metis \<open>s division_of cbox a a\<close> cbox_sing division_ofD(2) division_ofD(3) subset_singletonD that)
moreover have "s \<noteq> {}"
using \<open>s division_of cbox a a\<close> 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 cbox_division_memE:
assumes \<D>: "K \<in> \<D>" "\<D> division_of S"
obtains a b where "K = cbox a b" "K \<noteq> {}" "K \<subseteq> S"
using assms unfolding division_of_def by metis
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_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"
unfolding *
using division_ofD(6)[OF assms(1)] and division_ofD(6)[OF assms(2)] by auto
{
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 Int_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_Int:
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_Inter:
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_Int 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)"])
have ord: "a \<bullet> i \<le> c \<bullet> i" "c \<bullet> i \<le> d \<bullet> i" "d \<bullet> i \<le> b \<bullet> i" if "i \<in> Basis" for i
using incl nonempty that
unfolding box_eq_empty subset_box by (auto simp: not_le)
show "p division_of (cbox a b)"
proof (rule division_ofI)
show "finite p"
unfolding p_def by (auto intro!: finite_PiE)
{
fix k
assume "k \<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 fun_eq_iff)
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 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
proposition 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
have "q x division_of \<Union>(q x)" if x: "x \<in> p" for x
apply (rule division_ofI)
using division_ofD[OF q(1)[OF x]] by auto
then have di: "\<And>x. x \<in> p \<Longrightarrow> \<exists>d. d division_of \<Union>(q x - {x})"
by (meson Diff_subset division_of_subset)
have "\<exists>d. d division_of \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p)"
apply (rule elementary_Inter [OF finite_imageI[OF p(1)]])
apply (auto dest: di simp: False elementary_Inter [OF finite_imageI[OF p(1)]])
done
then obtain d where d: "d division_of \<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"
note qk=division_ofD[OF q(1)[OF K]]
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 Int_interior_Union_intervals])
show "\<And>T. T\<in>q K - {K} \<Longrightarrow> 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 by auto
qed (use qk in auto)} 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 Int_interior_Union_intervals)
apply (rule p open_interior ballI)+
apply simp_all
done
qed
then show ?thesis
by (meson Un_upper2 that)
qed
lemma elementary_bounded[dest]:
fixes S :: "'a::euclidean_space set"
shows "p division_of S \<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 bounded_subset_cbox_symmetric elementary_bounded)
proposition 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
with assms that show ?thesis by force
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 Int_interval by auto
have uv_sub: "cbox u v \<subseteq> cbox c d" using uv by auto
obtain p where pd: "p division_of cbox c d" and "cbox u v \<in> p"
by (rule partial_division_extend_1[OF uv_sub False[unfolded uv]])
note p = this division_ofD[OF pd]
have "interior (cbox a b \<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 Int_interior_Union_intervals)
using p(6) p(7)[OF p(2)] \<open>finite p\<close>
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_Union:
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 (cases "p = {} \<or> cbox a b = {}")
case True
obtain p where "p division_of (cbox a b)"
by (rule elementary_interval)
then show thesis
using True assms that by auto
next
case False
then have "p \<noteq> {}" "cbox a b \<noteq> {}"
by auto
note pdiv = division_ofD[OF assms]
show ?thesis
proof (cases "interior (cbox a b) = {}")
case True
show ?thesis
apply (rule that [of "insert (cbox a b) p", OF division_ofI])
using pdiv(1-4) True False apply auto
apply (metis IntI equals0D pdiv(5))
by (metis disjoint_iff_not_equal pdiv(5))
next
case False
have "\<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 pdiv(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 \<open>cbox a b \<noteq> {}\<close> 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 "*" pdiv(1) q(1) by auto
have "\<Union>?D = (\<Union>x\<in>p. \<Union>(insert (cbox a b) (q x)))"
by auto
also have "... = \<Union>{cbox a b \<union> t |t. t \<in> p}"
using q(6) by auto
also have "... = cbox a b \<union> \<Union>p"
using \<open>p \<noteq> {}\<close> by auto
finally show "\<Union>?D = cbox a b \<union> \<Union>p" .
show "K \<subseteq> cbox a b \<union> \<Union>p" "K \<noteq> {}" if "K \<in> ?D" for K
using q that by blast+
show "\<exists>a b. K = cbox a b" if "K \<in> ?D" for K
using q(4) that by auto
next
fix K' K
assume K: "K \<in> ?D" and 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' \<or> K = cbox a b \<or> K' = cbox a b")
case True
then show ?thesis
using True K' q(5) x' x by auto
next
case False
then have as': "K \<noteq> cbox a b" "K' \<noteq> cbox a b"
by auto
obtain c d where K: "K = cbox c d"
using q(4) x 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"
by (metis \<open>interior (cbox a b) \<noteq> {}\<close> 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 \<open>interior (cbox a b) \<noteq> {}\<close> c_d interior_subset_union_intervals q(2) x'(1) x'(2))
moreover have "interior x \<inter> interior x' = {}"
by (meson False assms division_ofD(5) x'(2) x(2))
ultimately show ?thesis
using \<open>interior K \<subseteq> interior x\<close> \<open>interior K' \<subseteq> interior x'\<close> by auto
qed
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_Int)
{
fix x
assume x: "x \<in> T" "x \<notin> S"
then obtain R where r: "R \<in> r1" "x \<in> R"
unfolding r1 using ab
by (meson division_contains r1(2) subsetCE)
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 Teq: "T = \<Union>p \<union> (\<Union>(r1 - p) \<inter> \<Union>q)"
unfolding divp divq using assms(3) by auto
have "interior S \<inter> interior (\<Union>(r1-p)) = {}"
proof (rule Int_interior_Union_intervals)
have *: "\<And>S. (\<And>x. x \<in> S \<Longrightarrow> False) \<Longrightarrow> S = {}"
by auto
show "interior S \<inter> interior m = {}" if "m \<in> r1 - p" for m
proof -
have "interior m \<inter> interior (\<Union>p) = {}"
proof (rule Int_interior_Union_intervals)
show "\<And>T. T\<in>p \<Longrightarrow> interior m \<inter> interior T = {}"
by (metis DiffD1 DiffD2 that r1(1) r1(7) rev_subsetD)
qed (use divp in auto)
then show "interior S \<inter> interior m = {}"
unfolding divp by auto
qed
qed (use r1 in auto)
then have "interior S \<inter> interior (\<Union>(r1-p) \<inter> (\<Union>q)) = {}"
using interior_subset by auto
then have div: "p \<union> r2 division_of \<Union>p \<union> \<Union>(r1 - p) \<inter> \<Union>q"
by (simp add: assms(1) division_disjoint_union divp(6) r2)
show ?thesis
apply (rule that[of "p \<union> r2"])
apply (auto simp: div Teq)
done
qed
lemma division_split:
fixes a :: "'a::euclidean_space"
assumes "p division_of (cbox a b)"
and k: "k\<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 obtain l where l: "K = l \<inter> {x. x \<bullet> k \<le> c}" "l \<in> p" "l \<inter> {x. x \<bullet> k \<le> c} \<noteq> {}"
by blast
obtain u v where uv: "l = cbox u v"
using assms(1) l(2) by blast
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 obtain l' where l': "K' = l' \<inter> {x. x \<bullet> k \<le> c}" "l' \<in> p" "l' \<inter> {x. x \<bullet> k \<le> c} \<noteq> {}"
by blast
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 obtain l where l: "K = l \<inter> {x. c \<le> x \<bullet> k}" "l \<in> p" "l \<inter> {x. c \<le> x \<bullet> k} \<noteq> {}"
by blast
obtain u v where uv: "l = cbox u v"
using l(2) p(4) by blast
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 obtain l' where l': "K' = l' \<inter> {x. c \<le> x \<bullet> k}" "l' \<in> p" "l' \<inter> {x. c \<le> x \<bullet> k} \<noteq> {}"
by blast
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\<^marker>\<open>tag important\<close> 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\<^marker>\<open>tag important\<close> 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 by fastforce
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 finite_subset[OF assms(2)]
unfolding tagged_partial_division_of_def
by blast
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_Un:
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_Union:
assumes "finite I"
and tag: "\<And>i. i\<in>I \<Longrightarrow> pfn i tagged_division_of i"
and disj: "\<And>i1 i2. \<lbrakk>i1 \<in> I; i2 \<in> I; i1 \<noteq> i2\<rbrakk> \<Longrightarrow> interior(i1) \<inter> interior(i2) = {}"
shows "\<Union>(pfn ` I) tagged_division_of (\<Union>I)"
proof (rule tagged_division_ofI)
note assm = tagged_division_ofD[OF tag]
show "finite (\<Union>(pfn ` I))"
using assms by auto
have "\<Union>{k. \<exists>x. (x, k) \<in> \<Union>(pfn ` I)} = \<Union>((\<lambda>i. \<Union>{k. \<exists>x. (x, k) \<in> pfn i}) ` I)"
by blast
also have "\<dots> = \<Union>I"
using assm(6) by auto
finally show "\<Union>{k. \<exists>x. (x, k) \<in> \<Union>(pfn ` I)} = \<Union>I" .
fix x k
assume xk: "(x, k) \<in> \<Union>(pfn ` I)"
then obtain i where i: "i \<in> I" "(x, k) \<in> pfn i"
by auto
show "x \<in> k" "\<exists>a b. k = cbox a b" "k \<subseteq> \<Union>I"
using assm(2-4)[OF i] using i(1) by auto
fix x' k'
assume xk': "(x', k') \<in> \<Union>(pfn ` I)" "(x, k) \<noteq> (x', k')"
then obtain i' where i': "i' \<in> I" "(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) disj interior_mono by blast
show "interior k \<inter> interior k' = {}"
proof (cases "i = i'")
case True then show ?thesis
using assm(5) i' i xk'(2) by blast
next
case False then show ?thesis
using "*" assm(3) i' i by auto
qed
qed
lemma tagged_partial_division_of_Union_self:
assumes "p tagged_partial_division_of s"
shows "p tagged_division_of (\<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
lemma tagged_division_Un_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_Un[OF assms(1-2)])
unfolding interval_split[OF k] interior_cbox
using k
apply (auto simp add: box_def elim!: ballE[where x=k])
done
qed
lemma tagged_division_Un_interval_real:
fixes a :: real
assumes "p1 tagged_division_of ({a .. b} \<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_Un_interval)
lemma tagged_division_split_left_inj:
assumes d: "d tagged_division_of i"
and tags: "(x1, K1) \<in> d" "(x2, K2) \<in> d"
and "K1 \<noteq> K2"
and eq: "K1 \<inter> {x. x \<bullet> k \<le> c} = K2 \<inter> {x. x \<bullet> k \<le> c}"
shows "interior (K1 \<inter> {x. x\<bullet>k \<le> c}) = {}"
proof -
have "interior (K1 \<inter> K2) = {} \<or> (x2, K2) = (x1, K1)"
using tags d by (metis (no_types) interior_Int tagged_division_ofD(5))
then show ?thesis
using eq \<open>K1 \<noteq> K2\<close> by (metis (no_types) inf_assoc inf_bot_left inf_left_idem interior_Int old.prod.inject)
qed
lemma tagged_division_split_right_inj:
assumes d: "d tagged_division_of i"
and tags: "(x1, K1) \<in> d" "(x2, K2) \<in> d"
and "K1 \<noteq> K2"
and eq: "K1 \<inter> {x. x\<bullet>k \<ge> c} = K2 \<inter> {x. x\<bullet>k \<ge> c}"
shows "interior (K1 \<inter> {x. x\<bullet>k \<ge> c}) = {}"
proof -
have "interior (K1 \<inter> K2) = {} \<or> (x2, K2) = (x1, K1)"
using tags d by (metis (no_types) interior_Int tagged_division_ofD(5))
then show ?thesis
using eq \<open>K1 \<noteq> K2\<close> by (metis (no_types) inf_assoc inf_bot_left inf_left_idem interior_Int old.prod.inject)
qed
lemma (in comm_monoid_set) over_tagged_division_lemma:
assumes "p tagged_division_of i"
and "\<And>u v. box u v = {} \<Longrightarrow> d (cbox u v) = \<^bold>1"
shows "F (\<lambda>(_, k). d k) p = F d (snd ` p)"
proof -
have *: "(\<lambda>(_ ,k). d k) = d \<circ> snd"
by (simp add: fun_eq_iff)
note assm = tagged_division_ofD[OF assms(1)]
show ?thesis
unfolding *
proof (rule reindex_nontrivial[symmetric])
show "finite p"
using assm by auto
fix x y
assume "x\<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"
using \<open>x \<in> p\<close> \<open>x \<noteq> y\<close> \<open>snd x = snd y\<close> [symmetric] by auto
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 "box a b = {}"
unfolding \<open>snd x = snd y\<close>[symmetric] ab 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
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
\<open>operative_division\<close>. Instances for the monoid are \<^typ>\<open>'a option\<close>, \<^typ>\<open>real\<close>, and
\<^typ>\<open>bool\<close>.\<close>
paragraph\<^marker>\<open>tag important\<close> \<open>Using additivity of lifted function to encode definedness.\<close>
text \<open>%whitespace\<close>
definition\<^marker>\<open>tag important\<close> 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\<^marker>\<open>tag important\<close> comm_monoid_lift_option:
assumes "comm_monoid f z"
shows "comm_monoid (lift_option f) (Some z)"
proof -
from assms interpret comm_monoid f z .
show ?thesis
by standard (auto simp: lift_option_def ac_simps split: bind_split)
qed
lemma comm_monoid_and: "comm_monoid HOL.conj True"
by standard auto
lemma comm_monoid_set_and: "comm_monoid_set HOL.conj True"
by (rule comm_monoid_set.intro) (fact comm_monoid_and)
paragraph \<open>Misc\<close>
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 bgauge_existence_lemma: "(\<forall>x\<in>s. \<exists>d::real. 0 < d \<and> q d x) \<longleftrightarrow> (\<forall>x. \<exists>d>0. x\<in>s \<longrightarrow> q d x)"
by (meson zero_less_one)
paragraph \<open>Division points\<close>
text \<open>%whitespace\<close>
definition\<^marker>\<open>tag important\<close> "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> \<not>(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
have "\<exists>i\<in>d. interval_lowerbound i \<bullet> x = y \<or> interval_upperbound i \<bullet> x = y"
if "a \<bullet> x < y" "y < (if x = k then c else b \<bullet> x)"
"interval_lowerbound i \<bullet> x = y \<or> interval_upperbound i \<bullet> x = y"
"i = l \<inter> {x. x \<bullet> k \<le> c}" "l \<in> d" "l \<inter> {x. x \<bullet> k \<le> c} \<noteq> {}"
"x \<in> Basis" for i l x y
proof -
obtain u v where l: "l = cbox u v"
using \<open>l \<in> d\<close> assms(1) by blast
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 that(6) unfolding l interval_split[OF k] box_ne_empty that .
have **: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i"
using l using that(6) unfolding box_ne_empty[symmetric] by auto
show ?thesis
apply (rule bexI[OF _ \<open>l \<in> d\<close>])
using that(1-3,5) \<open>x \<in> Basis\<close>
unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] that
apply (auto split: if_split_asm)
done
qed
moreover have "\<And>x y. \<lbrakk>y < (if x = k then c else b \<bullet> x)\<rbrakk> \<Longrightarrow> y < b \<bullet> x"
using \<open>c < b\<bullet>k\<close> by (auto split: if_split_asm)
ultimately show ?t1
unfolding division_points_def interval_split[OF k, of a b]
unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)]
unfolding * by force
have "\<And>x y i l. (if x = k then c else a \<bullet> x) < y \<Longrightarrow> a \<bullet> x < y"
using \<open>a\<bullet>k < c\<close> by (auto split: if_split_asm)
moreover have "\<exists>i\<in>d. interval_lowerbound i \<bullet> x = y \<or>
interval_upperbound i \<bullet> x = y"
if "(if x = k then c else a \<bullet> x) < y" "y < b \<bullet> x"
"interval_lowerbound i \<bullet> x = y \<or> interval_upperbound i \<bullet> x = y"
"i = l \<inter> {x. c \<le> x \<bullet> k}" "l \<in> d" "l \<inter> {x. c \<le> x \<bullet> k} \<noteq> {}"
"x \<in> Basis" for x y i l
proof -
obtain u v where l: "l = cbox u v"
using \<open>l \<in> d\<close> assm(4) by blast
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 that(6) unfolding l interval_split[OF k] box_ne_empty that .
have **: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i"
using l using that(6) unfolding box_ne_empty[symmetric] by auto
show "\<exists>i\<in>d. interval_lowerbound i \<bullet> x = y \<or> interval_upperbound i \<bullet> x = y"
apply (rule bexI[OF _ \<open>l \<in> d\<close>])
using that(1-3,5) \<open>x \<in> Basis\<close>
unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] that
apply (auto split: if_split_asm)
done
qed
ultimately show ?t2
unfolding division_points_def interval_split[OF k, of a b]
unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)]
unfolding *
by force
qed
lemma division_points_psubset:
fixes a :: "'a::euclidean_space"
assumes d: "d division_of (cbox a b)"
and altb: "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i" "a\<bullet>k < c" "c < b\<bullet>k"
and "l \<in> d"
and disj: "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 altb by (auto intro!:less_imp_le)
obtain u v where l: "l = cbox u v"
using d \<open>l \<in> d\<close> by blast
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"
apply (metis assms(5) box_ne_empty(1) cbox_division_memE d l)
by (metis assms(5) box_ne_empty(1) cbox_division_memE d l subset_box(1))
have *: "interval_upperbound (cbox a b \<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]
by (force simp add: *)
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]
by (force simp add: *)
moreover have "?D2 \<subseteq> ?D"
by (auto simp add: assms division_points_subset)
ultimately show "?D2 \<subset> ?D"
by blast
qed
lemma division_split_left_inj:
fixes S :: "'a::euclidean_space set"
assumes div: "\<D> division_of S"
and eq: "K1 \<inter> {x::'a. x\<bullet>k \<le> c} = K2 \<inter> {x. x\<bullet>k \<le> c}"
and "K1 \<in> \<D>" "K2 \<in> \<D>" "K1 \<noteq> K2"
shows "interior (K1 \<inter> {x. x\<bullet>k \<le> c}) = {}"
proof -
have "interior K2 \<inter> interior {a. a \<bullet> k \<le> c} = interior K1 \<inter> interior {a. a \<bullet> k \<le> c}"
by (metis (no_types) eq interior_Int)
moreover have "\<And>A. interior A \<inter> interior K2 = {} \<or> A = K2 \<or> A \<notin> \<D>"
by (meson div \<open>K2 \<in> \<D>\<close> division_of_def)
ultimately show ?thesis
using \<open>K1 \<in> \<D>\<close> \<open>K1 \<noteq> K2\<close> by auto
qed
lemma division_split_right_inj:
fixes S :: "'a::euclidean_space set"
assumes div: "\<D> division_of S"
and eq: "K1 \<inter> {x::'a. x\<bullet>k \<ge> c} = K2 \<inter> {x. x\<bullet>k \<ge> c}"
and "K1 \<in> \<D>" "K2 \<in> \<D>" "K1 \<noteq> K2"
shows "interior (K1 \<inter> {x. x\<bullet>k \<ge> c}) = {}"
proof -
have "interior K2 \<inter> interior {a. a \<bullet> k \<ge> c} = interior K1 \<inter> interior {a. a \<bullet> k \<ge> c}"
by (metis (no_types) eq interior_Int)
moreover have "\<And>A. interior A \<inter> interior K2 = {} \<or> A = K2 \<or> A \<notin> \<D>"
by (meson div \<open>K2 \<in> \<D>\<close> division_of_def)
ultimately show ?thesis
using \<open>K1 \<in> \<D>\<close> \<open>K1 \<noteq> K2\<close> by auto
qed
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] cong: image_cong_simp)
apply (rule equalityI)
apply blast
apply clarsimp
apply (rule_tac x="xa \<inter> {x. c + e \<ge> x \<bullet> k}" in exI)
apply auto
done
by (simp add: interval_split k interval_doublesplit)
qed
paragraph \<open>Operative\<close>
locale operative = comm_monoid_set +
fixes g :: "'b::euclidean_space set \<Rightarrow> 'a"
assumes box_empty_imp: "\<And>a b. box a b = {} \<Longrightarrow> g (cbox a b) = \<^bold>1"
and Basis_imp: "\<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})"
begin
lemma empty [simp]:
"g {} = \<^bold>1"
proof -
have *: "cbox One (-One) = ({}::'b set)"
by (auto simp: box_eq_empty inner_sum_left inner_Basis sum.If_cases ex_in_conv)
moreover have "box One (-One) = ({}::'b set)"
using box_subset_cbox[of One "-One"] by (auto simp: *)
ultimately show ?thesis
using box_empty_imp [of One "-One"] by simp
qed
lemma division:
"F g d = g (cbox a b)" if "d division_of (cbox a b)"
proof -
define C where [abs_def]: "C = card (division_points (cbox a b) d)"
then show ?thesis
using that proof (induction C arbitrary: a b d rule: less_induct)
case (less a b d)
show ?case
proof cases
assume "box a b = {}"
{ fix k assume "k\<in>d"
then obtain a' b' where k: "k = cbox a' b'"
using division_ofD(4)[OF less.prems] by blast
with \<open>k\<in>d\<close> division_ofD(2)[OF less.prems] have "cbox a' b' \<subseteq> cbox a b"
by auto
then have "box a' b' \<subseteq> box a b"
unfolding subset_box by auto
then have "g k = \<^bold>1"
using box_empty_imp [of a' b'] k by (simp add: \<open>box a b = {}\<close>) }
then show "box a b = {} \<Longrightarrow> F g d = g (cbox a b)"
by (auto intro!: neutral simp: box_empty_imp)
next
assume "box a b \<noteq> {}"
then have ab: "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i" and ab': "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
by (auto simp: box_ne_empty)
show "F g d = g (cbox a b)"
proof (cases "division_points (cbox a b) d = {}")
case True
{ fix u v and j :: 'b
assume j: "j \<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 simp: inner_simps)
note this[unfolded division_ofD(6)[OF \<open>d division_of cbox a b\<close>,symmetric] Union_iff]
then obtain i where i: "i \<in> d" "(1 / 2) *\<^sub>R (a + b) \<in> i" ..
obtain u v where uv: "i = cbox u v"
"\<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"
using d' i(1) by auto
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 obtain u v where uv: "x = cbox u v"
"\<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"
by blast
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 "box u v = {}"
using j unfolding box_eq_empty by (auto intro!: bexI[of _ j])
then show "g x = \<^bold>1"
unfolding uv(1) by (rule box_empty_imp)
qed
then show "F g d = g (cbox a b)"
using division_ofD[OF less.prems]
apply (subst deq)
apply (subst insert)
apply auto
done
next
case False
then have "\<exists>x. x \<in> division_points (cbox a b) d"
by auto
then obtain k c
where "k \<in> Basis" "interval_lowerbound (cbox a b) \<bullet> k < c"
"c < interval_upperbound (cbox a b) \<bullet> k"
"\<exists>i\<in>d. interval_lowerbound i \<bullet> k = c \<or> interval_upperbound i \<bullet> k = c"
unfolding division_points_def by auto
then obtain j where "a \<bullet> k < c" "c < b \<bullet> k"
and "j \<in> d" and j: "interval_lowerbound j \<bullet> k = c \<or> interval_upperbound j \<bullet> k = c"
by (metis division_of_trivial empty_iff interval_bounds' less.prems)
let ?lec = "{x. x\<bullet>k \<le> c}" let ?gec = "{x. x\<bullet>k \<ge> c}"
define d1 where "d1 = {l \<inter> ?lec | l. l \<in> d \<and> l \<inter> ?lec \<noteq> {}}"
define d2 where "d2 = {l \<inter> ?gec | l. l \<in> d \<and> l \<inter> ?gec \<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)"
have "division_points (cbox a b \<inter> ?lec) {l \<inter> ?lec |l. l \<in> d \<and> l \<inter> ?lec \<noteq> {}}
\<subset> division_points (cbox a b) d"
by (rule division_points_psubset[OF \<open>d division_of cbox a b\<close> ab \<open>a \<bullet> k < c\<close> \<open>c < b \<bullet> k\<close> \<open>j \<in> d\<close> j \<open>k \<in> Basis\<close>])
with division_points_finite[OF \<open>d division_of cbox a b\<close>]
have "card
(division_points (cbox a b \<inter> ?lec) {l \<inter> ?lec |l. l \<in> d \<and> l \<inter> ?lec \<noteq> {}})
< card (division_points (cbox a b) d)"
by (rule psubset_card_mono)
moreover have "division_points (cbox a b \<inter> {x. c \<le> x \<bullet> k}) {l \<inter> {x. c \<le> x \<bullet> k} |l. l \<in> d \<and> l \<inter> {x. c \<le> x \<bullet> k} \<noteq> {}}
\<subset> division_points (cbox a b) d"
by (rule division_points_psubset[OF \<open>d division_of cbox a b\<close> ab \<open>a \<bullet> k < c\<close> \<open>c < b \<bullet> k\<close> \<open>j \<in> d\<close> j \<open>k \<in> Basis\<close>])
with division_points_finite[OF \<open>d division_of cbox a b\<close>]
have "card (division_points (cbox a b \<inter> ?gec) {l \<inter> ?gec |l. l \<in> d \<and> l \<inter> ?gec \<noteq> {}})
< card (division_points (cbox a b) d)"
by (rule psubset_card_mono)
ultimately have *: "F g d1 = g (cbox a b \<inter> ?lec)" "F g d2 = g (cbox a b \<inter> ?gec)"
unfolding interval_split[OF \<open>k \<in> Basis\<close>]
apply (rule_tac[!] less.hyps)
using division_split[OF \<open>d division_of cbox a b\<close>, where k=k and c=c] \<open>k \<in> Basis\<close>
by (simp_all add: interval_split d1_def d2_def division_points_finite[OF \<open>d division_of cbox a b\<close>])
have fxk_le: "g (l \<inter> ?lec) = \<^bold>1"
if "l \<in> d" "y \<in> d" "l \<inter> ?lec = y \<inter> ?lec" "l \<noteq> y" for l y
proof -
obtain u v where leq: "l = cbox u v"
using \<open>l \<in> d\<close> less.prems by auto
have "interior (cbox u v \<inter> ?lec) = {}"
using that division_split_left_inj leq less.prems by blast
then show ?thesis
unfolding leq interval_split [OF \<open>k \<in> Basis\<close>]
by (auto intro: box_empty_imp)
qed
have fxk_ge: "g (l \<inter> {x. x \<bullet> k \<ge> c}) = \<^bold>1"
if "l \<in> d" "y \<in> d" "l \<inter> ?gec = y \<inter> ?gec" "l \<noteq> y" for l y
proof -
obtain u v where leq: "l = cbox u v"
using \<open>l \<in> d\<close> less.prems by auto
have "interior (cbox u v \<inter> ?gec) = {}"
using that division_split_right_inj leq less.prems by blast
then show ?thesis
unfolding leq interval_split[OF \<open>k \<in> Basis\<close>]
by (auto intro: box_empty_imp)
qed
have d1_alt: "d1 = (\<lambda>l. l \<inter> ?lec) ` {l \<in> d. l \<inter> ?lec \<noteq> {}}"
using d1_def by auto
have d2_alt: "d2 = (\<lambda>l. l \<inter> ?gec) ` {l \<in> d. l \<inter> ?gec \<noteq> {}}"
using d2_def by auto
have "g (cbox a b) = F g d1 \<^bold>* F g d2" (is "_ = ?prev")
unfolding * using \<open>k \<in> Basis\<close>
by (auto dest: Basis_imp)
also have "F g d1 = F (\<lambda>l. g (l \<inter> ?lec)) d"
unfolding d1_alt using division_of_finite[OF less.prems] fxk_le
by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left)
also have "F g d2 = F (\<lambda>l. g (l \<inter> ?gec)) d"
unfolding d2_alt using division_of_finite[OF less.prems] fxk_ge
by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left)
also have *: "\<forall>x\<in>d. g x = g (x \<inter> ?lec) \<^bold>* g (x \<inter> ?gec)"
unfolding forall_in_division[OF \<open>d division_of cbox a b\<close>]
using \<open>k \<in> Basis\<close>
by (auto dest: Basis_imp)
have "F (\<lambda>l. g (l \<inter> ?lec)) d \<^bold>* F (\<lambda>l. g (l \<inter> ?gec)) d = F g d"
using * by (simp add: distrib)
finally show ?thesis by auto
qed
qed
qed
qed
proposition tagged_division:
assumes "d tagged_division_of (cbox a b)"
shows "F (\<lambda>(_, l). g l) d = g (cbox a b)"
proof -
have "F (\<lambda>(_, k). g k) d = F g (snd ` d)"
using assms box_empty_imp by (rule over_tagged_division_lemma)
then show ?thesis
unfolding assms [THEN division_of_tagged_division, THEN division] .
qed
end
locale operative_real = comm_monoid_set +
fixes g :: "real set \<Rightarrow> 'a"
assumes neutral: "b \<le> a \<Longrightarrow> g {a..b} = \<^bold>1"
assumes coalesce_less: "a < c \<Longrightarrow> c < b \<Longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}"
begin
sublocale operative where g = g
rewrites "box = (greaterThanLessThan :: real \<Rightarrow> _)"
and "cbox = (atLeastAtMost :: real \<Rightarrow> _)"
and "\<And>x::real. x \<in> Basis \<longleftrightarrow> x = 1"
proof -
show "operative f z g"
proof
show "g (cbox a b) = \<^bold>1" if "box a b = {}" for a b
using that by (simp add: neutral)
show "g (cbox a b) = g (cbox a b \<inter> {x. x \<bullet> k \<le> c}) \<^bold>* g (cbox a b \<inter> {x. c \<le> x \<bullet> k})"
if "k \<in> Basis" for a b c k
proof -
from that have [simp]: "k = 1"
by simp
from neutral [of 0 1] neutral [of a a for a] coalesce_less
have [simp]: "g {} = \<^bold>1" "\<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
have "g {a..b} = g {a..min b c} \<^bold>* g {max a c..b}"
by (auto simp: min_def max_def le_less)
then show "g (cbox a b) = g (cbox a b \<inter> {x. x \<bullet> k \<le> c}) \<^bold>* g (cbox a b \<inter> {x. c \<le> x \<bullet> k})"
by (simp add: atMost_def [symmetric] atLeast_def [symmetric])
qed
qed
show "box = (greaterThanLessThan :: real \<Rightarrow> _)"
and "cbox = (atLeastAtMost :: real \<Rightarrow> _)"
and "\<And>x::real. x \<in> Basis \<longleftrightarrow> x = 1"
by (simp_all add: fun_eq_iff)
qed
lemma coalesce_less_eq:
"g {a..c} \<^bold>* g {c..b} = g {a..b}" if "a \<le> c" "c \<le> b"
proof (cases "c = a \<or> c = b")
case False
with that have "a < c" "c < b"
by auto
then show ?thesis
by (rule coalesce_less)
next
case True
with that box_empty_imp [of a a] box_empty_imp [of b b] show ?thesis
by safe simp_all
qed
end
lemma operative_realI:
"operative_real f z g" if "operative f z g"
proof -
interpret operative f z g
using that .
show ?thesis
proof
show "g {a..b} = z" if "b \<le> a" for a b
using that box_empty_imp by simp
show "f (g {a..c}) (g {c..b}) = g {a..b}" if "a < c" "c < b" for a b c
using that
using Basis_imp [of 1 a b c]
by (simp_all add: atMost_def [symmetric] atLeast_def [symmetric] max_def min_def)
qed
qed
subsection \<open>Special case of additivity we need for the FTC\<close>
(*fix add explanation here *)
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 "sum (\<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))"
interpret operative_real plus 0 ?f
rewrites "comm_monoid_set.F (+) 0 = sum"
by standard[1] (auto simp add: sum_def)
have p_td: "p tagged_division_of cbox a b"
using assms(2) box_real(2) by presburger
have **: "cbox a b \<noteq> {}"
using assms(1) by auto
then have "f b - f a = (\<Sum>(x, l)\<in>p. if l = {} then 0 else f (interval_upperbound l) - f (interval_lowerbound l))"
proof -
have "(if cbox a b = {} then 0 else f (interval_upperbound (cbox a b)) - f (interval_lowerbound (cbox a b))) = f b - f a"
using assms by auto
then show ?thesis
using p_td assms by (simp add: tagged_division)
qed
then show ?thesis
using assms by (auto intro!: sum.cong)
qed
subsection \<open>Fine-ness of a partition w.r.t. a gauge\<close>
definition\<^marker>\<open>tag important\<close> 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_Int: "(\<lambda>x. d1 x \<inter> d2 x) fine p \<longleftrightarrow> d1 fine p \<and> d2 fine p"
unfolding fine_def by auto
lemma fine_Inter:
"(\<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_Un: "d fine p1 \<Longrightarrow> d fine p2 \<Longrightarrow> d fine (p1 \<union> p2)"
unfolding fine_def by blast
lemma fine_Union: "(\<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>Some basic combining lemmas\<close>
lemma tagged_division_Union_exists:
assumes "finite I"
and "\<forall>i\<in>I. \<exists>p. p tagged_division_of i \<and> d fine p"
and "\<forall>i1\<in>I. \<forall>i2\<in>I. i1 \<noteq> i2 \<longrightarrow> interior i1 \<inter> interior i2 = {}"
and "\<Union>I = i"
obtains p where "p tagged_division_of i" and "d fine p"
proof -
obtain pfn where pfn:
"\<And>x. x \<in> I \<Longrightarrow> pfn x tagged_division_of x"
"\<And>x. x \<in> I \<Longrightarrow> d fine pfn x"
using bchoice[OF assms(2)] by auto
show thesis
apply (rule_tac p="\<Union>(pfn ` I)" in that)
using assms(1) assms(3) assms(4) pfn(1) tagged_division_Union apply force
by (metis (mono_tags, lifting) fine_Union imageE pfn(2))
qed
(* FIXME structure here, do not have one lemma in each subsection *)
subsection\<^marker>\<open>tag unimportant\<close> \<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
(* FIXME structure here, do not have one lemma in each subsection *)
subsection \<open>General bisection principle for intervals; might be useful elsewhere\<close>
(* FIXME move? *)
lemma interval_bisection_step:
fixes type :: "'a::euclidean_space"
assumes emp: "P {}"
and Un: "\<And>S T. \<lbrakk>P S; P T; interior(S) \<inter> interior(T) = {}\<rbrakk> \<Longrightarrow> P (S \<union> T)"
and non: "\<not> P (cbox a (b::'a))"
obtains c d where "\<not> P (cbox c d)"
and "\<And>i. i \<in> Basis \<Longrightarrow> 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 emp non by metis
then have ab: "\<And>i. i\<in>Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i"
by (force simp: mem_box)
have UN_cases: "\<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>)" for \<F>
proof (induct \<F> rule: finite_induct)
case empty show ?case
using emp by auto
next
case (insert x f)
then show ?case
unfolding Union_insert by (metis Int_interior_Union_intervals Un insert_iff open_interior)
qed
let ?ab = "\<lambda>i. (a\<bullet>i + b\<bullet>i) / 2"
let ?A = "{cbox c d | c d::'a. \<forall>i\<in>Basis. (c\<bullet>i = a\<bullet>i) \<and> (d\<bullet>i = ?ab i) \<or>
(c\<bullet>i = ?ab i) \<and> (d\<bullet>i = b\<bullet>i)}"
have "P (\<Union>?A)"
if "\<And>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 \<Longrightarrow> P (cbox c d)"
proof (rule UN_cases)
let ?B = "(\<lambda>S. cbox (\<Sum>i\<in>Basis. (if i \<in> S then a\<bullet>i else ?ab i) *\<^sub>R i::'a)
(\<Sum>i\<in>Basis. (if i \<in> S then ?ab i 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 = ?ab i \<or> c \<bullet> i = ?ab i \<and> d \<bullet> i = b \<bullet> i"
by blast
have "c = (\<Sum>i\<in>Basis. (if c \<bullet> i = a \<bullet> i then a \<bullet> i else ?ab i) *\<^sub>R i)"
"d = (\<Sum>i\<in>Basis. (if c \<bullet> i = a \<bullet> i then ?ab i else b \<bullet> i) *\<^sub>R i)"
using x(2) ab by (fastforce simp add: euclidean_eq_iff[where 'a='a])+
then show "x \<in> ?B"
unfolding x by (rule_tac x="{i. i\<in>Basis \<and> c\<bullet>i = a\<bullet>i}" in image_eqI) auto
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 = ?ab i \<or> c \<bullet> i = ?ab i \<and> d \<bullet> i = b \<bullet> i"
by blast
show "P S"
unfolding s using ab s(2) by (fastforce intro!: that)
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 = ?ab i \<or> e \<bullet> i = ?ab i \<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'] t(2)[OF i'] and i x
by auto
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 = ?ab i \<or> c \<bullet> i = ?ab i \<and> d \<bullet> i = b \<bullet> i"
by blast
then show "x\<in>cbox a b"
unfolding mem_box by force
next
fix x
assume x: "x \<in> cbox a b"
then have "\<forall>i\<in>Basis. \<exists>c d. (c = a\<bullet>i \<and> d = ?ab i \<or> c = ?ab i \<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 by (metis linear)
then obtain \<alpha> \<beta> where "\<forall>i\<in>Basis. (\<alpha> \<bullet> i = a \<bullet> i \<and> \<beta> \<bullet> i = ?ab i \<or>
\<alpha> \<bullet> i = ?ab i \<and> \<beta> \<bullet> i = b \<bullet> i) \<and> \<alpha> \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> \<beta> \<bullet> i"
by (auto simp: choice_Basis_iff)
then show "x\<in>\<Union>?A"
by (force simp add: mem_box)
qed
finally show thesis
by (metis (no_types, lifting) assms(3) that)
qed
lemma interval_bisection:
fixes type :: "'a::euclidean_space"
assumes "P {}"
and Un: "\<And>S T. \<lbrakk>P S; P T; interior(S) \<inter> interior(T) = {}\<rbrakk> \<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 False
obtain c d where "\<not> P (cbox c d)"
"\<And>i. i \<in> Basis \<Longrightarrow>
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 (blast intro: interval_bisection_step[of P, OF assms(1-2) False])
then show ?thesis
by (rule_tac x="(c,d)" in exI) auto
qed
qed
then obtain f where f:
"\<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)" by metis
define AB A B where ab_def: "AB n = (f ^^ n) (a,b)" "A n = fst(AB n)" "B n = snd(AB n)" for n
have [simp]: "A 0 = a" "B 0 = b" and ABRAW: "\<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 using \<open>\<not> P (cbox a b)\<close> f by auto
next
case (Suc n)
show ?case
unfolding S
apply (rule f[rule_format])
using Suc
unfolding S
apply auto
done
qed
qed
then have AB: "A(n)\<bullet>i \<le> A(Suc n)\<bullet>i" "A(Suc n)\<bullet>i \<le> B(Suc n)\<bullet>i"
"B(Suc n)\<bullet>i \<le> B(n)\<bullet>i" "2 * (B(Suc n)\<bullet>i - A(Suc n)\<bullet>i) \<le> B(n)\<bullet>i - A(n)\<bullet>i"
if "i\<in>Basis" for i n
using that by blast+
have notPAB: "\<not> P (cbox (A(Suc n)) (B(Suc n)))" for n
using ABRAW by blast
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 "(sum (\<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> sum (\<lambda>i. \<bar>(x - y)\<bullet>i\<bar>) Basis"
unfolding dist_norm by(rule norm_le_l1)
also have "\<dots> \<le> sum (\<lambda>i. B n\<bullet>i - A n\<bullet>i) Basis"
proof (rule sum_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> sum (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis / 2^n"
unfolding sum_divide_distrib
proof (rule sum_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(3) that 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 by (intro order_trans[OF step.IH] subset_box_imp) auto
qed simp
} note ABsubset = this
have "\<And>n. cbox (A n) (B n) \<noteq> {}"
by (meson AB dual_order.trans interval_not_empty)
then obtain x0 where x0: "\<And>n. x0 \<in> cbox (A n) (B n)"
using decreasing_closed_nest [OF closed_cbox] ABsubset interv by blast
show thesis
proof (rule that[rule_format, of x0])
show "x0\<in>cbox a b"
using \<open>A 0 = a\<close> \<open>B 0 = b\<close> x0 by blast
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))"
proof (cases "0 < n")
case True then show ?thesis
by (metis Suc_pred' notPAB)
next
case False then show ?thesis
using \<open>A 0 = a\<close> \<open>B 0 = b\<close> \<open>\<not> P (cbox a b)\<close> by blast
qed
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"
using ABsubset \<open>A 0 = a\<close> \<open>B 0 = b\<close> by blast
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: (*FIXME rename(?) *)
fixes a b :: "'a::euclidean_space"
assumes "gauge g"
obtains p where "p tagged_division_of (cbox a b)" "g fine p"
proof (cases "\<exists>p. p tagged_division_of (cbox a b) \<and> g fine p")
case True
then show ?thesis
using that by auto
next
case False
assume "\<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 _ _ False])
apply (simp add: fine_def)
apply (metis tagged_division_Un fine_Un)
apply auto
done
obtain e where e: "e > 0" "ball x e \<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 ?thesis
using tagged_division_of_self[OF c_d(1)] using c_d by auto
qed
lemma fine_division_exists_real:
fixes a b :: real
assumes "gauge g"
obtains p where "p tagged_division_of {a .. b}" "g fine p"
by (metis assms box_real(2) fine_division_exists)
subsection \<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 ptag: "p tagged_division_of (cbox a b)"
and "gauge d"
obtains q where "q tagged_division_of (cbox a b)"
and "d fine q"
and "\<forall>(x,k) \<in> p. k \<subseteq> d(x) \<longrightarrow> (x,k) \<in> q"
proof -
have p: "finite p" "p tagged_partial_division_of (cbox a b)"
using ptag unfolding tagged_division_of_def by auto
have "(\<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))"
if "finite p" "p tagged_partial_division_of (cbox a b)" "gauge d" for p
using that
proof (induct p)
case empty
show ?case
by (force simp add: fine_def)
next
case (insert xk p)
obtain x k where xk: "xk = (x, k)"
using surj_pair[of xk] by metis
obtain q1 where q1: "q1 tagged_division_of \<Union>{k. \<exists>x. (x, k) \<in> p}"
and "d fine q1"
and q1I: "\<And>x k. \<lbrakk>(x, k)\<in>p; k \<subseteq> d x\<rbrakk> \<Longrightarrow> (x, k) \<in> q1"
using case_prodD tagged_partial_division_subset[OF insert(4) subset_insertI]
by (metis (mono_tags, lifting) insert.hyps(3) insert.prems(2))
have *: "\<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)]
obtain u v where uv: "k = cbox u v"
using p(4) xk by blast
have "finite {k. \<exists>x. (x, k) \<in> p}"
apply (rule finite_subset[of _ "snd ` p"])
using image_iff apply fastforce
using insert.hyps(1) by blast
then have int: "interior (cbox u v) \<inter> interior (\<Union>{k. \<exists>x. (x, k) \<in> p}) = {}"
proof (rule Int_interior_Union_intervals)
show "open (interior (cbox u v))"
by auto
show "\<And>T. T \<in> {k. \<exists>x. (x, k) \<in> p} \<Longrightarrow> \<exists>a b. T = cbox a b"
using p(4) by auto
show "\<And>T. T \<in> {k. \<exists>x. (x, k) \<in> p} \<Longrightarrow> interior (cbox u v) \<inter> interior T = {}"
by clarify (metis insert.hyps(2) insert_iff interior_cbox p(5) uv xk)
qed
show ?case
proof (cases "cbox u v \<subseteq> d x")
case True
have "{(x, cbox u v)} tagged_division_of cbox u v"
by (simp add: p(2) uv xk tagged_division_of_self)
then have "{(x, cbox u v)} \<union> q1 tagged_division_of \<Union>{k. \<exists>x. (x, k) \<in> insert xk p}"
unfolding * uv by (metis (no_types, lifting) int q1 tagged_division_Un)
with True show ?thesis
apply (rule_tac x="{(x,cbox u v)} \<union> q1" in exI)
using \<open>d fine q1\<close> fine_def q1I uv xk apply fastforce
done
next
case False
obtain q2 where q2: "q2 tagged_division_of cbox u v" "d fine q2"
using fine_division_exists[OF assms(2)] by blast
show ?thesis
apply (rule_tac x="q2 \<union> q1" in exI)
apply (intro conjI)
unfolding * uv
apply (rule tagged_division_Un q2 q1 int fine_Un)+
apply (auto intro: q1 q2 fine_Un \<open>d fine q1\<close> simp add: False q1I uv xk)
done
qed
qed
with p obtain q where q: "q tagged_division_of \<Union>{k. \<exists>x. (x, k) \<in> p}" "d fine q" "\<forall>(x, k)\<in>p. k \<subseteq> d x \<longrightarrow> (x, k) \<in> q"
by (meson \<open>gauge d\<close>)
with ptag that show ?thesis by auto
qed
subsubsection\<^marker>\<open>tag important\<close> \<open>Covering lemma\<close>
text\<open> Some technical lemmas used in the approximation results that follow. Proof of the covering
lemma is an obvious multidimensional generalization of Lemma 3, p65 of Swartz's
"Introduction to Gauge Integrals". \<close>
proposition covering_lemma:
assumes "S \<subseteq> cbox a b" "box a b \<noteq> {}" "gauge g"
obtains \<D> where
"countable \<D>" "\<Union>\<D> \<subseteq> cbox a b"
"\<And>K. K \<in> \<D> \<Longrightarrow> interior K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
"pairwise (\<lambda>A B. interior A \<inter> interior B = {}) \<D>"
"\<And>K. K \<in> \<D> \<Longrightarrow> \<exists>x \<in> S \<inter> K. K \<subseteq> g x"
"\<And>u v. cbox u v \<in> \<D> \<Longrightarrow> \<exists>n. \<forall>i \<in> Basis. v \<bullet> i - u \<bullet> i = (b \<bullet> i - a \<bullet> i) / 2^n"
"S \<subseteq> \<Union>\<D>"
proof -
have aibi: "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i" and normab: "0 < norm(b - a)"
using \<open>box a b \<noteq> {}\<close> box_eq_empty box_sing by fastforce+
let ?K0 = "\<lambda>(n, f::'a\<Rightarrow>nat).
cbox (\<Sum>i \<in> Basis. (a \<bullet> i + (f i / 2^n) * (b \<bullet> i - a \<bullet> i)) *\<^sub>R i)
(\<Sum>i \<in> Basis. (a \<bullet> i + ((f i + 1) / 2^n) * (b \<bullet> i - a \<bullet> i)) *\<^sub>R i)"
let ?D0 = "?K0 ` (SIGMA n:UNIV. Pi\<^sub>E Basis (\<lambda>i::'a. lessThan (2^n)))"
obtain \<D>0 where count: "countable \<D>0"
and sub: "\<Union>\<D>0 \<subseteq> cbox a b"
and int: "\<And>K. K \<in> \<D>0 \<Longrightarrow> (interior K \<noteq> {}) \<and> (\<exists>c d. K = cbox c d)"
and intdj: "\<And>A B. \<lbrakk>A \<in> \<D>0; B \<in> \<D>0\<rbrakk> \<Longrightarrow> A \<subseteq> B \<or> B \<subseteq> A \<or> interior A \<inter> interior B = {}"
and SK: "\<And>x. x \<in> S \<Longrightarrow> \<exists>K \<in> \<D>0. x \<in> K \<and> K \<subseteq> g x"
and cbox: "\<And>u v. cbox u v \<in> \<D>0 \<Longrightarrow> \<exists>n. \<forall>i \<in> Basis. v \<bullet> i - u \<bullet> i = (b \<bullet> i - a \<bullet> i) / 2^n"
and fin: "\<And>K. K \<in> \<D>0 \<Longrightarrow> finite {L \<in> \<D>0. K \<subseteq> L}"
proof
show "countable ?D0"
by (simp add: countable_PiE)
next
show "\<Union>?D0 \<subseteq> cbox a b"
apply (simp add: UN_subset_iff)
apply (intro conjI allI ballI subset_box_imp)
apply (simp add: field_simps)
apply (auto intro: mult_right_mono aibi)
apply (force simp: aibi scaling_mono nat_less_real_le dest: PiE_mem intro: mult_right_mono)
done
next
show "\<And>K. K \<in> ?D0 \<Longrightarrow> interior K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
using \<open>box a b \<noteq> {}\<close>
by (clarsimp simp: box_eq_empty) (fastforce simp add: field_split_simps dest: PiE_mem)
next
have realff: "(real w) * 2^m < (real v) * 2^n \<longleftrightarrow> w * 2^m < v * 2^n" for m n v w
using of_nat_less_iff less_imp_of_nat_less by fastforce
have *: "\<forall>v w. ?K0(m,v) \<subseteq> ?K0(n,w) \<or> ?K0(n,w) \<subseteq> ?K0(m,v) \<or> interior(?K0(m,v)) \<inter> interior(?K0(n,w)) = {}"
for m n \<comment> \<open>The symmetry argument requires a single HOL formula\<close>
proof (rule linorder_wlog [where a=m and b=n], intro allI impI)
fix v w m and n::nat
assume "m \<le> n" \<comment> \<open>WLOG we can assume \<^term>\<open>m \<le> n\<close>, when the first disjunct becomes impossible\<close>
have "?K0(n,w) \<subseteq> ?K0(m,v) \<or> interior(?K0(m,v)) \<inter> interior(?K0(n,w)) = {}"
apply (simp add: subset_box disjoint_interval)
apply (rule ccontr)
apply (clarsimp simp add: aibi mult_le_cancel_right divide_le_cancel not_less not_le)
apply (drule_tac x=i in bspec, assumption)
using \<open>m\<le>n\<close> realff [of _ _ "1+_"] realff [of "1+_"_ "1+_"]
apply (auto simp: divide_simps add.commute not_le nat_le_iff_add realff)
apply (simp_all add: power_add)
apply (metis (no_types, hide_lams) mult_Suc mult_less_cancel2 not_less_eq mult.assoc)
apply (metis (no_types, hide_lams) mult_Suc mult_less_cancel2 not_less_eq mult.assoc)
done
then show "?K0(m,v) \<subseteq> ?K0(n,w) \<or> ?K0(n,w) \<subseteq> ?K0(m,v) \<or> interior(?K0(m,v)) \<inter> interior(?K0(n,w)) = {}"
by meson
qed auto
show "\<And>A B. \<lbrakk>A \<in> ?D0; B \<in> ?D0\<rbrakk> \<Longrightarrow> A \<subseteq> B \<or> B \<subseteq> A \<or> interior A \<inter> interior B = {}"
apply (erule imageE SigmaE)+
using * by simp
next
show "\<exists>K \<in> ?D0. x \<in> K \<and> K \<subseteq> g x" if "x \<in> S" for x
proof (simp only: bex_simps split_paired_Bex_Sigma)
show "\<exists>n. \<exists>f \<in> Basis \<rightarrow>\<^sub>E {..<2 ^ n}. x \<in> ?K0(n,f) \<and> ?K0(n,f) \<subseteq> g x"
proof -
obtain e where "0 < e"
and e: "\<And>y. (\<And>i. i \<in> Basis \<Longrightarrow> \<bar>x \<bullet> i - y \<bullet> i\<bar> \<le> e) \<Longrightarrow> y \<in> g x"
proof -
have "x \<in> g x" "open (g x)"
using \<open>gauge g\<close> by (auto simp: gauge_def)
then obtain \<epsilon> where "0 < \<epsilon>" and \<epsilon>: "ball x \<epsilon> \<subseteq> g x"
using openE by blast
have "norm (x - y) < \<epsilon>"
if "(\<And>i. i \<in> Basis \<Longrightarrow> \<bar>x \<bullet> i - y \<bullet> i\<bar> \<le> \<epsilon> / (2 * real DIM('a)))" for y
proof -
have "norm (x - y) \<le> (\<Sum>i\<in>Basis. \<bar>x \<bullet> i - y \<bullet> i\<bar>)"
by (metis (no_types, lifting) inner_diff_left norm_le_l1 sum.cong)
also have "... \<le> DIM('a) * (\<epsilon> / (2 * real DIM('a)))"
by (meson sum_bounded_above that)
also have "... = \<epsilon> / 2"
by (simp add: field_split_simps)
also have "... < \<epsilon>"
by (simp add: \<open>0 < \<epsilon>\<close>)
finally show ?thesis .
qed
then show ?thesis
by (rule_tac e = "\<epsilon> / 2 / DIM('a)" in that) (simp_all add: \<open>0 < \<epsilon>\<close> dist_norm subsetD [OF \<epsilon>])
qed
have xab: "x \<in> cbox a b"
using \<open>x \<in> S\<close> \<open>S \<subseteq> cbox a b\<close> by blast
obtain n where n: "norm (b - a) / 2^n < e"
using real_arch_pow_inv [of "e / norm(b - a)" "1/2"] normab \<open>0 < e\<close>
by (auto simp: field_split_simps)
then have "norm (b - a) < e * 2^n"
by (auto simp: field_split_simps)
then have bai: "b \<bullet> i - a \<bullet> i < e * 2 ^ n" if "i \<in> Basis" for i
proof -
have "b \<bullet> i - a \<bullet> i \<le> norm (b - a)"
by (metis abs_of_nonneg dual_order.trans inner_diff_left linear norm_ge_zero Basis_le_norm that)
also have "... < e * 2 ^ n"
using \<open>norm (b - a) < e * 2 ^ n\<close> by blast
finally show ?thesis .
qed
have D: "(a + n \<le> x \<and> x \<le> a + m) \<Longrightarrow> (a + n \<le> y \<and> y \<le> a + m) \<Longrightarrow> abs(x - y) \<le> m - n"
for a m n x and y::real
by auto
have "\<forall>i\<in>Basis. \<exists>k<2 ^ n. (a \<bullet> i + real k * (b \<bullet> i - a \<bullet> i) / 2 ^ n \<le> x \<bullet> i \<and>
x \<bullet> i \<le> a \<bullet> i + (real k + 1) * (b \<bullet> i - a \<bullet> i) / 2 ^ n)"
proof
fix i::'a assume "i \<in> Basis"
consider "x \<bullet> i = b \<bullet> i" | "x \<bullet> i < b \<bullet> i"
using \<open>i \<in> Basis\<close> mem_box(2) xab by force
then show "\<exists>k<2 ^ n. (a \<bullet> i + real k * (b \<bullet> i - a \<bullet> i) / 2 ^ n \<le> x \<bullet> i \<and>
x \<bullet> i \<le> a \<bullet> i + (real k + 1) * (b \<bullet> i - a \<bullet> i) / 2 ^ n)"
proof cases
case 1 then show ?thesis
by (rule_tac x = "2^n - 1" in exI) (auto simp: algebra_simps field_split_simps of_nat_diff \<open>i \<in> Basis\<close> aibi)
next
case 2
then have abi_less: "a \<bullet> i < b \<bullet> i"
using \<open>i \<in> Basis\<close> xab by (auto simp: mem_box)
let ?k = "nat \<lfloor>2 ^ n * (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i)\<rfloor>"
show ?thesis
proof (intro exI conjI)
show "?k < 2 ^ n"
using aibi xab \<open>i \<in> Basis\<close>
by (force simp: nat_less_iff floor_less_iff field_split_simps 2 mem_box)
next
have "a \<bullet> i + real ?k * (b \<bullet> i - a \<bullet> i) / 2 ^ n \<le>
a \<bullet> i + (2 ^ n * (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i)) * (b \<bullet> i - a \<bullet> i) / 2 ^ n"
apply (intro add_left_mono mult_right_mono divide_right_mono of_nat_floor)
using aibi [OF \<open>i \<in> Basis\<close>] xab 2
apply (simp_all add: \<open>i \<in> Basis\<close> mem_box field_split_simps)
done
also have "... = x \<bullet> i"
using abi_less by (simp add: field_split_simps)
finally show "a \<bullet> i + real ?k * (b \<bullet> i - a \<bullet> i) / 2 ^ n \<le> x \<bullet> i" .
next
have "x \<bullet> i \<le> a \<bullet> i + (2 ^ n * (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i)) * (b \<bullet> i - a \<bullet> i) / 2 ^ n"
- using abi_less by (simp add: field_split_simps algebra_simps)
+ using abi_less by (simp add: field_split_simps)
also have "... \<le> a \<bullet> i + (real ?k + 1) * (b \<bullet> i - a \<bullet> i) / 2 ^ n"
apply (intro add_left_mono mult_right_mono divide_right_mono of_nat_floor)
using aibi [OF \<open>i \<in> Basis\<close>] xab
apply (auto simp: \<open>i \<in> Basis\<close> mem_box divide_simps)
done
finally show "x \<bullet> i \<le> a \<bullet> i + (real ?k + 1) * (b \<bullet> i - a \<bullet> i) / 2 ^ n" .
qed
qed
qed
then have "\<exists>f\<in>Basis \<rightarrow>\<^sub>E {..<2 ^ n}. x \<in> ?K0(n,f)"
apply (simp add: mem_box Bex_def)
apply (clarify dest!: bchoice)
apply (rule_tac x="restrict f Basis" in exI, simp)
done
moreover have "\<And>f. x \<in> ?K0(n,f) \<Longrightarrow> ?K0(n,f) \<subseteq> g x"
apply (clarsimp simp add: mem_box)
apply (rule e)
apply (drule bspec D, assumption)+
apply (erule order_trans)
apply (simp add: divide_simps)
using bai apply (force simp add: algebra_simps)
done
ultimately show ?thesis by auto
qed
qed
next
show "\<And>u v. cbox u v \<in> ?D0 \<Longrightarrow> \<exists>n. \<forall>i \<in> Basis. v \<bullet> i - u \<bullet> i = (b \<bullet> i - a \<bullet> i) / 2^n"
by (force simp: eq_cbox box_eq_empty field_simps dest!: aibi)
next
obtain j::'a where "j \<in> Basis"
using nonempty_Basis by blast
have "finite {L \<in> ?D0. ?K0(n,f) \<subseteq> L}" if "f \<in> Basis \<rightarrow>\<^sub>E {..<2 ^ n}" for n f
proof (rule finite_subset)
let ?B = "(\<lambda>(n, f::'a\<Rightarrow>nat). cbox (\<Sum>i\<in>Basis. (a \<bullet> i + (f i) / 2^n * (b \<bullet> i - a \<bullet> i)) *\<^sub>R i)
(\<Sum>i\<in>Basis. (a \<bullet> i + ((f i) + 1) / 2^n * (b \<bullet> i - a \<bullet> i)) *\<^sub>R i))
` (SIGMA m:atMost n. Pi\<^sub>E Basis (\<lambda>i::'a. lessThan (2^m)))"
have "?K0(m,g) \<in> ?B" if "g \<in> Basis \<rightarrow>\<^sub>E {..<2 ^ m}" "?K0(n,f) \<subseteq> ?K0(m,g)" for m g
proof -
have dd: "w / m \<le> v / n \<and> (v+1) / n \<le> (w+1) / m
\<Longrightarrow> inverse n \<le> inverse m" for w m v n::real
- by (auto simp: field_split_simps algebra_simps)
+ by (auto simp: field_split_simps)
have bjaj: "b \<bullet> j - a \<bullet> j > 0"
using \<open>j \<in> Basis\<close> \<open>box a b \<noteq> {}\<close> box_eq_empty(1) by fastforce
have "((g j) / 2 ^ m) * (b \<bullet> j - a \<bullet> j) \<le> ((f j) / 2 ^ n) * (b \<bullet> j - a \<bullet> j) \<and>
(((f j) + 1) / 2 ^ n) * (b \<bullet> j - a \<bullet> j) \<le> (((g j) + 1) / 2 ^ m) * (b \<bullet> j - a \<bullet> j)"
- using that \<open>j \<in> Basis\<close> by (simp add: subset_box algebra_simps field_split_simps aibi)
+ using that \<open>j \<in> Basis\<close> by (simp add: subset_box field_split_simps aibi)
then have "((g j) / 2 ^ m) \<le> ((f j) / 2 ^ n) \<and>
((real(f j) + 1) / 2 ^ n) \<le> ((real(g j) + 1) / 2 ^ m)"
by (metis bjaj mult.commute of_nat_1 of_nat_add real_mult_le_cancel_iff2)
then have "inverse (2^n) \<le> (inverse (2^m) :: real)"
by (rule dd)
then have "m \<le> n"
by auto
show ?thesis
by (rule imageI) (simp add: \<open>m \<le> n\<close> that)
qed
then show "{L \<in> ?D0. ?K0(n,f) \<subseteq> L} \<subseteq> ?B"
by auto
show "finite ?B"
by (intro finite_imageI finite_SigmaI finite_atMost finite_lessThan finite_PiE finite_Basis)
qed
then show "finite {L \<in> ?D0. K \<subseteq> L}" if "K \<in> ?D0" for K
using that by auto
qed
let ?D1 = "{K \<in> \<D>0. \<exists>x \<in> S \<inter> K. K \<subseteq> g x}"
obtain \<D> where count: "countable \<D>"
and sub: "\<Union>\<D> \<subseteq> cbox a b" "S \<subseteq> \<Union>\<D>"
and int: "\<And>K. K \<in> \<D> \<Longrightarrow> (interior K \<noteq> {}) \<and> (\<exists>c d. K = cbox c d)"
and intdj: "\<And>A B. \<lbrakk>A \<in> \<D>; B \<in> \<D>\<rbrakk> \<Longrightarrow> A \<subseteq> B \<or> B \<subseteq> A \<or> interior A \<inter> interior B = {}"
and SK: "\<And>K. K \<in> \<D> \<Longrightarrow> \<exists>x. x \<in> S \<inter> K \<and> K \<subseteq> g x"
and cbox: "\<And>u v. cbox u v \<in> \<D> \<Longrightarrow> \<exists>n. \<forall>i \<in> Basis. v \<bullet> i - u \<bullet> i = (b \<bullet> i - a \<bullet> i) / 2^n"
and fin: "\<And>K. K \<in> \<D> \<Longrightarrow> finite {L. L \<in> \<D> \<and> K \<subseteq> L}"
proof
show "countable ?D1" using count countable_subset
by (simp add: count countable_subset)
show "\<Union>?D1 \<subseteq> cbox a b"
using sub by blast
show "S \<subseteq> \<Union>?D1"
using SK by (force simp:)
show "\<And>K. K \<in> ?D1 \<Longrightarrow> (interior K \<noteq> {}) \<and> (\<exists>c d. K = cbox c d)"
using int by blast
show "\<And>A B. \<lbrakk>A \<in> ?D1; B \<in> ?D1\<rbrakk> \<Longrightarrow> A \<subseteq> B \<or> B \<subseteq> A \<or> interior A \<inter> interior B = {}"
using intdj by blast
show "\<And>K. K \<in> ?D1 \<Longrightarrow> \<exists>x. x \<in> S \<inter> K \<and> K \<subseteq> g x"
by auto
show "\<And>u v. cbox u v \<in> ?D1 \<Longrightarrow> \<exists>n. \<forall>i \<in> Basis. v \<bullet> i - u \<bullet> i = (b \<bullet> i - a \<bullet> i) / 2^n"
using cbox by blast
show "\<And>K. K \<in> ?D1 \<Longrightarrow> finite {L. L \<in> ?D1 \<and> K \<subseteq> L}"
using fin by simp (metis (mono_tags, lifting) Collect_mono rev_finite_subset)
qed
let ?\<D> = "{K \<in> \<D>. \<forall>K'. K' \<in> \<D> \<and> K \<noteq> K' \<longrightarrow> \<not>(K \<subseteq> K')}"
show ?thesis
proof (rule that)
show "countable ?\<D>"
by (blast intro: countable_subset [OF _ count])
show "\<Union>?\<D> \<subseteq> cbox a b"
using sub by blast
show "S \<subseteq> \<Union>?\<D>"
proof clarsimp
fix x
assume "x \<in> S"
then obtain X where "x \<in> X" "X \<in> \<D>" using \<open>S \<subseteq> \<Union>\<D>\<close> by blast
let ?R = "{(K,L). K \<in> \<D> \<and> L \<in> \<D> \<and> L \<subset> K}"
have irrR: "irrefl ?R" by (force simp: irrefl_def)
have traR: "trans ?R" by (force simp: trans_def)
have finR: "\<And>x. finite {y. (y, x) \<in> ?R}"
by simp (metis (mono_tags, lifting) fin \<open>X \<in> \<D>\<close> finite_subset mem_Collect_eq psubset_imp_subset subsetI)
have "{X \<in> \<D>. x \<in> X} \<noteq> {}"
using \<open>X \<in> \<D>\<close> \<open>x \<in> X\<close> by blast
then obtain Y where "Y \<in> {X \<in> \<D>. x \<in> X}" "\<And>Y'. (Y', Y) \<in> ?R \<Longrightarrow> Y' \<notin> {X \<in> \<D>. x \<in> X}"
by (rule wfE_min' [OF wf_finite_segments [OF irrR traR finR]]) blast
then show "\<exists>Y. Y \<in> \<D> \<and> (\<forall>K'. K' \<in> \<D> \<and> Y \<noteq> K' \<longrightarrow> \<not> Y \<subseteq> K') \<and> x \<in> Y"
by blast
qed
show "\<And>K. K \<in> ?\<D> \<Longrightarrow> interior K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
by (blast intro: dest: int)
show "pairwise (\<lambda>A B. interior A \<inter> interior B = {}) ?\<D>"
using intdj by (simp add: pairwise_def) metis
show "\<And>K. K \<in> ?\<D> \<Longrightarrow> \<exists>x \<in> S \<inter> K. K \<subseteq> g x"
using SK by force
show "\<And>u v. cbox u v \<in> ?\<D> \<Longrightarrow> \<exists>n. \<forall>i\<in>Basis. v \<bullet> i - u \<bullet> i = (b \<bullet> i - a \<bullet> i) / 2^n"
using cbox by force
qed
qed
subsection \<open>Division filter\<close>
text \<open>Divisions over all gauges towards finer divisions.\<close>
definition\<^marker>\<open>tag important\<close> division_filter :: "'a::euclidean_space set \<Rightarrow> ('a \<times> 'a set) set filter"
where "division_filter s = (INF g\<in>{g. gauge g}. principal {p. p tagged_division_of s \<and> g fine p})"
proposition eventually_division_filter:
"(\<forall>\<^sub>F p in division_filter s. P p) \<longleftrightarrow>
(\<exists>g. gauge g \<and> (\<forall>p. p tagged_division_of s \<and> g fine p \<longrightarrow> P p))"
unfolding division_filter_def
proof (subst eventually_INF_base; clarsimp)
fix g1 g2 :: "'a \<Rightarrow> 'a set" show "gauge g1 \<Longrightarrow> gauge g2 \<Longrightarrow> \<exists>x. gauge x \<and>
{p. p tagged_division_of s \<and> x fine p} \<subseteq> {p. p tagged_division_of s \<and> g1 fine p} \<and>
{p. p tagged_division_of s \<and> x fine p} \<subseteq> {p. p tagged_division_of s \<and> g2 fine p}"
by (intro exI[of _ "\<lambda>x. g1 x \<inter> g2 x"]) (auto simp: fine_Int)
qed (auto simp: eventually_principal)
lemma division_filter_not_empty: "division_filter (cbox a b) \<noteq> bot"
unfolding trivial_limit_def eventually_division_filter
by (auto elim: fine_division_exists)
lemma eventually_division_filter_tagged_division:
"eventually (\<lambda>p. p tagged_division_of s) (division_filter s)"
unfolding eventually_division_filter by (intro exI[of _ "\<lambda>x. ball x 1"]) auto
end
diff --git a/src/HOL/Analysis/Weierstrass_Theorems.thy b/src/HOL/Analysis/Weierstrass_Theorems.thy
--- a/src/HOL/Analysis/Weierstrass_Theorems.thy
+++ b/src/HOL/Analysis/Weierstrass_Theorems.thy
@@ -1,1417 +1,1417 @@
section \<open>Bernstein-Weierstrass and Stone-Weierstrass\<close>
text\<open>By L C Paulson (2015)\<close>
theory Weierstrass_Theorems
imports Uniform_Limit Path_Connected Derivative
begin
subsection \<open>Bernstein polynomials\<close>
definition\<^marker>\<open>tag important\<close> Bernstein :: "[nat,nat,real] \<Rightarrow> real" where
"Bernstein n k x \<equiv> of_nat (n choose k) * x ^ k * (1 - x) ^ (n - k)"
lemma Bernstein_nonneg: "\<lbrakk>0 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> 0 \<le> Bernstein n k x"
by (simp add: Bernstein_def)
lemma Bernstein_pos: "\<lbrakk>0 < x; x < 1; k \<le> n\<rbrakk> \<Longrightarrow> 0 < Bernstein n k x"
by (simp add: Bernstein_def)
lemma sum_Bernstein [simp]: "(\<Sum>k\<le>n. Bernstein n k x) = 1"
using binomial_ring [of x "1-x" n]
by (simp add: Bernstein_def)
lemma binomial_deriv1:
"(\<Sum>k\<le>n. (of_nat k * of_nat (n choose k)) * a^(k-1) * b^(n-k)) = real_of_nat n * (a+b) ^ (n-1)"
apply (rule DERIV_unique [where f = "\<lambda>a. (a+b)^n" and x=a])
apply (subst binomial_ring)
apply (rule derivative_eq_intros sum.cong | simp add: atMost_atLeast0)+
done
lemma binomial_deriv2:
"(\<Sum>k\<le>n. (of_nat k * of_nat (k-1) * of_nat (n choose k)) * a^(k-2) * b^(n-k)) =
of_nat n * of_nat (n-1) * (a+b::real) ^ (n-2)"
apply (rule DERIV_unique [where f = "\<lambda>a. of_nat n * (a+b::real) ^ (n-1)" and x=a])
apply (subst binomial_deriv1 [symmetric])
apply (rule derivative_eq_intros sum.cong | simp add: Num.numeral_2_eq_2)+
done
lemma sum_k_Bernstein [simp]: "(\<Sum>k\<le>n. real k * Bernstein n k x) = of_nat n * x"
apply (subst binomial_deriv1 [of n x "1-x", simplified, symmetric])
apply (simp add: sum_distrib_right)
apply (auto simp: Bernstein_def algebra_simps power_eq_if intro!: sum.cong)
done
lemma sum_kk_Bernstein [simp]: "(\<Sum>k\<le>n. real k * (real k - 1) * Bernstein n k x) = real n * (real n - 1) * x\<^sup>2"
proof -
have "(\<Sum>k\<le>n. real k * (real k - 1) * Bernstein n k x) =
(\<Sum>k\<le>n. real k * real (k - Suc 0) * real (n choose k) * x ^ (k - 2) * (1 - x) ^ (n - k) * x\<^sup>2)"
proof (rule sum.cong [OF refl], simp)
fix k
assume "k \<le> n"
then consider "k = 0" | "k = 1" | k' where "k = Suc (Suc k')"
by (metis One_nat_def not0_implies_Suc)
then show "k = 0 \<or>
(real k - 1) * Bernstein n k x =
real (k - Suc 0) *
(real (n choose k) * (x ^ (k - 2) * ((1 - x) ^ (n - k) * x\<^sup>2)))"
by cases (auto simp add: Bernstein_def power2_eq_square algebra_simps)
qed
also have "... = real_of_nat n * real_of_nat (n - Suc 0) * x\<^sup>2"
by (subst binomial_deriv2 [of n x "1-x", simplified, symmetric]) (simp add: sum_distrib_right)
also have "... = n * (n - 1) * x\<^sup>2"
by auto
finally show ?thesis
by auto
qed
subsection \<open>Explicit Bernstein version of the 1D Weierstrass approximation theorem\<close>
theorem Bernstein_Weierstrass:
fixes f :: "real \<Rightarrow> real"
assumes contf: "continuous_on {0..1} f" and e: "0 < e"
shows "\<exists>N. \<forall>n x. N \<le> n \<and> x \<in> {0..1}
\<longrightarrow> \<bar>f x - (\<Sum>k\<le>n. f(k/n) * Bernstein n k x)\<bar> < e"
proof -
have "bounded (f ` {0..1})"
using compact_continuous_image compact_imp_bounded contf by blast
then obtain M where M: "\<And>x. 0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> \<bar>f x\<bar> \<le> M"
by (force simp add: bounded_iff)
then have Mge0: "0 \<le> M" by force
have ucontf: "uniformly_continuous_on {0..1} f"
using compact_uniformly_continuous contf by blast
then obtain d where d: "d>0" "\<And>x x'. \<lbrakk> x \<in> {0..1}; x' \<in> {0..1}; \<bar>x' - x\<bar> < d\<rbrakk> \<Longrightarrow> \<bar>f x' - f x\<bar> < e/2"
apply (rule uniformly_continuous_onE [where e = "e/2"])
using e by (auto simp: dist_norm)
{ fix n::nat and x::real
assume n: "Suc (nat\<lceil>4*M/(e*d\<^sup>2)\<rceil>) \<le> n" and x: "0 \<le> x" "x \<le> 1"
have "0 < n" using n by simp
have ed0: "- (e * d\<^sup>2) < 0"
using e \<open>0<d\<close> by simp
also have "... \<le> M * 4"
using \<open>0\<le>M\<close> by simp
finally have [simp]: "real_of_int (nat \<lceil>4 * M / (e * d\<^sup>2)\<rceil>) = real_of_int \<lceil>4 * M / (e * d\<^sup>2)\<rceil>"
using \<open>0\<le>M\<close> e \<open>0<d\<close>
by (simp add: field_simps)
have "4*M/(e*d\<^sup>2) + 1 \<le> real (Suc (nat\<lceil>4*M/(e*d\<^sup>2)\<rceil>))"
by (simp add: real_nat_ceiling_ge)
also have "... \<le> real n"
using n by (simp add: field_simps)
finally have nbig: "4*M/(e*d\<^sup>2) + 1 \<le> real n" .
have sum_bern: "(\<Sum>k\<le>n. (x - k/n)\<^sup>2 * Bernstein n k x) = x * (1 - x) / n"
proof -
have *: "\<And>a b x::real. (a - b)\<^sup>2 * x = a * (a - 1) * x + (1 - 2 * b) * a * x + b * b * x"
by (simp add: algebra_simps power2_eq_square)
have "(\<Sum>k\<le>n. (k - n * x)\<^sup>2 * Bernstein n k x) = n * x * (1 - x)"
apply (simp add: * sum.distrib)
apply (simp flip: sum_distrib_left add: mult.assoc)
apply (simp add: algebra_simps power2_eq_square)
done
then have "(\<Sum>k\<le>n. (k - n * x)\<^sup>2 * Bernstein n k x)/n^2 = x * (1 - x) / n"
by (simp add: power2_eq_square)
then show ?thesis
- using n by (simp add: sum_divide_distrib field_split_simps mult.commute power2_commute)
+ using n by (simp add: sum_divide_distrib field_split_simps power2_commute)
qed
{ fix k
assume k: "k \<le> n"
then have kn: "0 \<le> k / n" "k / n \<le> 1"
by (auto simp: field_split_simps)
consider (lessd) "\<bar>x - k / n\<bar> < d" | (ged) "d \<le> \<bar>x - k / n\<bar>"
by linarith
then have "\<bar>(f x - f (k/n))\<bar> \<le> e/2 + 2 * M / d\<^sup>2 * (x - k/n)\<^sup>2"
proof cases
case lessd
then have "\<bar>(f x - f (k/n))\<bar> < e/2"
using d x kn by (simp add: abs_minus_commute)
also have "... \<le> (e/2 + 2 * M / d\<^sup>2 * (x - k/n)\<^sup>2)"
using Mge0 d by simp
finally show ?thesis by simp
next
case ged
then have dle: "d\<^sup>2 \<le> (x - k/n)\<^sup>2"
by (metis d(1) less_eq_real_def power2_abs power_mono)
have "\<bar>(f x - f (k/n))\<bar> \<le> \<bar>f x\<bar> + \<bar>f (k/n)\<bar>"
by (rule abs_triangle_ineq4)
also have "... \<le> M+M"
by (meson M add_mono_thms_linordered_semiring(1) kn x)
also have "... \<le> 2 * M * ((x - k/n)\<^sup>2 / d\<^sup>2)"
apply simp
apply (rule Rings.ordered_semiring_class.mult_left_mono [of 1 "((x - k/n)\<^sup>2 / d\<^sup>2)", simplified])
using dle \<open>d>0\<close> \<open>M\<ge>0\<close> by auto
also have "... \<le> e/2 + 2 * M / d\<^sup>2 * (x - k/n)\<^sup>2"
using e by simp
finally show ?thesis .
qed
} note * = this
have "\<bar>f x - (\<Sum>k\<le>n. f(k / n) * Bernstein n k x)\<bar> \<le> \<bar>\<Sum>k\<le>n. (f x - f(k / n)) * Bernstein n k x\<bar>"
by (simp add: sum_subtractf sum_distrib_left [symmetric] algebra_simps)
also have "... \<le> (\<Sum>k\<le>n. (e/2 + (2 * M / d\<^sup>2) * (x - k / n)\<^sup>2) * Bernstein n k x)"
apply (rule order_trans [OF sum_abs sum_mono])
using *
apply (simp add: abs_mult Bernstein_nonneg x mult_right_mono)
done
also have "... \<le> e/2 + (2 * M) / (d\<^sup>2 * n)"
apply (simp only: sum.distrib Rings.semiring_class.distrib_right sum_distrib_left [symmetric] mult.assoc sum_bern)
using \<open>d>0\<close> x
apply (simp add: divide_simps Mge0 mult_le_one mult_left_le)
done
also have "... < e"
apply (simp add: field_simps)
using \<open>d>0\<close> nbig e \<open>n>0\<close>
- apply (simp add: field_split_simps algebra_simps)
+ apply (simp add: field_split_simps)
using ed0 by linarith
finally have "\<bar>f x - (\<Sum>k\<le>n. f (real k / real n) * Bernstein n k x)\<bar> < e" .
}
then show ?thesis
by auto
qed
subsection \<open>General Stone-Weierstrass theorem\<close>
text\<open>Source:
Bruno Brosowski and Frank Deutsch.
An Elementary Proof of the Stone-Weierstrass Theorem.
Proceedings of the American Mathematical Society
Volume 81, Number 1, January 1981.
DOI: 10.2307/2043993 \<^url>\<open>https://www.jstor.org/stable/2043993\<close>\<close>
locale function_ring_on =
fixes R :: "('a::t2_space \<Rightarrow> real) set" and S :: "'a set"
assumes compact: "compact S"
assumes continuous: "f \<in> R \<Longrightarrow> continuous_on S f"
assumes add: "f \<in> R \<Longrightarrow> g \<in> R \<Longrightarrow> (\<lambda>x. f x + g x) \<in> R"
assumes mult: "f \<in> R \<Longrightarrow> g \<in> R \<Longrightarrow> (\<lambda>x. f x * g x) \<in> R"
assumes const: "(\<lambda>_. c) \<in> R"
assumes separable: "x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> x \<noteq> y \<Longrightarrow> \<exists>f\<in>R. f x \<noteq> f y"
begin
lemma minus: "f \<in> R \<Longrightarrow> (\<lambda>x. - f x) \<in> R"
by (frule mult [OF const [of "-1"]]) simp
lemma diff: "f \<in> R \<Longrightarrow> g \<in> R \<Longrightarrow> (\<lambda>x. f x - g x) \<in> R"
unfolding diff_conv_add_uminus by (metis add minus)
lemma power: "f \<in> R \<Longrightarrow> (\<lambda>x. f x ^ n) \<in> R"
by (induct n) (auto simp: const mult)
lemma sum: "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> f i \<in> R\<rbrakk> \<Longrightarrow> (\<lambda>x. \<Sum>i \<in> I. f i x) \<in> R"
by (induct I rule: finite_induct; simp add: const add)
lemma prod: "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> f i \<in> R\<rbrakk> \<Longrightarrow> (\<lambda>x. \<Prod>i \<in> I. f i x) \<in> R"
by (induct I rule: finite_induct; simp add: const mult)
definition\<^marker>\<open>tag important\<close> normf :: "('a::t2_space \<Rightarrow> real) \<Rightarrow> real"
where "normf f \<equiv> SUP x\<in>S. \<bar>f x\<bar>"
lemma normf_upper: "\<lbrakk>continuous_on S f; x \<in> S\<rbrakk> \<Longrightarrow> \<bar>f x\<bar> \<le> normf f"
apply (simp add: normf_def)
apply (rule cSUP_upper, assumption)
by (simp add: bounded_imp_bdd_above compact compact_continuous_image compact_imp_bounded continuous_on_rabs)
lemma normf_least: "S \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> \<bar>f x\<bar> \<le> M) \<Longrightarrow> normf f \<le> M"
by (simp add: normf_def cSUP_least)
end
lemma (in function_ring_on) one:
assumes U: "open U" and t0: "t0 \<in> S" "t0 \<in> U" and t1: "t1 \<in> S-U"
shows "\<exists>V. open V \<and> t0 \<in> V \<and> S \<inter> V \<subseteq> U \<and>
(\<forall>e>0. \<exists>f \<in> R. f ` S \<subseteq> {0..1} \<and> (\<forall>t \<in> S \<inter> V. f t < e) \<and> (\<forall>t \<in> S - U. f t > 1 - e))"
proof -
have "\<exists>pt \<in> R. pt t0 = 0 \<and> pt t > 0 \<and> pt ` S \<subseteq> {0..1}" if t: "t \<in> S - U" for t
proof -
have "t \<noteq> t0" using t t0 by auto
then obtain g where g: "g \<in> R" "g t \<noteq> g t0"
using separable t0 by (metis Diff_subset subset_eq t)
define h where [abs_def]: "h x = g x - g t0" for x
have "h \<in> R"
unfolding h_def by (fast intro: g const diff)
then have hsq: "(\<lambda>w. (h w)\<^sup>2) \<in> R"
by (simp add: power2_eq_square mult)
have "h t \<noteq> h t0"
by (simp add: h_def g)
then have "h t \<noteq> 0"
by (simp add: h_def)
then have ht2: "0 < (h t)^2"
by simp
also have "... \<le> normf (\<lambda>w. (h w)\<^sup>2)"
using t normf_upper [where x=t] continuous [OF hsq] by force
finally have nfp: "0 < normf (\<lambda>w. (h w)\<^sup>2)" .
define p where [abs_def]: "p x = (1 / normf (\<lambda>w. (h w)\<^sup>2)) * (h x)^2" for x
have "p \<in> R"
unfolding p_def by (fast intro: hsq const mult)
moreover have "p t0 = 0"
by (simp add: p_def h_def)
moreover have "p t > 0"
using nfp ht2 by (simp add: p_def)
moreover have "\<And>x. x \<in> S \<Longrightarrow> p x \<in> {0..1}"
using nfp normf_upper [OF continuous [OF hsq] ] by (auto simp: p_def)
ultimately show "\<exists>pt \<in> R. pt t0 = 0 \<and> pt t > 0 \<and> pt ` S \<subseteq> {0..1}"
by auto
qed
then obtain pf where pf: "\<And>t. t \<in> S-U \<Longrightarrow> pf t \<in> R \<and> pf t t0 = 0 \<and> pf t t > 0"
and pf01: "\<And>t. t \<in> S-U \<Longrightarrow> pf t ` S \<subseteq> {0..1}"
by metis
have com_sU: "compact (S-U)"
using compact closed_Int_compact U by (simp add: Diff_eq compact_Int_closed open_closed)
have "\<And>t. t \<in> S-U \<Longrightarrow> \<exists>A. open A \<and> A \<inter> S = {x\<in>S. 0 < pf t x}"
apply (rule open_Collect_positive)
by (metis pf continuous)
then obtain Uf where Uf: "\<And>t. t \<in> S-U \<Longrightarrow> open (Uf t) \<and> (Uf t) \<inter> S = {x\<in>S. 0 < pf t x}"
by metis
then have open_Uf: "\<And>t. t \<in> S-U \<Longrightarrow> open (Uf t)"
by blast
have tUft: "\<And>t. t \<in> S-U \<Longrightarrow> t \<in> Uf t"
using pf Uf by blast
then have *: "S-U \<subseteq> (\<Union>x \<in> S-U. Uf x)"
by blast
obtain subU where subU: "subU \<subseteq> S - U" "finite subU" "S - U \<subseteq> (\<Union>x \<in> subU. Uf x)"
by (blast intro: that compactE_image [OF com_sU open_Uf *])
then have [simp]: "subU \<noteq> {}"
using t1 by auto
then have cardp: "card subU > 0" using subU
by (simp add: card_gt_0_iff)
define p where [abs_def]: "p x = (1 / card subU) * (\<Sum>t \<in> subU. pf t x)" for x
have pR: "p \<in> R"
unfolding p_def using subU pf by (fast intro: pf const mult sum)
have pt0 [simp]: "p t0 = 0"
using subU pf by (auto simp: p_def intro: sum.neutral)
have pt_pos: "p t > 0" if t: "t \<in> S-U" for t
proof -
obtain i where i: "i \<in> subU" "t \<in> Uf i" using subU t by blast
show ?thesis
using subU i t
apply (clarsimp simp: p_def field_split_simps)
apply (rule sum_pos2 [OF \<open>finite subU\<close>])
using Uf t pf01 apply auto
apply (force elim!: subsetCE)
done
qed
have p01: "p x \<in> {0..1}" if t: "x \<in> S" for x
proof -
have "0 \<le> p x"
using subU cardp t
apply (simp add: p_def field_split_simps sum_nonneg)
apply (rule sum_nonneg)
using pf01 by force
moreover have "p x \<le> 1"
using subU cardp t
apply (simp add: p_def field_split_simps sum_nonneg)
apply (rule sum_bounded_above [where 'a=real and K=1, simplified])
using pf01 by force
ultimately show ?thesis
by auto
qed
have "compact (p ` (S-U))"
by (meson Diff_subset com_sU compact_continuous_image continuous continuous_on_subset pR)
then have "open (- (p ` (S-U)))"
by (simp add: compact_imp_closed open_Compl)
moreover have "0 \<in> - (p ` (S-U))"
by (metis (no_types) ComplI image_iff not_less_iff_gr_or_eq pt_pos)
ultimately obtain delta0 where delta0: "delta0 > 0" "ball 0 delta0 \<subseteq> - (p ` (S-U))"
by (auto simp: elim!: openE)
then have pt_delta: "\<And>x. x \<in> S-U \<Longrightarrow> p x \<ge> delta0"
by (force simp: ball_def dist_norm dest: p01)
define \<delta> where "\<delta> = delta0/2"
have "delta0 \<le> 1" using delta0 p01 [of t1] t1
by (force simp: ball_def dist_norm dest: p01)
with delta0 have \<delta>01: "0 < \<delta>" "\<delta> < 1"
by (auto simp: \<delta>_def)
have pt_\<delta>: "\<And>x. x \<in> S-U \<Longrightarrow> p x \<ge> \<delta>"
using pt_delta delta0 by (force simp: \<delta>_def)
have "\<exists>A. open A \<and> A \<inter> S = {x\<in>S. p x < \<delta>/2}"
by (rule open_Collect_less_Int [OF continuous [OF pR] continuous_on_const])
then obtain V where V: "open V" "V \<inter> S = {x\<in>S. p x < \<delta>/2}"
by blast
define k where "k = nat\<lfloor>1/\<delta>\<rfloor> + 1"
have "k>0" by (simp add: k_def)
have "k-1 \<le> 1/\<delta>"
using \<delta>01 by (simp add: k_def)
with \<delta>01 have "k \<le> (1+\<delta>)/\<delta>"
by (auto simp: algebra_simps add_divide_distrib)
also have "... < 2/\<delta>"
using \<delta>01 by (auto simp: field_split_simps)
finally have k2\<delta>: "k < 2/\<delta>" .
have "1/\<delta> < k"
using \<delta>01 unfolding k_def by linarith
with \<delta>01 k2\<delta> have k\<delta>: "1 < k*\<delta>" "k*\<delta> < 2"
by (auto simp: field_split_simps)
define q where [abs_def]: "q n t = (1 - p t ^ n) ^ (k^n)" for n t
have qR: "q n \<in> R" for n
by (simp add: q_def const diff power pR)
have q01: "\<And>n t. t \<in> S \<Longrightarrow> q n t \<in> {0..1}"
using p01 by (simp add: q_def power_le_one algebra_simps)
have qt0 [simp]: "\<And>n. n>0 \<Longrightarrow> q n t0 = 1"
using t0 pf by (simp add: q_def power_0_left)
{ fix t and n::nat
assume t: "t \<in> S \<inter> V"
with \<open>k>0\<close> V have "k * p t < k * \<delta> / 2"
by force
then have "1 - (k * \<delta> / 2)^n \<le> 1 - (k * p t)^n"
using \<open>k>0\<close> p01 t by (simp add: power_mono)
also have "... \<le> q n t"
using Bernoulli_inequality [of "- ((p t)^n)" "k^n"]
apply (simp add: q_def)
by (metis IntE atLeastAtMost_iff p01 power_le_one power_mult_distrib t)
finally have "1 - (k * \<delta> / 2) ^ n \<le> q n t" .
} note limitV = this
{ fix t and n::nat
assume t: "t \<in> S - U"
with \<open>k>0\<close> U have "k * \<delta> \<le> k * p t"
by (simp add: pt_\<delta>)
with k\<delta> have kpt: "1 < k * p t"
by (blast intro: less_le_trans)
have ptn_pos: "0 < p t ^ n"
using pt_pos [OF t] by simp
have ptn_le: "p t ^ n \<le> 1"
by (meson DiffE atLeastAtMost_iff p01 power_le_one t)
have "q n t = (1/(k^n * (p t)^n)) * (1 - p t ^ n) ^ (k^n) * k^n * (p t)^n"
using pt_pos [OF t] \<open>k>0\<close> by (simp add: q_def)
also have "... \<le> (1/(k * (p t))^n) * (1 - p t ^ n) ^ (k^n) * (1 + k^n * (p t)^n)"
using pt_pos [OF t] \<open>k>0\<close>
apply simp
apply (simp only: times_divide_eq_right [symmetric])
apply (rule mult_left_mono [of "1::real", simplified])
apply (simp_all add: power_mult_distrib)
apply (rule zero_le_power)
using ptn_le by linarith
also have "... \<le> (1/(k * (p t))^n) * (1 - p t ^ n) ^ (k^n) * (1 + (p t)^n) ^ (k^n)"
apply (rule mult_left_mono [OF Bernoulli_inequality [of "p t ^ n" "k^n"]])
using \<open>k>0\<close> ptn_pos ptn_le
apply (auto simp: power_mult_distrib)
done
also have "... = (1/(k * (p t))^n) * (1 - p t ^ (2*n)) ^ (k^n)"
using pt_pos [OF t] \<open>k>0\<close>
by (simp add: algebra_simps power_mult power2_eq_square flip: power_mult_distrib)
also have "... \<le> (1/(k * (p t))^n) * 1"
apply (rule mult_left_mono [OF power_le_one])
using pt_pos \<open>k>0\<close> p01 power_le_one t apply auto
done
also have "... \<le> (1 / (k*\<delta>))^n"
using \<open>k>0\<close> \<delta>01 power_mono pt_\<delta> t
by (fastforce simp: field_simps)
finally have "q n t \<le> (1 / (real k * \<delta>)) ^ n " .
} note limitNonU = this
define NN
where "NN e = 1 + nat \<lceil>max (ln e / ln (real k * \<delta> / 2)) (- ln e / ln (real k * \<delta>))\<rceil>" for e
have NN: "of_nat (NN e) > ln e / ln (real k * \<delta> / 2)" "of_nat (NN e) > - ln e / ln (real k * \<delta>)"
if "0<e" for e
unfolding NN_def by linarith+
have NN1: "\<And>e. e>0 \<Longrightarrow> (k * \<delta> / 2)^NN e < e"
apply (subst Transcendental.ln_less_cancel_iff [symmetric])
prefer 3 apply (subst ln_realpow)
using \<open>k>0\<close> \<open>\<delta>>0\<close> NN k\<delta>
apply (force simp add: field_simps)+
done
have NN0: "(1/(k*\<delta>)) ^ (NN e) < e" if "e>0" for e
proof -
have "0 < ln (real k) + ln \<delta>"
using \<delta>01(1) \<open>0 < k\<close> k\<delta>(1) ln_gt_zero ln_mult by fastforce
then have "real (NN e) * ln (1 / (real k * \<delta>)) < ln e"
using k\<delta>(1) NN(2) [of e] that by (simp add: ln_div divide_simps)
then have "exp (real (NN e) * ln (1 / (real k * \<delta>))) < e"
by (metis exp_less_mono exp_ln that)
then show ?thesis
by (simp add: \<delta>01(1) \<open>0 < k\<close> exp_of_nat_mult)
qed
{ fix t and e::real
assume "e>0"
have "t \<in> S \<inter> V \<Longrightarrow> 1 - q (NN e) t < e" "t \<in> S - U \<Longrightarrow> q (NN e) t < e"
proof -
assume t: "t \<in> S \<inter> V"
show "1 - q (NN e) t < e"
by (metis add.commute diff_le_eq not_le limitV [OF t] less_le_trans [OF NN1 [OF \<open>e>0\<close>]])
next
assume t: "t \<in> S - U"
show "q (NN e) t < e"
using limitNonU [OF t] less_le_trans [OF NN0 [OF \<open>e>0\<close>]] not_le by blast
qed
} then have "\<And>e. e > 0 \<Longrightarrow> \<exists>f\<in>R. f ` S \<subseteq> {0..1} \<and> (\<forall>t \<in> S \<inter> V. f t < e) \<and> (\<forall>t \<in> S - U. 1 - e < f t)"
using q01
by (rule_tac x="\<lambda>x. 1 - q (NN e) x" in bexI) (auto simp: algebra_simps intro: diff const qR)
moreover have t0V: "t0 \<in> V" "S \<inter> V \<subseteq> U"
using pt_\<delta> t0 U V \<delta>01 by fastforce+
ultimately show ?thesis using V t0V
by blast
qed
text\<open>Non-trivial case, with \<^term>\<open>A\<close> and \<^term>\<open>B\<close> both non-empty\<close>
lemma (in function_ring_on) two_special:
assumes A: "closed A" "A \<subseteq> S" "a \<in> A"
and B: "closed B" "B \<subseteq> S" "b \<in> B"
and disj: "A \<inter> B = {}"
and e: "0 < e" "e < 1"
shows "\<exists>f \<in> R. f ` S \<subseteq> {0..1} \<and> (\<forall>x \<in> A. f x < e) \<and> (\<forall>x \<in> B. f x > 1 - e)"
proof -
{ fix w
assume "w \<in> A"
then have "open ( - B)" "b \<in> S" "w \<notin> B" "w \<in> S"
using assms by auto
then have "\<exists>V. open V \<and> w \<in> V \<and> S \<inter> V \<subseteq> -B \<and>
(\<forall>e>0. \<exists>f \<in> R. f ` S \<subseteq> {0..1} \<and> (\<forall>x \<in> S \<inter> V. f x < e) \<and> (\<forall>x \<in> S \<inter> B. f x > 1 - e))"
using one [of "-B" w b] assms \<open>w \<in> A\<close> by simp
}
then obtain Vf where Vf:
"\<And>w. w \<in> A \<Longrightarrow> open (Vf w) \<and> w \<in> Vf w \<and> S \<inter> Vf w \<subseteq> -B \<and>
(\<forall>e>0. \<exists>f \<in> R. f ` S \<subseteq> {0..1} \<and> (\<forall>x \<in> S \<inter> Vf w. f x < e) \<and> (\<forall>x \<in> S \<inter> B. f x > 1 - e))"
by metis
then have open_Vf: "\<And>w. w \<in> A \<Longrightarrow> open (Vf w)"
by blast
have tVft: "\<And>w. w \<in> A \<Longrightarrow> w \<in> Vf w"
using Vf by blast
then have sum_max_0: "A \<subseteq> (\<Union>x \<in> A. Vf x)"
by blast
have com_A: "compact A" using A
by (metis compact compact_Int_closed inf.absorb_iff2)
obtain subA where subA: "subA \<subseteq> A" "finite subA" "A \<subseteq> (\<Union>x \<in> subA. Vf x)"
by (blast intro: that compactE_image [OF com_A open_Vf sum_max_0])
then have [simp]: "subA \<noteq> {}"
using \<open>a \<in> A\<close> by auto
then have cardp: "card subA > 0" using subA
by (simp add: card_gt_0_iff)
have "\<And>w. w \<in> A \<Longrightarrow> \<exists>f \<in> R. f ` S \<subseteq> {0..1} \<and> (\<forall>x \<in> S \<inter> Vf w. f x < e / card subA) \<and> (\<forall>x \<in> S \<inter> B. f x > 1 - e / card subA)"
using Vf e cardp by simp
then obtain ff where ff:
"\<And>w. w \<in> A \<Longrightarrow> ff w \<in> R \<and> ff w ` S \<subseteq> {0..1} \<and>
(\<forall>x \<in> S \<inter> Vf w. ff w x < e / card subA) \<and> (\<forall>x \<in> S \<inter> B. ff w x > 1 - e / card subA)"
by metis
define pff where [abs_def]: "pff x = (\<Prod>w \<in> subA. ff w x)" for x
have pffR: "pff \<in> R"
unfolding pff_def using subA ff by (auto simp: intro: prod)
moreover
have pff01: "pff x \<in> {0..1}" if t: "x \<in> S" for x
proof -
have "0 \<le> pff x"
using subA cardp t
apply (simp add: pff_def field_split_simps sum_nonneg)
apply (rule Groups_Big.linordered_semidom_class.prod_nonneg)
using ff by fastforce
moreover have "pff x \<le> 1"
using subA cardp t
apply (simp add: pff_def field_split_simps sum_nonneg)
apply (rule prod_mono [where g = "\<lambda>x. 1", simplified])
using ff by fastforce
ultimately show ?thesis
by auto
qed
moreover
{ fix v x
assume v: "v \<in> subA" and x: "x \<in> Vf v" "x \<in> S"
from subA v have "pff x = ff v x * (\<Prod>w \<in> subA - {v}. ff w x)"
unfolding pff_def by (metis prod.remove)
also have "... \<le> ff v x * 1"
apply (rule Rings.ordered_semiring_class.mult_left_mono)
apply (rule prod_mono [where g = "\<lambda>x. 1", simplified])
using ff [THEN conjunct2, THEN conjunct1] v subA x
apply auto
apply (meson atLeastAtMost_iff contra_subsetD imageI)
apply (meson atLeastAtMost_iff contra_subsetD image_eqI)
using atLeastAtMost_iff by blast
also have "... < e / card subA"
using ff [THEN conjunct2, THEN conjunct2, THEN conjunct1] v subA x
by auto
also have "... \<le> e"
using cardp e by (simp add: field_split_simps)
finally have "pff x < e" .
}
then have "\<And>x. x \<in> A \<Longrightarrow> pff x < e"
using A Vf subA by (metis UN_E contra_subsetD)
moreover
{ fix x
assume x: "x \<in> B"
then have "x \<in> S"
using B by auto
have "1 - e \<le> (1 - e / card subA) ^ card subA"
using Bernoulli_inequality [of "-e / card subA" "card subA"] e cardp
by (auto simp: field_simps)
also have "... = (\<Prod>w \<in> subA. 1 - e / card subA)"
by (simp add: subA(2))
also have "... < pff x"
apply (simp add: pff_def)
apply (rule prod_mono_strict [where f = "\<lambda>x. 1 - e / card subA", simplified])
apply (simp_all add: subA(2))
apply (intro ballI conjI)
using e apply (force simp: field_split_simps)
using ff [THEN conjunct2, THEN conjunct2, THEN conjunct2] subA B x
apply blast
done
finally have "1 - e < pff x" .
}
ultimately
show ?thesis by blast
qed
lemma (in function_ring_on) two:
assumes A: "closed A" "A \<subseteq> S"
and B: "closed B" "B \<subseteq> S"
and disj: "A \<inter> B = {}"
and e: "0 < e" "e < 1"
shows "\<exists>f \<in> R. f ` S \<subseteq> {0..1} \<and> (\<forall>x \<in> A. f x < e) \<and> (\<forall>x \<in> B. f x > 1 - e)"
proof (cases "A \<noteq> {} \<and> B \<noteq> {}")
case True then show ?thesis
apply (simp flip: ex_in_conv)
using assms
apply safe
apply (force simp add: intro!: two_special)
done
next
case False with e show ?thesis
apply simp
apply (erule disjE)
apply (rule_tac [2] x="\<lambda>x. 0" in bexI)
apply (rule_tac x="\<lambda>x. 1" in bexI)
apply (auto simp: const)
done
qed
text\<open>The special case where \<^term>\<open>f\<close> is non-negative and \<^term>\<open>e<1/3\<close>\<close>
lemma (in function_ring_on) Stone_Weierstrass_special:
assumes f: "continuous_on S f" and fpos: "\<And>x. x \<in> S \<Longrightarrow> f x \<ge> 0"
and e: "0 < e" "e < 1/3"
shows "\<exists>g \<in> R. \<forall>x\<in>S. \<bar>f x - g x\<bar> < 2*e"
proof -
define n where "n = 1 + nat \<lceil>normf f / e\<rceil>"
define A where "A j = {x \<in> S. f x \<le> (j - 1/3)*e}" for j :: nat
define B where "B j = {x \<in> S. f x \<ge> (j + 1/3)*e}" for j :: nat
have ngt: "(n-1) * e \<ge> normf f" "n\<ge>1"
using e
- apply (simp_all add: n_def field_simps of_nat_Suc)
+ apply (simp_all add: n_def field_simps)
by (metis real_nat_ceiling_ge mult.commute not_less pos_less_divide_eq)
then have ge_fx: "(n-1) * e \<ge> f x" if "x \<in> S" for x
using f normf_upper that by fastforce
{ fix j
have A: "closed (A j)" "A j \<subseteq> S"
apply (simp_all add: A_def Collect_restrict)
apply (rule continuous_on_closed_Collect_le [OF f continuous_on_const])
apply (simp add: compact compact_imp_closed)
done
have B: "closed (B j)" "B j \<subseteq> S"
apply (simp_all add: B_def Collect_restrict)
apply (rule continuous_on_closed_Collect_le [OF continuous_on_const f])
apply (simp add: compact compact_imp_closed)
done
have disj: "(A j) \<inter> (B j) = {}"
using e by (auto simp: A_def B_def field_simps)
have "\<exists>f \<in> R. f ` S \<subseteq> {0..1} \<and> (\<forall>x \<in> A j. f x < e/n) \<and> (\<forall>x \<in> B j. f x > 1 - e/n)"
apply (rule two)
using e A B disj ngt
apply simp_all
done
}
then obtain xf where xfR: "\<And>j. xf j \<in> R" and xf01: "\<And>j. xf j ` S \<subseteq> {0..1}"
and xfA: "\<And>x j. x \<in> A j \<Longrightarrow> xf j x < e/n"
and xfB: "\<And>x j. x \<in> B j \<Longrightarrow> xf j x > 1 - e/n"
by metis
define g where [abs_def]: "g x = e * (\<Sum>i\<le>n. xf i x)" for x
have gR: "g \<in> R"
unfolding g_def by (fast intro: mult const sum xfR)
have gge0: "\<And>x. x \<in> S \<Longrightarrow> g x \<ge> 0"
using e xf01 by (simp add: g_def zero_le_mult_iff image_subset_iff sum_nonneg)
have A0: "A 0 = {}"
using fpos e by (fastforce simp: A_def)
have An: "A n = S"
using e ngt f normf_upper by (fastforce simp: A_def field_simps of_nat_diff)
have Asub: "A j \<subseteq> A i" if "i\<ge>j" for i j
using e that apply (clarsimp simp: A_def)
apply (erule order_trans, simp)
done
{ fix t
assume t: "t \<in> S"
define j where "j = (LEAST j. t \<in> A j)"
have jn: "j \<le> n"
using t An by (simp add: Least_le j_def)
have Aj: "t \<in> A j"
using t An by (fastforce simp add: j_def intro: LeastI)
then have Ai: "t \<in> A i" if "i\<ge>j" for i
using Asub [OF that] by blast
then have fj1: "f t \<le> (j - 1/3)*e"
by (simp add: A_def)
then have Anj: "t \<notin> A i" if "i<j" for i
using Aj \<open>i<j\<close>
apply (simp add: j_def)
using not_less_Least by blast
have j1: "1 \<le> j"
using A0 Aj j_def not_less_eq_eq by (fastforce simp add: j_def)
then have Anj: "t \<notin> A (j-1)"
using Least_le by (fastforce simp add: j_def)
then have fj2: "(j - 4/3)*e < f t"
using j1 t by (simp add: A_def of_nat_diff)
have ***: "xf i t \<le> e/n" if "i\<ge>j" for i
using xfA [OF Ai] that by (simp add: less_eq_real_def)
{ fix i
assume "i+2 \<le> j"
then obtain d where "i+2+d = j"
using le_Suc_ex that by blast
then have "t \<in> B i"
using Anj e ge_fx [OF t] \<open>1 \<le> n\<close> fpos [OF t] t
apply (simp add: A_def B_def)
- apply (clarsimp simp add: field_simps of_nat_diff not_le of_nat_Suc)
+ apply (clarsimp simp add: field_simps of_nat_diff not_le)
apply (rule order_trans [of _ "e * 2 + (e * (real d * 3) + e * (real i * 3))"])
apply auto
done
then have "xf i t > 1 - e/n"
by (rule xfB)
} note **** = this
have xf_le1: "\<And>i. xf i t \<le> 1"
using xf01 t by force
have "g t = e * (\<Sum>i<j. xf i t) + e * (\<Sum>i=j..n. xf i t)"
using j1 jn e
apply (simp add: g_def flip: distrib_left)
apply (subst sum.union_disjoint [symmetric])
apply (auto simp: ivl_disj_un)
done
also have "... \<le> e*j + e * ((Suc n - j)*e/n)"
apply (rule add_mono)
apply (simp_all only: mult_le_cancel_left_pos e)
apply (rule sum_bounded_above [OF xf_le1, where A = "lessThan j", simplified])
using sum_bounded_above [of "{j..n}" "\<lambda>i. xf i t", OF ***]
apply simp
done
also have "... \<le> j*e + e*(n - j + 1)*e/n "
using \<open>1 \<le> n\<close> e by (simp add: field_simps del: of_nat_Suc)
also have "... \<le> j*e + e*e"
using \<open>1 \<le> n\<close> e j1 by (simp add: field_simps del: of_nat_Suc)
also have "... < (j + 1/3)*e"
using e by (auto simp: field_simps)
finally have gj1: "g t < (j + 1 / 3) * e" .
have gj2: "(j - 4/3)*e < g t"
proof (cases "2 \<le> j")
case False
then have "j=1" using j1 by simp
with t gge0 e show ?thesis by force
next
case True
then have "(j - 4/3)*e < (j-1)*e - e^2"
using e by (auto simp: of_nat_diff algebra_simps power2_eq_square)
also have "... < (j-1)*e - ((j - 1)/n) * e^2"
using e True jn by (simp add: power2_eq_square field_simps)
also have "... = e * (j-1) * (1 - e/n)"
by (simp add: power2_eq_square field_simps)
also have "... \<le> e * (\<Sum>i\<le>j-2. xf i t)"
using e
apply simp
apply (rule order_trans [OF _ sum_bounded_below [OF less_imp_le [OF ****]]])
using True
- apply (simp_all add: of_nat_Suc of_nat_diff)
+ apply (simp_all add: of_nat_diff)
done
also have "... \<le> g t"
using jn e
using e xf01 t
apply (simp add: g_def zero_le_mult_iff image_subset_iff sum_nonneg)
apply (rule Groups_Big.sum_mono2, auto)
done
finally show ?thesis .
qed
have "\<bar>f t - g t\<bar> < 2 * e"
using fj1 fj2 gj1 gj2 by (simp add: abs_less_iff field_simps)
}
then show ?thesis
by (rule_tac x=g in bexI) (auto intro: gR)
qed
text\<open>The ``unpretentious'' formulation\<close>
proposition (in function_ring_on) Stone_Weierstrass_basic:
assumes f: "continuous_on S f" and e: "e > 0"
shows "\<exists>g \<in> R. \<forall>x\<in>S. \<bar>f x - g x\<bar> < e"
proof -
have "\<exists>g \<in> R. \<forall>x\<in>S. \<bar>(f x + normf f) - g x\<bar> < 2 * min (e/2) (1/4)"
apply (rule Stone_Weierstrass_special)
apply (rule Limits.continuous_on_add [OF f Topological_Spaces.continuous_on_const])
using normf_upper [OF f] apply force
apply (simp add: e, linarith)
done
then obtain g where "g \<in> R" "\<forall>x\<in>S. \<bar>g x - (f x + normf f)\<bar> < e"
by force
then show ?thesis
apply (rule_tac x="\<lambda>x. g x - normf f" in bexI)
apply (auto simp: algebra_simps intro: diff const)
done
qed
theorem (in function_ring_on) Stone_Weierstrass:
assumes f: "continuous_on S f"
shows "\<exists>F\<in>UNIV \<rightarrow> R. LIM n sequentially. F n :> uniformly_on S f"
proof -
{ fix e::real
assume e: "0 < e"
then obtain N::nat where N: "0 < N" "0 < inverse N" "inverse N < e"
by (auto simp: real_arch_inverse [of e])
{ fix n :: nat and x :: 'a and g :: "'a \<Rightarrow> real"
assume n: "N \<le> n" "\<forall>x\<in>S. \<bar>f x - g x\<bar> < 1 / (1 + real n)"
assume x: "x \<in> S"
have "\<not> real (Suc n) < inverse e"
using \<open>N \<le> n\<close> N using less_imp_inverse_less by force
then have "1 / (1 + real n) \<le> e"
- using e by (simp add: field_simps of_nat_Suc)
+ using e by (simp add: field_simps)
then have "\<bar>f x - g x\<bar> < e"
using n(2) x by auto
} note * = this
have "\<forall>\<^sub>F n in sequentially. \<forall>x\<in>S. \<bar>f x - (SOME g. g \<in> R \<and> (\<forall>x\<in>S. \<bar>f x - g x\<bar> < 1 / (1 + real n))) x\<bar> < e"
apply (rule eventually_sequentiallyI [of N])
apply (auto intro: someI2_bex [OF Stone_Weierstrass_basic [OF f]] *)
done
} then
show ?thesis
apply (rule_tac x="\<lambda>n::nat. SOME g. g \<in> R \<and> (\<forall>x\<in>S. \<bar>f x - g x\<bar> < 1 / (1 + n))" in bexI)
prefer 2 apply (force intro: someI2_bex [OF Stone_Weierstrass_basic [OF f]])
unfolding uniform_limit_iff
apply (auto simp: dist_norm abs_minus_commute)
done
qed
text\<open>A HOL Light formulation\<close>
corollary Stone_Weierstrass_HOL:
fixes R :: "('a::t2_space \<Rightarrow> real) set" and S :: "'a set"
assumes "compact S" "\<And>c. P(\<lambda>x. c::real)"
"\<And>f. P f \<Longrightarrow> continuous_on S f"
"\<And>f g. P(f) \<and> P(g) \<Longrightarrow> P(\<lambda>x. f x + g x)" "\<And>f g. P(f) \<and> P(g) \<Longrightarrow> P(\<lambda>x. f x * g x)"
"\<And>x y. x \<in> S \<and> y \<in> S \<and> x \<noteq> y \<Longrightarrow> \<exists>f. P(f) \<and> f x \<noteq> f y"
"continuous_on S f"
"0 < e"
shows "\<exists>g. P(g) \<and> (\<forall>x \<in> S. \<bar>f x - g x\<bar> < e)"
proof -
interpret PR: function_ring_on "Collect P"
apply unfold_locales
using assms
by auto
show ?thesis
using PR.Stone_Weierstrass_basic [OF \<open>continuous_on S f\<close> \<open>0 < e\<close>]
by blast
qed
subsection \<open>Polynomial functions\<close>
inductive real_polynomial_function :: "('a::real_normed_vector \<Rightarrow> real) \<Rightarrow> bool" where
linear: "bounded_linear f \<Longrightarrow> real_polynomial_function f"
| const: "real_polynomial_function (\<lambda>x. c)"
| add: "\<lbrakk>real_polynomial_function f; real_polynomial_function g\<rbrakk> \<Longrightarrow> real_polynomial_function (\<lambda>x. f x + g x)"
| mult: "\<lbrakk>real_polynomial_function f; real_polynomial_function g\<rbrakk> \<Longrightarrow> real_polynomial_function (\<lambda>x. f x * g x)"
declare real_polynomial_function.intros [intro]
definition\<^marker>\<open>tag important\<close> polynomial_function :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> bool"
where
"polynomial_function p \<equiv> (\<forall>f. bounded_linear f \<longrightarrow> real_polynomial_function (f o p))"
lemma real_polynomial_function_eq: "real_polynomial_function p = polynomial_function p"
unfolding polynomial_function_def
proof
assume "real_polynomial_function p"
then show " \<forall>f. bounded_linear f \<longrightarrow> real_polynomial_function (f \<circ> p)"
proof (induction p rule: real_polynomial_function.induct)
case (linear h) then show ?case
by (auto simp: bounded_linear_compose real_polynomial_function.linear)
next
case (const h) then show ?case
by (simp add: real_polynomial_function.const)
next
case (add h) then show ?case
by (force simp add: bounded_linear_def linear_add real_polynomial_function.add)
next
case (mult h) then show ?case
by (force simp add: real_bounded_linear const real_polynomial_function.mult)
qed
next
assume [rule_format, OF bounded_linear_ident]: "\<forall>f. bounded_linear f \<longrightarrow> real_polynomial_function (f \<circ> p)"
then show "real_polynomial_function p"
by (simp add: o_def)
qed
lemma polynomial_function_const [iff]: "polynomial_function (\<lambda>x. c)"
by (simp add: polynomial_function_def o_def const)
lemma polynomial_function_bounded_linear:
"bounded_linear f \<Longrightarrow> polynomial_function f"
by (simp add: polynomial_function_def o_def bounded_linear_compose real_polynomial_function.linear)
lemma polynomial_function_id [iff]: "polynomial_function(\<lambda>x. x)"
by (simp add: polynomial_function_bounded_linear)
lemma polynomial_function_add [intro]:
"\<lbrakk>polynomial_function f; polynomial_function g\<rbrakk> \<Longrightarrow> polynomial_function (\<lambda>x. f x + g x)"
by (auto simp: polynomial_function_def bounded_linear_def linear_add real_polynomial_function.add o_def)
lemma polynomial_function_mult [intro]:
assumes f: "polynomial_function f" and g: "polynomial_function g"
shows "polynomial_function (\<lambda>x. f x *\<^sub>R g x)"
using g
apply (auto simp: polynomial_function_def bounded_linear_def Real_Vector_Spaces.linear.scaleR const real_polynomial_function.mult o_def)
apply (rule mult)
using f
apply (auto simp: real_polynomial_function_eq)
done
lemma polynomial_function_cmul [intro]:
assumes f: "polynomial_function f"
shows "polynomial_function (\<lambda>x. c *\<^sub>R f x)"
by (rule polynomial_function_mult [OF polynomial_function_const f])
lemma polynomial_function_minus [intro]:
assumes f: "polynomial_function f"
shows "polynomial_function (\<lambda>x. - f x)"
using polynomial_function_cmul [OF f, of "-1"] by simp
lemma polynomial_function_diff [intro]:
"\<lbrakk>polynomial_function f; polynomial_function g\<rbrakk> \<Longrightarrow> polynomial_function (\<lambda>x. f x - g x)"
unfolding add_uminus_conv_diff [symmetric]
by (metis polynomial_function_add polynomial_function_minus)
lemma polynomial_function_sum [intro]:
"\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> polynomial_function (\<lambda>x. f x i)\<rbrakk> \<Longrightarrow> polynomial_function (\<lambda>x. sum (f x) I)"
by (induct I rule: finite_induct) auto
lemma real_polynomial_function_minus [intro]:
"real_polynomial_function f \<Longrightarrow> real_polynomial_function (\<lambda>x. - f x)"
using polynomial_function_minus [of f]
by (simp add: real_polynomial_function_eq)
lemma real_polynomial_function_diff [intro]:
"\<lbrakk>real_polynomial_function f; real_polynomial_function g\<rbrakk> \<Longrightarrow> real_polynomial_function (\<lambda>x. f x - g x)"
using polynomial_function_diff [of f]
by (simp add: real_polynomial_function_eq)
lemma real_polynomial_function_sum [intro]:
"\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> real_polynomial_function (\<lambda>x. f x i)\<rbrakk> \<Longrightarrow> real_polynomial_function (\<lambda>x. sum (f x) I)"
using polynomial_function_sum [of I f]
by (simp add: real_polynomial_function_eq)
lemma real_polynomial_function_power [intro]:
"real_polynomial_function f \<Longrightarrow> real_polynomial_function (\<lambda>x. f x ^ n)"
by (induct n) (simp_all add: const mult)
lemma real_polynomial_function_compose [intro]:
assumes f: "polynomial_function f" and g: "real_polynomial_function g"
shows "real_polynomial_function (g o f)"
using g
apply (induction g rule: real_polynomial_function.induct)
using f
apply (simp_all add: polynomial_function_def o_def const add mult)
done
lemma polynomial_function_compose [intro]:
assumes f: "polynomial_function f" and g: "polynomial_function g"
shows "polynomial_function (g o f)"
using g real_polynomial_function_compose [OF f]
by (auto simp: polynomial_function_def o_def)
lemma sum_max_0:
fixes x::real (*in fact "'a::comm_ring_1"*)
shows "(\<Sum>i\<le>max m n. x^i * (if i \<le> m then a i else 0)) = (\<Sum>i\<le>m. x^i * a i)"
proof -
have "(\<Sum>i\<le>max m n. x^i * (if i \<le> m then a i else 0)) = (\<Sum>i\<le>max m n. (if i \<le> m then x^i * a i else 0))"
by (auto simp: algebra_simps intro: sum.cong)
also have "... = (\<Sum>i\<le>m. (if i \<le> m then x^i * a i else 0))"
by (rule sum.mono_neutral_right) auto
also have "... = (\<Sum>i\<le>m. x^i * a i)"
by (auto simp: algebra_simps intro: sum.cong)
finally show ?thesis .
qed
lemma real_polynomial_function_imp_sum:
assumes "real_polynomial_function f"
shows "\<exists>a n::nat. f = (\<lambda>x. \<Sum>i\<le>n. a i * x ^ i)"
using assms
proof (induct f)
case (linear f)
then show ?case
apply (clarsimp simp add: real_bounded_linear)
apply (rule_tac x="\<lambda>i. if i=0 then 0 else c" in exI)
apply (rule_tac x=1 in exI)
apply (simp add: mult_ac)
done
next
case (const c)
show ?case
apply (rule_tac x="\<lambda>i. c" in exI)
apply (rule_tac x=0 in exI)
- apply (auto simp: mult_ac of_nat_Suc)
+ apply (auto simp: mult_ac)
done
case (add f1 f2)
then obtain a1 n1 a2 n2 where
"f1 = (\<lambda>x. \<Sum>i\<le>n1. a1 i * x ^ i)" "f2 = (\<lambda>x. \<Sum>i\<le>n2. a2 i * x ^ i)"
by auto
then show ?case
apply (rule_tac x="\<lambda>i. (if i \<le> n1 then a1 i else 0) + (if i \<le> n2 then a2 i else 0)" in exI)
apply (rule_tac x="max n1 n2" in exI)
using sum_max_0 [where m=n1 and n=n2] sum_max_0 [where m=n2 and n=n1]
apply (simp add: sum.distrib algebra_simps max.commute)
done
case (mult f1 f2)
then obtain a1 n1 a2 n2 where
"f1 = (\<lambda>x. \<Sum>i\<le>n1. a1 i * x ^ i)" "f2 = (\<lambda>x. \<Sum>i\<le>n2. a2 i * x ^ i)"
by auto
then obtain b1 b2 where
"f1 = (\<lambda>x. \<Sum>i\<le>n1. b1 i * x ^ i)" "f2 = (\<lambda>x. \<Sum>i\<le>n2. b2 i * x ^ i)"
"b1 = (\<lambda>i. if i\<le>n1 then a1 i else 0)" "b2 = (\<lambda>i. if i\<le>n2 then a2 i else 0)"
by auto
then show ?case
apply (rule_tac x="\<lambda>i. \<Sum>k\<le>i. b1 k * b2 (i - k)" in exI)
apply (rule_tac x="n1+n2" in exI)
using polynomial_product [of n1 b1 n2 b2]
apply (simp add: Set_Interval.atLeast0AtMost)
done
qed
lemma real_polynomial_function_iff_sum:
"real_polynomial_function f \<longleftrightarrow> (\<exists>a n::nat. f = (\<lambda>x. \<Sum>i\<le>n. a i * x ^ i))"
apply (rule iffI)
apply (erule real_polynomial_function_imp_sum)
apply (auto simp: linear mult const real_polynomial_function_power real_polynomial_function_sum)
done
lemma polynomial_function_iff_Basis_inner:
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space"
shows "polynomial_function f \<longleftrightarrow> (\<forall>b\<in>Basis. real_polynomial_function (\<lambda>x. inner (f x) b))"
(is "?lhs = ?rhs")
unfolding polynomial_function_def
proof (intro iffI allI impI)
assume "\<forall>h. bounded_linear h \<longrightarrow> real_polynomial_function (h \<circ> f)"
then show ?rhs
by (force simp add: bounded_linear_inner_left o_def)
next
fix h :: "'b \<Rightarrow> real"
assume rp: "\<forall>b\<in>Basis. real_polynomial_function (\<lambda>x. f x \<bullet> b)" and h: "bounded_linear h"
have "real_polynomial_function (h \<circ> (\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b))"
apply (rule real_polynomial_function_compose [OF _ linear [OF h]])
using rp
apply (auto simp: real_polynomial_function_eq polynomial_function_mult)
done
then show "real_polynomial_function (h \<circ> f)"
by (simp add: euclidean_representation_sum_fun)
qed
subsection \<open>Stone-Weierstrass theorem for polynomial functions\<close>
text\<open>First, we need to show that they are continuous, differentiable and separable.\<close>
lemma continuous_real_polymonial_function:
assumes "real_polynomial_function f"
shows "continuous (at x) f"
using assms
by (induct f) (auto simp: linear_continuous_at)
lemma continuous_polymonial_function:
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space"
assumes "polynomial_function f"
shows "continuous (at x) f"
apply (rule euclidean_isCont)
using assms apply (simp add: polynomial_function_iff_Basis_inner)
apply (force dest: continuous_real_polymonial_function intro: isCont_scaleR)
done
lemma continuous_on_polymonial_function:
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space"
assumes "polynomial_function f"
shows "continuous_on S f"
using continuous_polymonial_function [OF assms] continuous_at_imp_continuous_on
by blast
lemma has_real_derivative_polynomial_function:
assumes "real_polynomial_function p"
shows "\<exists>p'. real_polynomial_function p' \<and>
(\<forall>x. (p has_real_derivative (p' x)) (at x))"
using assms
proof (induct p)
case (linear p)
then show ?case
by (force simp: real_bounded_linear const intro!: derivative_eq_intros)
next
case (const c)
show ?case
by (rule_tac x="\<lambda>x. 0" in exI) auto
case (add f1 f2)
then obtain p1 p2 where
"real_polynomial_function p1" "\<And>x. (f1 has_real_derivative p1 x) (at x)"
"real_polynomial_function p2" "\<And>x. (f2 has_real_derivative p2 x) (at x)"
by auto
then show ?case
apply (rule_tac x="\<lambda>x. p1 x + p2 x" in exI)
apply (auto intro!: derivative_eq_intros)
done
case (mult f1 f2)
then obtain p1 p2 where
"real_polynomial_function p1" "\<And>x. (f1 has_real_derivative p1 x) (at x)"
"real_polynomial_function p2" "\<And>x. (f2 has_real_derivative p2 x) (at x)"
by auto
then show ?case
using mult
apply (rule_tac x="\<lambda>x. f1 x * p2 x + f2 x * p1 x" in exI)
apply (auto intro!: derivative_eq_intros)
done
qed
lemma has_vector_derivative_polynomial_function:
fixes p :: "real \<Rightarrow> 'a::euclidean_space"
assumes "polynomial_function p"
obtains p' where "polynomial_function p'" "\<And>x. (p has_vector_derivative (p' x)) (at x)"
proof -
{ fix b :: 'a
assume "b \<in> Basis"
then
obtain p' where p': "real_polynomial_function p'" and pd: "\<And>x. ((\<lambda>x. p x \<bullet> b) has_real_derivative p' x) (at x)"
using assms [unfolded polynomial_function_iff_Basis_inner, rule_format] \<open>b \<in> Basis\<close>
has_real_derivative_polynomial_function
by blast
have "\<exists>q. polynomial_function q \<and> (\<forall>x. ((\<lambda>u. (p u \<bullet> b) *\<^sub>R b) has_vector_derivative q x) (at x))"
apply (rule_tac x="\<lambda>x. p' x *\<^sub>R b" in exI)
using \<open>b \<in> Basis\<close> p'
apply (simp add: polynomial_function_iff_Basis_inner inner_Basis)
apply (auto intro: derivative_eq_intros pd)
done
}
then obtain qf where qf:
"\<And>b. b \<in> Basis \<Longrightarrow> polynomial_function (qf b)"
"\<And>b x. b \<in> Basis \<Longrightarrow> ((\<lambda>u. (p u \<bullet> b) *\<^sub>R b) has_vector_derivative qf b x) (at x)"
by metis
show ?thesis
apply (rule_tac p'="\<lambda>x. \<Sum>b\<in>Basis. qf b x" in that)
apply (force intro: qf)
apply (subst euclidean_representation_sum_fun [of p, symmetric])
apply (auto intro: has_vector_derivative_sum qf)
done
qed
lemma real_polynomial_function_separable:
fixes x :: "'a::euclidean_space"
assumes "x \<noteq> y" shows "\<exists>f. real_polynomial_function f \<and> f x \<noteq> f y"
proof -
have "real_polynomial_function (\<lambda>u. \<Sum>b\<in>Basis. (inner (x-u) b)^2)"
apply (rule real_polynomial_function_sum)
apply (auto simp: algebra_simps real_polynomial_function_power real_polynomial_function_diff
const linear bounded_linear_inner_left)
done
then show ?thesis
apply (intro exI conjI, assumption)
using assms
apply (force simp add: euclidean_eq_iff [of x y] sum_nonneg_eq_0_iff algebra_simps)
done
qed
lemma Stone_Weierstrass_real_polynomial_function:
fixes f :: "'a::euclidean_space \<Rightarrow> real"
assumes "compact S" "continuous_on S f" "0 < e"
obtains g where "real_polynomial_function g" "\<And>x. x \<in> S \<Longrightarrow> \<bar>f x - g x\<bar> < e"
proof -
interpret PR: function_ring_on "Collect real_polynomial_function"
apply unfold_locales
using assms continuous_on_polymonial_function real_polynomial_function_eq
apply (auto intro: real_polynomial_function_separable)
done
show ?thesis
using PR.Stone_Weierstrass_basic [OF \<open>continuous_on S f\<close> \<open>0 < e\<close>] that
by blast
qed
theorem Stone_Weierstrass_polynomial_function:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes S: "compact S"
and f: "continuous_on S f"
and e: "0 < e"
shows "\<exists>g. polynomial_function g \<and> (\<forall>x \<in> S. norm(f x - g x) < e)"
proof -
{ fix b :: 'b
assume "b \<in> Basis"
have "\<exists>p. real_polynomial_function p \<and> (\<forall>x \<in> S. \<bar>f x \<bullet> b - p x\<bar> < e / DIM('b))"
apply (rule exE [OF Stone_Weierstrass_real_polynomial_function [OF S _, of "\<lambda>x. f x \<bullet> b" "e / card Basis"]])
using e f
apply (auto intro: continuous_intros)
done
}
then obtain pf where pf:
"\<And>b. b \<in> Basis \<Longrightarrow> real_polynomial_function (pf b) \<and> (\<forall>x \<in> S. \<bar>f x \<bullet> b - pf b x\<bar> < e / DIM('b))"
apply (rule bchoice [rule_format, THEN exE])
apply assumption
apply (force simp add: intro: that)
done
have "polynomial_function (\<lambda>x. \<Sum>b\<in>Basis. pf b x *\<^sub>R b)"
using pf
by (simp add: polynomial_function_sum polynomial_function_mult real_polynomial_function_eq)
moreover
{ fix x
assume "x \<in> S"
have "norm (\<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b - pf b x *\<^sub>R b) \<le> (\<Sum>b\<in>Basis. norm ((f x \<bullet> b) *\<^sub>R b - pf b x *\<^sub>R b))"
by (rule norm_sum)
also have "... < of_nat DIM('b) * (e / DIM('b))"
apply (rule sum_bounded_above_strict)
apply (simp add: Real_Vector_Spaces.scaleR_diff_left [symmetric] pf \<open>x \<in> S\<close>)
apply (rule DIM_positive)
done
also have "... = e"
by (simp add: field_simps)
finally have "norm (\<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b - pf b x *\<^sub>R b) < e" .
}
ultimately
show ?thesis
apply (subst euclidean_representation_sum_fun [of f, symmetric])
apply (rule_tac x="\<lambda>x. \<Sum>b\<in>Basis. pf b x *\<^sub>R b" in exI)
apply (auto simp flip: sum_subtractf)
done
qed
proposition Stone_Weierstrass_uniform_limit:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes S: "compact S"
and f: "continuous_on S f"
obtains g where "uniform_limit S g f sequentially" "\<And>n. polynomial_function (g n)"
proof -
have pos: "inverse (Suc n) > 0" for n by auto
obtain g where g: "\<And>n. polynomial_function (g n)" "\<And>x n. x \<in> S \<Longrightarrow> norm(f x - g n x) < inverse (Suc n)"
using Stone_Weierstrass_polynomial_function[OF S f pos]
by metis
have "uniform_limit S g f sequentially"
proof (rule uniform_limitI)
fix e::real assume "0 < e"
with LIMSEQ_inverse_real_of_nat have "\<forall>\<^sub>F n in sequentially. inverse (Suc n) < e"
by (rule order_tendstoD)
moreover have "\<forall>\<^sub>F n in sequentially. \<forall>x\<in>S. dist (g n x) (f x) < inverse (Suc n)"
using g by (simp add: dist_norm norm_minus_commute)
ultimately show "\<forall>\<^sub>F n in sequentially. \<forall>x\<in>S. dist (g n x) (f x) < e"
by (eventually_elim) auto
qed
then show ?thesis using g(1) ..
qed
subsection\<open>Polynomial functions as paths\<close>
text\<open>One application is to pick a smooth approximation to a path,
or just pick a smooth path anyway in an open connected set\<close>
lemma path_polynomial_function:
fixes g :: "real \<Rightarrow> 'b::euclidean_space"
shows "polynomial_function g \<Longrightarrow> path g"
by (simp add: path_def continuous_on_polymonial_function)
lemma path_approx_polynomial_function:
fixes g :: "real \<Rightarrow> 'b::euclidean_space"
assumes "path g" "0 < e"
shows "\<exists>p. polynomial_function p \<and>
pathstart p = pathstart g \<and>
pathfinish p = pathfinish g \<and>
(\<forall>t \<in> {0..1}. norm(p t - g t) < e)"
proof -
obtain q where poq: "polynomial_function q" and noq: "\<And>x. x \<in> {0..1} \<Longrightarrow> norm (g x - q x) < e/4"
using Stone_Weierstrass_polynomial_function [of "{0..1}" g "e/4"] assms
by (auto simp: path_def)
have pf: "polynomial_function (\<lambda>t. q t + (g 0 - q 0) + t *\<^sub>R (g 1 - q 1 - (g 0 - q 0)))"
by (force simp add: poq)
have *: "\<And>t. t \<in> {0..1} \<Longrightarrow> norm (((q t - g t) + (g 0 - q 0)) + (t *\<^sub>R (g 1 - q 1) + t *\<^sub>R (q 0 - g 0))) < (e/4 + e/4) + (e/4+e/4)"
apply (intro Real_Vector_Spaces.norm_add_less)
using noq
apply (auto simp: norm_minus_commute intro: le_less_trans [OF mult_left_le_one_le noq] simp del: less_divide_eq_numeral1)
done
show ?thesis
apply (intro exI conjI)
apply (rule pf)
using *
apply (auto simp add: pathstart_def pathfinish_def algebra_simps)
done
qed
proposition connected_open_polynomial_connected:
fixes S :: "'a::euclidean_space set"
assumes S: "open S" "connected S"
and "x \<in> S" "y \<in> S"
shows "\<exists>g. polynomial_function g \<and> path_image g \<subseteq> S \<and>
pathstart g = x \<and> pathfinish g = y"
proof -
have "path_connected S" using assms
by (simp add: connected_open_path_connected)
with \<open>x \<in> S\<close> \<open>y \<in> S\<close> obtain p where p: "path p" "path_image p \<subseteq> S" "pathstart p = x" "pathfinish p = y"
by (force simp: path_connected_def)
have "\<exists>e. 0 < e \<and> (\<forall>x \<in> path_image p. ball x e \<subseteq> S)"
proof (cases "S = UNIV")
case True then show ?thesis
by (simp add: gt_ex)
next
case False
then have "- S \<noteq> {}" by blast
then show ?thesis
apply (rule_tac x="setdist (path_image p) (-S)" in exI)
using S p
apply (simp add: setdist_gt_0_compact_closed compact_path_image open_closed)
using setdist_le_dist [of _ "path_image p" _ "-S"]
by fastforce
qed
then obtain e where "0 < e"and eb: "\<And>x. x \<in> path_image p \<Longrightarrow> ball x e \<subseteq> S"
by auto
show ?thesis
using path_approx_polynomial_function [OF \<open>path p\<close> \<open>0 < e\<close>]
apply clarify
apply (intro exI conjI, assumption)
using p
apply (fastforce simp add: dist_norm path_image_def norm_minus_commute intro: eb [THEN subsetD])+
done
qed
lemma differentiable_componentwise_within:
"f differentiable (at a within S) \<longleftrightarrow>
(\<forall>i \<in> Basis. (\<lambda>x. f x \<bullet> i) differentiable at a within S)"
proof -
{ assume "\<forall>i\<in>Basis. \<exists>D. ((\<lambda>x. f x \<bullet> i) has_derivative D) (at a within S)"
then obtain f' where f':
"\<And>i. i \<in> Basis \<Longrightarrow> ((\<lambda>x. f x \<bullet> i) has_derivative f' i) (at a within S)"
by metis
have eq: "(\<lambda>x. (\<Sum>j\<in>Basis. f' j x *\<^sub>R j) \<bullet> i) = f' i" if "i \<in> Basis" for i
using that by (simp add: inner_add_left inner_add_right)
have "\<exists>D. \<forall>i\<in>Basis. ((\<lambda>x. f x \<bullet> i) has_derivative (\<lambda>x. D x \<bullet> i)) (at a within S)"
apply (rule_tac x="\<lambda>x::'a. (\<Sum>j\<in>Basis. f' j x *\<^sub>R j) :: 'b" in exI)
apply (simp add: eq f')
done
}
then show ?thesis
apply (simp add: differentiable_def)
using has_derivative_componentwise_within
by blast
qed
lemma polynomial_function_inner [intro]:
fixes i :: "'a::euclidean_space"
shows "polynomial_function g \<Longrightarrow> polynomial_function (\<lambda>x. g x \<bullet> i)"
apply (subst euclidean_representation [where x=i, symmetric])
apply (force simp: inner_sum_right polynomial_function_iff_Basis_inner polynomial_function_sum)
done
text\<open> Differentiability of real and vector polynomial functions.\<close>
lemma differentiable_at_real_polynomial_function:
"real_polynomial_function f \<Longrightarrow> f differentiable (at a within S)"
by (induction f rule: real_polynomial_function.induct)
(simp_all add: bounded_linear_imp_differentiable)
lemma differentiable_on_real_polynomial_function:
"real_polynomial_function p \<Longrightarrow> p differentiable_on S"
by (simp add: differentiable_at_imp_differentiable_on differentiable_at_real_polynomial_function)
lemma differentiable_at_polynomial_function:
fixes f :: "_ \<Rightarrow> 'a::euclidean_space"
shows "polynomial_function f \<Longrightarrow> f differentiable (at a within S)"
by (metis differentiable_at_real_polynomial_function polynomial_function_iff_Basis_inner differentiable_componentwise_within)
lemma differentiable_on_polynomial_function:
fixes f :: "_ \<Rightarrow> 'a::euclidean_space"
shows "polynomial_function f \<Longrightarrow> f differentiable_on S"
by (simp add: differentiable_at_polynomial_function differentiable_on_def)
lemma vector_eq_dot_span:
assumes "x \<in> span B" "y \<in> span B" and i: "\<And>i. i \<in> B \<Longrightarrow> i \<bullet> x = i \<bullet> y"
shows "x = y"
proof -
have "\<And>i. i \<in> B \<Longrightarrow> orthogonal (x - y) i"
by (simp add: i inner_commute inner_diff_right orthogonal_def)
moreover have "x - y \<in> span B"
by (simp add: assms span_diff)
ultimately have "x - y = 0"
using orthogonal_to_span orthogonal_self by blast
then show ?thesis by simp
qed
lemma orthonormal_basis_expand:
assumes B: "pairwise orthogonal B"
and 1: "\<And>i. i \<in> B \<Longrightarrow> norm i = 1"
and "x \<in> span B"
and "finite B"
shows "(\<Sum>i\<in>B. (x \<bullet> i) *\<^sub>R i) = x"
proof (rule vector_eq_dot_span [OF _ \<open>x \<in> span B\<close>])
show "(\<Sum>i\<in>B. (x \<bullet> i) *\<^sub>R i) \<in> span B"
by (simp add: span_clauses span_sum)
show "i \<bullet> (\<Sum>i\<in>B. (x \<bullet> i) *\<^sub>R i) = i \<bullet> x" if "i \<in> B" for i
proof -
have [simp]: "i \<bullet> j = (if j = i then 1 else 0)" if "j \<in> B" for j
using B 1 that \<open>i \<in> B\<close>
by (force simp: norm_eq_1 orthogonal_def pairwise_def)
have "i \<bullet> (\<Sum>i\<in>B. (x \<bullet> i) *\<^sub>R i) = (\<Sum>j\<in>B. x \<bullet> j * (i \<bullet> j))"
by (simp add: inner_sum_right)
also have "... = (\<Sum>j\<in>B. if j = i then x \<bullet> i else 0)"
by (rule sum.cong; simp)
also have "... = i \<bullet> x"
- by (simp add: \<open>finite B\<close> that inner_commute sum.delta)
+ by (simp add: \<open>finite B\<close> that inner_commute)
finally show ?thesis .
qed
qed
theorem Stone_Weierstrass_polynomial_function_subspace:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "compact S"
and contf: "continuous_on S f"
and "0 < e"
and "subspace T" "f ` S \<subseteq> T"
obtains g where "polynomial_function g" "g ` S \<subseteq> T"
"\<And>x. x \<in> S \<Longrightarrow> norm(f x - g x) < e"
proof -
obtain B where "B \<subseteq> T" and orthB: "pairwise orthogonal B"
and B1: "\<And>x. x \<in> B \<Longrightarrow> norm x = 1"
and "independent B" and cardB: "card B = dim T"
and spanB: "span B = T"
using orthonormal_basis_subspace \<open>subspace T\<close> by metis
then have "finite B"
by (simp add: independent_imp_finite)
then obtain n::nat and b where "B = b ` {i. i < n}" "inj_on b {i. i < n}"
using finite_imp_nat_seg_image_inj_on by metis
with cardB have "n = card B" "dim T = n"
by (auto simp: card_image)
have fx: "(\<Sum>i\<in>B. (f x \<bullet> i) *\<^sub>R i) = f x" if "x \<in> S" for x
apply (rule orthonormal_basis_expand [OF orthB B1 _ \<open>finite B\<close>])
using \<open>f ` S \<subseteq> T\<close> spanB that by auto
have cont: "continuous_on S (\<lambda>x. \<Sum>i\<in>B. (f x \<bullet> i) *\<^sub>R i)"
by (intro continuous_intros contf)
obtain g where "polynomial_function g"
and g: "\<And>x. x \<in> S \<Longrightarrow> norm ((\<Sum>i\<in>B. (f x \<bullet> i) *\<^sub>R i) - g x) < e / (n+2)"
using Stone_Weierstrass_polynomial_function [OF \<open>compact S\<close> cont, of "e / real (n + 2)"] \<open>0 < e\<close>
by auto
with fx have g: "\<And>x. x \<in> S \<Longrightarrow> norm (f x - g x) < e / (n+2)"
by auto
show ?thesis
proof
show "polynomial_function (\<lambda>x. \<Sum>i\<in>B. (g x \<bullet> i) *\<^sub>R i)"
apply (rule polynomial_function_sum)
apply (simp add: \<open>finite B\<close>)
using \<open>polynomial_function g\<close> by auto
show "(\<lambda>x. \<Sum>i\<in>B. (g x \<bullet> i) *\<^sub>R i) ` S \<subseteq> T"
using \<open>B \<subseteq> T\<close>
by (blast intro: subspace_sum subspace_mul \<open>subspace T\<close>)
show "norm (f x - (\<Sum>i\<in>B. (g x \<bullet> i) *\<^sub>R i)) < e" if "x \<in> S" for x
proof -
have orth': "pairwise (\<lambda>i j. orthogonal ((f x \<bullet> i) *\<^sub>R i - (g x \<bullet> i) *\<^sub>R i)
((f x \<bullet> j) *\<^sub>R j - (g x \<bullet> j) *\<^sub>R j)) B"
apply (rule pairwise_mono [OF orthB])
apply (auto simp: orthogonal_def inner_diff_right inner_diff_left)
done
then have "(norm (\<Sum>i\<in>B. (f x \<bullet> i) *\<^sub>R i - (g x \<bullet> i) *\<^sub>R i))\<^sup>2 =
(\<Sum>i\<in>B. (norm ((f x \<bullet> i) *\<^sub>R i - (g x \<bullet> i) *\<^sub>R i))\<^sup>2)"
by (simp add: norm_sum_Pythagorean [OF \<open>finite B\<close> orth'])
also have "... = (\<Sum>i\<in>B. (norm (((f x - g x) \<bullet> i) *\<^sub>R i))\<^sup>2)"
by (simp add: algebra_simps)
also have "... \<le> (\<Sum>i\<in>B. (norm (f x - g x))\<^sup>2)"
apply (rule sum_mono)
apply (simp add: B1)
apply (rule order_trans [OF Cauchy_Schwarz_ineq])
by (simp add: B1 dot_square_norm)
also have "... = n * norm (f x - g x)^2"
by (simp add: \<open>n = card B\<close>)
also have "... \<le> n * (e / (n+2))^2"
apply (rule mult_left_mono)
apply (meson dual_order.order_iff_strict g norm_ge_zero power_mono that, simp)
done
also have "... \<le> e^2 / (n+2)"
using \<open>0 < e\<close> by (simp add: divide_simps power2_eq_square)
also have "... < e^2"
using \<open>0 < e\<close> by (simp add: divide_simps)
finally have "(norm (\<Sum>i\<in>B. (f x \<bullet> i) *\<^sub>R i - (g x \<bullet> i) *\<^sub>R i))\<^sup>2 < e^2" .
then have "(norm (\<Sum>i\<in>B. (f x \<bullet> i) *\<^sub>R i - (g x \<bullet> i) *\<^sub>R i)) < e"
apply (rule power2_less_imp_less)
using \<open>0 < e\<close> by auto
then show ?thesis
using fx that by (simp add: sum_subtractf)
qed
qed
qed
hide_fact linear add mult const
end

File Metadata

Mime Type
application/octet-stream
Expires
Sun, May 5, 5:59 PM (1 d, 23 h)
Storage Engine
chunks
Storage Format
Chunks
Storage Handle
J1hmpaMpo9Iw
Default Alt Text
(4 MB)

Event Timeline