Theory Improper_Integral
section ‹Continuity of the indefinite integral; improper integral theorem›
theory "Improper_Integral"
imports Equivalence_Lebesgue_Henstock_Integration
begin
subsection ‹Equiintegrability›
text‹The definition here only really makes sense for an elementary set.
We just use compact intervals in applications below.›
definition equiintegrable_on (infixr "equiintegrable'_on" 46)
where "F equiintegrable_on I ≡
(∀f ∈ F. f integrable_on I) ∧
(∀e > 0. ∃γ. gauge γ ∧
(∀f 𝒟. f ∈ F ∧ 𝒟 tagged_division_of I ∧ γ fine 𝒟
⟶ norm ((∑(x,K) ∈ 𝒟. content K *⇩R f x) - integral I f) < e))"
lemma equiintegrable_on_integrable:
"⟦F equiintegrable_on I; f ∈ F⟧ ⟹ f integrable_on I"
using equiintegrable_on_def by metis
lemma equiintegrable_on_sing [simp]:
"{f} equiintegrable_on cbox a b ⟷ f integrable_on cbox a b"
by (simp add: equiintegrable_on_def has_integral_integral has_integral integrable_on_def)
lemma equiintegrable_on_subset: "⟦F equiintegrable_on I; G ⊆ F⟧ ⟹ 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 ∪ G) equiintegrable_on I"
unfolding equiintegrable_on_def
proof (intro conjI impI allI)
show "∀f∈F ∪ G. f integrable_on I"
using assms unfolding equiintegrable_on_def by blast
show "∃γ. gauge γ ∧
(∀f 𝒟. f ∈ F ∪ G ∧
𝒟 tagged_division_of I ∧ γ fine 𝒟 ⟶
norm ((∑(x,K) ∈ 𝒟. content K *⇩R f x) - integral I f) < ε)"
if "ε > 0" for ε
proof -
obtain γ1 where "gauge γ1"
and γ1: "⋀f 𝒟. f ∈ F ∧ 𝒟 tagged_division_of I ∧ γ1 fine 𝒟
⟹ norm ((∑(x,K) ∈ 𝒟. content K *⇩R f x) - integral I f) < ε"
using assms ‹ε > 0› unfolding equiintegrable_on_def by auto
obtain γ2 where "gauge γ2"
and γ2: "⋀f 𝒟. f ∈ G ∧ 𝒟 tagged_division_of I ∧ γ2 fine 𝒟
⟹ norm ((∑(x,K) ∈ 𝒟. content K *⇩R f x) - integral I f) < ε"
using assms ‹ε > 0› unfolding equiintegrable_on_def by auto
have "gauge (λx. γ1 x ∩ γ2 x)"
using ‹gauge γ1› ‹gauge γ2› by blast
moreover have "∀f 𝒟. f ∈ F ∪ G ∧ 𝒟 tagged_division_of I ∧ (λx. γ1 x ∩ γ2 x) fine 𝒟 ⟶
norm ((∑(x,K) ∈ 𝒟. content K *⇩R f x) - integral I f) < ε"
using γ1 γ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 "(⋃c ∈ {-k..k}. ⋃f ∈ F. {(λx. c *⇩R f x)}) equiintegrable_on I"
unfolding equiintegrable_on_def
proof (intro conjI impI allI ballI)
show "f integrable_on I"
if "f ∈ (⋃c∈{- k..k}. ⋃f∈F. {λx. c *⇩R f x})"
for f :: "'a ⇒ 'b"
using that assms equiintegrable_on_integrable integrable_cmul by blast
show "∃γ. gauge γ ∧ (∀f 𝒟. f ∈ (⋃c∈{- k..k}. ⋃f∈F. {λx. c *⇩R f x}) ∧ 𝒟 tagged_division_of I
∧ γ fine 𝒟 ⟶ norm ((∑(x, K)∈𝒟. content K *⇩R f x) - integral I f) < ε)"
if "ε > 0" for ε
proof -
obtain γ where "gauge γ"
and γ: "⋀f 𝒟. ⟦f ∈ F; 𝒟 tagged_division_of I; γ fine 𝒟⟧
⟹ norm ((∑(x,K) ∈ 𝒟. content K *⇩R f x) - integral I f) < ε / (¦k¦ + 1)"
using assms ‹ε > 0› 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 ((∑(x, K)∈𝒟. content K *⇩R c *⇩R (f x)) - integral I (λx. c *⇩R f x)) < ε"
if c: "c ∈ {- k..k}"
and "f ∈ F" "𝒟 tagged_division_of I" "γ fine 𝒟"
for 𝒟 c f
proof -
have "norm ((∑x∈𝒟. case x of (x, K) ⇒ content K *⇩R c *⇩R f x) - integral I (λx. c *⇩R f x))
= ¦c¦ * norm ((∑x∈𝒟. case x of (x, K) ⇒ content K *⇩R f x) - integral I f)"
by (simp add: algebra_simps scale_sum_right case_prod_unfold flip: norm_scaleR)
also have "… ≤ (¦k¦ + 1) * norm ((∑x∈𝒟. case x of (x, K) ⇒ content K *⇩R f x) - integral I f)"
using c by (auto simp: mult_right_mono)
also have "… < (¦k¦ + 1) * (ε / (¦k¦ + 1))"
by (rule mult_strict_left_mono) (use γ less_eq_real_def that in auto)
also have "… = ε"
by auto
finally show ?thesis .
qed
ultimately show ?thesis
by (rule_tac x="γ" in exI) auto
qed
qed
lemma equiintegrable_add:
assumes F: "F equiintegrable_on I" and G: "G equiintegrable_on I"
shows "(⋃f ∈ F. ⋃g ∈ G. {(λ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 ∈ (⋃f∈F. ⋃g∈G. {λx. f x + g x})" for f
using that equiintegrable_on_integrable assms by (auto intro: integrable_add)
show "∃γ. gauge γ ∧ (∀f 𝒟. f ∈ (⋃f∈F. ⋃g∈G. {λx. f x + g x}) ∧ 𝒟 tagged_division_of I
∧ γ fine 𝒟 ⟶ norm ((∑(x, K)∈𝒟. content K *⇩R f x) - integral I f) < ε)"
if "ε > 0" for ε
proof -
obtain γ1 where "gauge γ1"
and γ1: "⋀f 𝒟. ⟦f ∈ F; 𝒟 tagged_division_of I; γ1 fine 𝒟⟧
⟹ norm ((∑(x,K) ∈ 𝒟. content K *⇩R f x) - integral I f) < ε/2"
using assms ‹ε > 0› unfolding equiintegrable_on_def by (meson half_gt_zero_iff)
obtain γ2 where "gauge γ2"
and γ2: "⋀g 𝒟. ⟦g ∈ G; 𝒟 tagged_division_of I; γ2 fine 𝒟⟧
⟹ norm ((∑(x,K) ∈ 𝒟. content K *⇩R g x) - integral I g) < ε/2"
using assms ‹ε > 0› unfolding equiintegrable_on_def by (meson half_gt_zero_iff)
have "gauge (λx. γ1 x ∩ γ2 x)"
using ‹gauge γ1› ‹gauge γ2› by blast
moreover have "norm ((∑(x,K) ∈ 𝒟. content K *⇩R h x) - integral I h) < ε"
if h: "h ∈ (⋃f∈F. ⋃g∈G. {λx. f x + g x})"
and 𝒟: "𝒟 tagged_division_of I" and fine: "(λx. γ1 x ∩ γ2 x) fine 𝒟"
for h 𝒟
proof -
obtain f g where "f ∈ F" "g ∈ G" and heq: "h = (λ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 ((∑(x,K) ∈ 𝒟. content K *⇩R h x) - integral I h)
= norm ((∑(x,K) ∈ 𝒟. content K *⇩R f x + content K *⇩R g x) - (integral I f + integral I g))"
by (simp add: heq algebra_simps integral_add int)
also have "… = norm (((∑(x,K) ∈ 𝒟. content K *⇩R f x) - integral I f + (∑(x,K) ∈ 𝒟. content K *⇩R g x) - integral I g))"
by (simp add: sum.distrib algebra_simps case_prod_unfold)
also have "… ≤ norm ((∑(x,K) ∈ 𝒟. content K *⇩R f x) - integral I f) + norm ((∑(x,K) ∈ 𝒟. content K *⇩R g x) - integral I g)"
by (metis (mono_tags) add_diff_eq norm_triangle_ineq)
also have "… < ε/2 + ε/2"
using γ1 [OF ‹f ∈ F› 𝒟] γ2 [OF ‹g ∈ G› 𝒟] 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 "(⋃f ∈ F. {(λ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 "(⋃f ∈ F. ⋃g ∈ G. {(λ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 ⇒ 'b::euclidean_space) set"
assumes "F equiintegrable_on cbox a b"
shows "(⋃I ∈ Collect finite. ⋃c ∈ {c. (∀i ∈ I. c i ≥ 0) ∧ sum c I = 1}.
⋃f ∈ I → F. {(λx. sum (λi::'j. c i *⇩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 ∈ ?G" for f
using that assms by (auto simp: equiintegrable_on_def intro!: integrable_sum integrable_cmul)
show "∃γ. gauge γ
∧ (∀g 𝒟. g ∈ ?G ∧ 𝒟 tagged_division_of cbox a b ∧ γ fine 𝒟
⟶ norm ((∑(x,K) ∈ 𝒟. content K *⇩R g x) - integral (cbox a b) g) < ε)"
if "ε > 0" for ε
proof -
obtain γ where "gauge γ"
and γ: "⋀f 𝒟. ⟦f ∈ F; 𝒟 tagged_division_of cbox a b; γ fine 𝒟⟧
⟹ norm ((∑(x,K) ∈ 𝒟. content K *⇩R f x) - integral (cbox a b) f) < ε / 2"
using assms ‹ε > 0› unfolding equiintegrable_on_def by (meson half_gt_zero_iff)
moreover have "norm ((∑(x,K) ∈ 𝒟. content K *⇩R g x) - integral (cbox a b) g) < ε"
if g: "g ∈ ?G"
and 𝒟: "𝒟 tagged_division_of cbox a b"
and fine: "γ fine 𝒟"
for 𝒟 g
proof -
obtain I c f where "finite I" and 0: "⋀i::'j. i ∈ I ⟹ 0 ≤ c i"
and 1: "sum c I = 1" and f: "f ∈ I → F" and geq: "g = (λx. ∑i∈I. c i *⇩R f i x)"
using g by auto
have fi_int: "f i integrable_on cbox a b" if "i ∈ I" for i
by (metis Pi_iff assms equiintegrable_on_def f that)
have *: "integral (cbox a b) (λx. c i *⇩R f i x) = (∑(x, K)∈𝒟. integral K (λx. c i *⇩R f i x))"
if "i ∈ 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 𝒟 integrable_cmul integral_combine_tagged_division_topdown)
qed
have "finite 𝒟"
using 𝒟 by blast
have swap: "(∑(x,K)∈𝒟. content K *⇩R (∑i∈I. c i *⇩R f i x))
= (∑i∈I. c i *⇩R (∑(x,K)∈𝒟. content K *⇩R f i x))"
by (simp add: scale_sum_right case_prod_unfold algebra_simps) (rule sum.swap)
have "norm ((∑(x, K)∈𝒟. content K *⇩R g x) - integral (cbox a b) g)
= norm ((∑i∈I. c i *⇩R ((∑(x,K)∈𝒟. content K *⇩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 ‹finite I› sum_subtractf flip: sum_diff)
also have "… ≤ (∑i∈I. c i * ε / 2)"
proof (rule sum_norm_le)
show "norm (c i *⇩R ((∑(xa, K)∈𝒟. content K *⇩R f i xa) - integral (cbox a b) (f i))) ≤ c i * ε / 2"
if "i ∈ I" for i
proof -
have "norm ((∑(x, K)∈𝒟. content K *⇩R f i x) - integral (cbox a b) (f i)) ≤ ε/2"
using γ [OF _ 𝒟 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 "… < ε"
using 1 ‹ε > 0› by (simp add: flip: sum_divide_distrib sum_distrib_right)
finally show ?thesis .
qed
ultimately show ?thesis
by (rule_tac x="γ" in exI) auto
qed
qed
corollary equiintegrable_sum_real:
fixes F :: "(real ⇒ 'b::euclidean_space) set"
assumes "F equiintegrable_on {a..b}"
shows "(⋃I ∈ Collect finite. ⋃c ∈ {c. (∀i ∈ I. c i ≥ 0) ∧ sum c I = 1}.
⋃f ∈ I → F. {(λx. sum (λi. c i *⇩R f i x) I)})
equiintegrable_on {a..b}"
using equiintegrable_sum [of F a b] assms by auto
text‹ Basic combining theorems for the interval of integration.›
lemma equiintegrable_on_null [simp]:
"content(cbox a b) = 0 ⟹ F equiintegrable_on cbox a b"
unfolding equiintegrable_on_def
by (metis diff_zero gauge_trivial integrable_on_null integral_null norm_zero sum_content_null)
text‹ Main limit theorem for an equiintegrable sequence.›
theorem equiintegrable_limit:
fixes g :: "'a :: euclidean_space ⇒ 'b :: banach"
assumes feq: "range f equiintegrable_on cbox a b"
and to_g: "⋀x. x ∈ cbox a b ⟹ (λn. f n x) ⇢ g x"
shows "g integrable_on cbox a b ∧ (λn. integral (cbox a b) (f n)) ⇢ integral (cbox a b) g"
proof -
have "Cauchy (λ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 γ where "gauge γ"
and γ: "⋀n 𝒟. ⟦𝒟 tagged_division_of cbox a b; γ fine 𝒟⟧
⟹ norm((∑(x,K) ∈ 𝒟. content K *⇩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 𝒟 where 𝒟: "𝒟 tagged_division_of (cbox a b)" and "γ fine 𝒟" "finite 𝒟"
by (meson ‹gauge γ› fine_division_exists tagged_division_of_finite)
with γ have δT: "⋀n. dist ((∑(x,K)∈𝒟. content K *⇩R f n x)) (integral (cbox a b) (f n)) < e/3"
by (force simp: dist_norm)
have "(λn. ∑(x,K)∈𝒟. content K *⇩R f n x) ⇢ (∑(x,K)∈𝒟. content K *⇩R g x)"
using 𝒟 to_g by (auto intro!: tendsto_sum tendsto_scaleR)
then have "Cauchy (λn. ∑(x,K)∈𝒟. content K *⇩R f n x)"
by (meson convergent_eq_Cauchy)
with e3 obtain M where
M: "⋀m n. ⟦m≥M; n≥M⟧ ⟹ dist (∑(x,K)∈𝒟. content K *⇩R f m x) (∑(x,K)∈𝒟. content K *⇩R f n x)
< e/3"
unfolding Cauchy_def by blast
have "⋀m n. ⟦m≥M; n≥M;
dist (∑(x,K)∈𝒟. content K *⇩R f m x) (∑(x,K)∈𝒟. content K *⇩R f n x) < e/3⟧
⟹ dist (integral (cbox a b) (f m)) (integral (cbox a b) (f n)) < e"
by (metis δT dist_commute dist_triangle_third [OF _ _ δT])
then show "∃M. ∀m≥M. ∀n≥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: "(λn. integral (cbox a b) (f n)) ⇢ 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 γ where "gauge γ"
and γ: "⋀n 𝒟. ⟦𝒟 tagged_division_of cbox a b; γ fine 𝒟⟧
⟹ norm((∑(x,K)∈𝒟. content K *⇩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 ((∑(x,K)∈𝒟. content K *⇩R g x) - L) < e"
if "𝒟 tagged_division_of cbox a b" "γ fine 𝒟" for 𝒟
proof -
have "norm ((∑(x,K)∈𝒟. content K *⇩R g x) - L) ≤ e/2"
proof (rule Lim_norm_ubound)
show "(λn. (∑(x,K)∈𝒟. content K *⇩R f n x) - integral (cbox a b) (f n)) ⇢ (∑(x,K)∈𝒟. content K *⇩R g x) - L"
using to_g that L
by (intro tendsto_diff tendsto_sum) (auto simp: tag_in_interval tendsto_scaleR)
show "∀⇩F n in sequentially.
norm ((∑(x,K) ∈ 𝒟. content K *⇩R f n x) - integral (cbox a b) (f n)) ≤ e/2"
by (intro eventuallyI less_imp_le γ that)
qed auto
with ‹0 < e› show ?thesis
by linarith
qed
ultimately
show "∃γ. gauge γ ∧
(∀𝒟. 𝒟 tagged_division_of cbox a b ∧ γ fine 𝒟 ⟶
norm ((∑(x,K)∈𝒟. content K *⇩R g x) - L) < e)"
by meson
qed
with L show ?thesis
by (simp add: ‹(λn. integral (cbox a b) (f n)) ⇢ L› has_integral_integrable_integral)
qed
lemma equiintegrable_reflect:
assumes "F equiintegrable_on cbox a b"
shows "(λf. f ∘ uminus) ` F equiintegrable_on cbox (-b) (-a)"
proof -
have §: "∃γ. gauge γ ∧
(∀f 𝒟. f ∈ (λf. f ∘ uminus) ` F ∧ 𝒟 tagged_division_of cbox (- b) (- a) ∧ γ fine 𝒟 ⟶
norm ((∑(x,K) ∈ 𝒟. content K *⇩R f x) - integral (cbox (- b) (- a)) f) < e)"
if "gauge γ" and
γ: "⋀f 𝒟. ⟦f ∈ F; 𝒟 tagged_division_of cbox a b; γ fine 𝒟⟧ ⟹
norm ((∑(x,K) ∈ 𝒟. content K *⇩R f x) - integral (cbox a b) f) < e" for e γ
proof (intro exI, safe)
show "gauge (λx. uminus ` γ (-x))"
by (metis ‹gauge γ› gauge_reflect)
show "norm ((∑(x,K) ∈ 𝒟. content K *⇩R (f ∘ uminus) x) - integral (cbox (- b) (- a)) (f ∘ uminus)) < e"
if "f ∈ F" and tag: "𝒟 tagged_division_of cbox (- b) (- a)"
and fine: "(λx. uminus ` γ (- x)) fine 𝒟" for f 𝒟
proof -
have 1: "(λ(x,K). (- x, uminus ` K)) ` 𝒟 tagged_partial_division_of cbox a b"
if "𝒟 tagged_partial_division_of cbox (- b) (- a)"
proof -
have "- y ∈ cbox a b"
if "⋀x K. (x,K) ∈ 𝒟 ⟹ x ∈ K ∧ K ⊆ cbox (- b) (- a) ∧ (∃a b. K = cbox a b)"
"(x, Y) ∈ 𝒟" "y ∈ Y" for x Y y
proof -
have "y ∈ uminus ` cbox a b"
using that by auto
then show "- y ∈ 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: "∃K. (∃x. (x,K) ∈ (λ(x,K). (- x, uminus ` K)) ` 𝒟) ∧ x ∈ K"
if "⋃{K. ∃x. (x,K) ∈ 𝒟} = cbox (- b) (- a)" "x ∈ cbox a b" for x
proof -
have xm: "x ∈ uminus ` ⋃{A. ∃a. (a, A) ∈ 𝒟}"
by (simp add: that)
then obtain a X where "-x ∈ X" "(a, X) ∈ 𝒟"
by auto
then show ?thesis
by (metis (no_types, lifting) add.inverse_inverse image_iff pair_imageI)
qed
have 3: "⋀x X y. ⟦𝒟 tagged_partial_division_of cbox (- b) (- a); (x, X) ∈ 𝒟; y ∈ X⟧ ⟹ - y ∈ cbox a b"
by (metis (no_types, lifting) equation_minus_iff imageE subsetD tagged_partial_division_ofD(3) uminus_interval_vector)
have tag': "(λ(x,K). (- x, uminus ` K)) ` 𝒟 tagged_division_of cbox a b"
using tag by (auto simp: tagged_division_of_def dest: 1 2 3)
have fine': "γ fine (λ(x,K). (- x, uminus ` K)) ` 𝒟"
using fine by (fastforce simp: fine_def)
have inj: "inj_on (λ(x,K). (- x, uminus ` K)) 𝒟"
unfolding inj_on_def by force
have eq: "content (uminus ` I) = content I"
if I: "(x, I) ∈ 𝒟" and fnz: "f (- x) ≠ 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 "(∑(x,K) ∈ (λ(x,K). (- x, uminus ` K)) ` 𝒟. content K *⇩R f x) =
(∑(x,K) ∈ 𝒟. content K *⇩R f (- x))"
by (auto simp add: eq sum.reindex [OF inj] intro!: sum.cong)
then show ?thesis
using γ [OF ‹f ∈ F› tag' fine'] integral_reflect
by (metis (mono_tags, lifting) Henstock_Kurzweil_Integration.integral_cong comp_apply split_def sum.cong)
qed
qed
show ?thesis
using assms
apply (auto simp: equiintegrable_on_def)
subgoal for f
by (metis (mono_tags, lifting) comp_apply integrable_eq integrable_reflect)
using § by fastforce
qed
subsection‹Subinterval restrictions for equiintegrable families›
text‹First, some technical lemmas about minimizing a "flat" part of a sum over a division.›
lemma lemma0:
assumes "i ∈ Basis"
shows "content (cbox u v) / (interval_upperbound (cbox u v) ∙ i - interval_lowerbound (cbox u v) ∙ i) =
(if content (cbox u v) = 0 then 0
else ∏j ∈ Basis - {i}. interval_upperbound (cbox u v) ∙ j - interval_lowerbound (cbox u v) ∙ 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: "𝒟 division_of S" and S: "S ⊆ cbox a b" and i: "i ∈ Basis"
and mt: "⋀K. K ∈ 𝒟 ⟹ content K ≠ 0"
and disj: "(∀K ∈ 𝒟. K ∩ {x. x ∙ i = a ∙ i} ≠ {}) ∨ (∀K ∈ 𝒟. K ∩ {x. x ∙ i = b ∙ i} ≠ {})"
shows "(b ∙ i - a ∙ i) * (∑K∈𝒟. content K / (interval_upperbound K ∙ i - interval_lowerbound K ∙ i))
≤ content(cbox a b)" (is "?lhs ≤ ?rhs")
proof -
have "finite 𝒟"
using div by blast
define extend where
"extend ≡ λK. cbox (∑j ∈ Basis. if j = i then (a ∙ i) *⇩R i else (interval_lowerbound K ∙ j) *⇩R j)
(∑j ∈ Basis. if j = i then (b ∙ i) *⇩R i else (interval_upperbound K ∙ j) *⇩R j)"
have div_subset_cbox: "⋀K. K ∈ 𝒟 ⟹ K ⊆ cbox a b"
using S div by auto
have "⋀K. K ∈ 𝒟 ⟹ K ≠ {}"
using div by blast
have extend_cbox: "⋀K. K ∈ 𝒟 ⟹ ∃a b. extend K = cbox a b"
using extend_def by blast
have extend: "extend K ≠ {}" "extend K ⊆ cbox a b" if K: "K ∈ 𝒟" for K
proof -
obtain u v where K: "K = cbox u v" "K ≠ {}" "K ⊆ cbox a b"
using K cbox_division_memE [OF _ div] by (meson div_subset_cbox)
with i show "extend K ⊆ cbox a b"
by (auto simp: extend_def subset_box box_ne_empty)
have "a ∙ i ≤ b ∙ i"
using K by (metis bot.extremum_uniqueI box_ne_empty(1) i)
with K show "extend K ≠ {}"
by (simp add: extend_def i box_ne_empty)
qed
have int_extend_disjoint:
"interior(extend K1) ∩ interior(extend K2) = {}" if K: "K1 ∈ 𝒟" "K2 ∈ 𝒟" "K1 ≠ K2" for K1 K2
proof -
obtain u v where K1: "K1 = cbox u v" "K1 ≠ {}" "K1 ⊆ 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 ≠ {}" "K2 ⊆ cbox a b"
using K cbox_division_memE [OF _ div] by (meson div_subset_cbox)
have cboxes: "cbox u v ∈ 𝒟" "cbox w z ∈ 𝒟" "cbox u v ≠ cbox w z"
using K1 K2 that by auto
with div have "interior (cbox u v) ∩ interior (cbox w z) = {}"
by blast
moreover
have "∃x. x ∈ box u v ∧ x ∈ box w z"
if "x ∈ interior (extend K1)" "x ∈ interior (extend K2)" for x
proof -
have "a ∙ i < x ∙ i" "x ∙ i < b ∙ i"
and ux: "⋀k. k ∈ Basis - {i} ⟹ u ∙ k < x ∙ k"
and xv: "⋀k. k ∈ Basis - {i} ⟹ x ∙ k < v ∙ k"
and wx: "⋀k. k ∈ Basis - {i} ⟹ w ∙ k < x ∙ k"
and xz: "⋀k. k ∈ Basis - {i} ⟹ x ∙ k < z ∙ k"
using that K1 K2 i by (auto simp: extend_def box_ne_empty mem_box)
have "box u v ≠ {}" "box w z ≠ {}"
using cboxes interior_cbox by (auto simp: content_eq_0_interior dest: mt)
then obtain q s
where q: "⋀k. k ∈ Basis ⟹ w ∙ k < q ∙ k ∧ q ∙ k < z ∙ k"
and s: "⋀k. k ∈ Basis ⟹ u ∙ k < s ∙ k ∧ s ∙ k < v ∙ k"
by (meson all_not_in_conv mem_box(1))
show ?thesis using disj
proof
assume "∀K∈𝒟. K ∩ {x. x ∙ i = a ∙ i} ≠ {}"
then have uva: "(cbox u v) ∩ {x. x ∙ i = a ∙ i} ≠ {}"
and wza: "(cbox w z) ∩ {x. x ∙ i = a ∙ i} ≠ {}"
using cboxes by (auto simp: content_eq_0_interior)
then obtain r t where "r ∙ i = a ∙ i" and r: "⋀k. k ∈ Basis ⟹ w ∙ k ≤ r ∙ k ∧ r ∙ k ≤ z ∙ k"
and "t ∙ i = a ∙ i" and t: "⋀k. k ∈ Basis ⟹ u ∙ k ≤ t ∙ k ∧ t ∙ k ≤ v ∙ k"
by (fastforce simp: mem_box)
have u: "u ∙ i < q ∙ i"
using i K2(1) K2(3) ‹t ∙ i = a ∙ i› q s t [OF i] by (force simp: subset_box)
have w: "w ∙ i < s ∙ i"
using i K1(1) K1(3) ‹r ∙ i = a ∙ i› s r [OF i] by (force simp: subset_box)
define ξ where "ξ ≡ (∑j ∈ Basis. if j = i then min (q ∙ i) (s ∙ i) *⇩R i else (x ∙ j) *⇩R j)"
have [simp]: "ξ ∙ j = (if j = i then min (q ∙ j) (s ∙ j) else x ∙ j)" if "j ∈ Basis" for j
unfolding ξ_def
by (intro sum_if_inner that ‹i ∈ Basis›)
show ?thesis
proof (intro exI conjI)
have "min (q ∙ i) (s ∙ i) < v ∙ i"
using i s by fastforce
with ‹i ∈ Basis› s u ux xv
show "ξ ∈ box u v"
by (force simp: mem_box)
have "min (q ∙ i) (s ∙ i) < z ∙ i"
using i q by force
with ‹i ∈ Basis› q w wx xz
show "ξ ∈ box w z"
by (force simp: mem_box)
qed
next
assume "∀K∈𝒟. K ∩ {x. x ∙ i = b ∙ i} ≠ {}"
then have uva: "(cbox u v) ∩ {x. x ∙ i = b ∙ i} ≠ {}"
and wza: "(cbox w z) ∩ {x. x ∙ i = b ∙ i} ≠ {}"
using cboxes by (auto simp: content_eq_0_interior)
then obtain r t where "r ∙ i = b ∙ i" and r: "⋀k. k ∈ Basis ⟹ w ∙ k ≤ r ∙ k ∧ r ∙ k ≤ z ∙ k"
and "t ∙ i = b ∙ i" and t: "⋀k. k ∈ Basis ⟹ u ∙ k ≤ t ∙ k ∧ t ∙ k ≤ v ∙ k"
by (fastforce simp: mem_box)
have z: "s ∙ i < z ∙ i"
using K1(1) K1(3) ‹r ∙ i = b ∙ i› r [OF i] i s by (force simp: subset_box)
have v: "q ∙ i < v ∙ i"
using K2(1) K2(3) ‹t ∙ i = b ∙ i› t [OF i] i q by (force simp: subset_box)
define ξ where "ξ ≡ (∑j ∈ Basis. if j = i then max (q ∙ i) (s ∙ i) *⇩R i else (x ∙ j) *⇩R j)"
have [simp]: "ξ ∙ j = (if j = i then max (q ∙ j) (s ∙ j) else x ∙ j)" if "j ∈ Basis" for j
unfolding ξ_def
by (intro sum_if_inner that ‹i ∈ Basis›)
show ?thesis
proof (intro exI conjI)
show "ξ ∈ box u v"
using ‹i ∈ Basis› s by (force simp: mem_box ux v xv)
show "ξ ∈ box w z"
using ‹i ∈ Basis› q by (force simp: mem_box wx xz z)
qed
qed
qed
ultimately show ?thesis by auto
qed
define interv_diff where "interv_diff ≡ λK. λi::'a. interval_upperbound K ∙ i - interval_lowerbound K ∙ i"
have "?lhs = (∑K∈𝒟. (b ∙ i - a ∙ i) * content K / (interv_diff K i))"
by (simp add: sum_distrib_left interv_diff_def)
also have "… = sum (content ∘ extend) 𝒟"
proof (rule sum.cong [OF refl])
fix K assume "K ∈ 𝒟"
then obtain u v where K: "K = cbox u v" "cbox u v ≠ {}" "K ⊆ cbox a b"
using cbox_division_memE [OF _ div] div_subset_cbox by metis
then have uv: "u ∙ i < v ∙ i"
using mt [OF ‹K ∈ 𝒟›] ‹i ∈ Basis› content_eq_0 by fastforce
have "insert i (Basis ∩ -{i}) = Basis"
using ‹i ∈ Basis› by auto
then have "(b ∙ i - a ∙ i) * content K / (interv_diff K i)
= (b ∙ i - a ∙ i) * (∏i ∈ insert i (Basis ∩ -{i}). v ∙ i - u ∙ i) / (interv_diff (cbox u v) i)"
using K box_ne_empty(1) content_cbox by fastforce
also have "... = (∏x∈Basis. if x = i then b ∙ x - a ∙ x
else (interval_upperbound (cbox u v) - interval_lowerbound (cbox u v)) ∙ x)"
using ‹i ∈ Basis› K uv by (simp add: prod.If_cases interv_diff_def) (simp add: algebra_simps)
also have "... = (∏k∈Basis.
(∑j∈Basis. if j = i then (b ∙ i - a ∙ i) *⇩R i
else ((interval_upperbound (cbox u v) - interval_lowerbound (cbox u v)) ∙ j) *⇩R j) ∙ k)"
using ‹i ∈ Basis› by (subst prod.cong [OF refl sum_if_inner]; simp)
also have "... = (∏k∈Basis.
(∑j∈Basis. if j = i then (b ∙ i) *⇩R i else (interval_upperbound (cbox u v) ∙ j) *⇩R j) ∙ k -
(∑j∈Basis. if j = i then (a ∙ i) *⇩R i else (interval_lowerbound (cbox u v) ∙ j) *⇩R j) ∙ k)"
using ‹i ∈ Basis›
by (intro prod.cong [OF refl]) (subst sum_if_inner; simp add: algebra_simps)+
also have "... = (content ∘ extend) K"
using ‹i ∈ Basis› K box_ne_empty ‹K ∈ 𝒟› extend(1)
by (auto simp add: extend_def content_cbox_if)
finally show "(b ∙ i - a ∙ i) * content K / (interv_diff K i) = (content ∘ extend) K" .
qed
also have "... = sum content (extend ` 𝒟)"
proof -
have "⟦K1 ∈ 𝒟; K2 ∈ 𝒟; K1 ≠ K2; extend K1 = extend K2⟧ ⟹ 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 ‹finite 𝒟›])
qed
also have "... ≤ ?rhs"
proof (rule subadditive_content_division)
show "extend ` 𝒟 division_of ⋃ (extend ` 𝒟)"
using int_extend_disjoint by (auto simp: division_of_def ‹finite 𝒟› extend extend_cbox)
show "⋃ (extend ` 𝒟) ⊆ cbox a b"
using extend by fastforce
qed
finally show ?thesis .
qed
proposition sum_content_area_over_thin_division:
assumes div: "𝒟 division_of S" and S: "S ⊆ cbox a b" and i: "i ∈ Basis"
and "a ∙ i ≤ c" "c ≤ b ∙ i"
and nonmt: "⋀K. K ∈ 𝒟 ⟹ K ∩ {x. x ∙ i = c} ≠ {}"
shows "(b ∙ i - a ∙ i) * (∑K∈𝒟. content K / (interval_upperbound K ∙ i - interval_lowerbound K ∙ i))
≤ 2 * content(cbox a b)"
proof (cases "content(cbox a b) = 0")
case True
have "(∑K∈𝒟. content K / (interval_upperbound K ∙ i - interval_lowerbound K ∙ 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 ∙ i < b ∙ i" if "i ∈ Basis" for i
using content_pos_lt_eq that by blast
have "finite 𝒟"
using div by blast
define Dlec where "Dlec ≡ {L ∈ (λL. L ∩ {x. x ∙ i ≤ c}) ` 𝒟. content L ≠ 0}"
define Dgec where "Dgec ≡ {L ∈ (λL. L ∩ {x. x ∙ i ≥ c}) ` 𝒟. content L ≠ 0}"
define a' where "a' ≡ (∑j∈Basis. (if j = i then c else a ∙ j) *⇩R j)"
define b' where "b' ≡ (∑j∈Basis. (if j = i then c else b ∙ j) *⇩R j)"
define interv_diff where "interv_diff ≡ λK. λi::'a. interval_upperbound K ∙ i - interval_lowerbound K ∙ i"
have Dlec_cbox: "⋀K. K ∈ Dlec ⟹ ∃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: "⟦content (L ∩ {x. x ∙ i ≤ c}) ≠ 0; L ∈ 𝒟⟧ ⟹ ∃a b. L ∩ {x. x ∙ i ≤ c} = cbox a b" for L
using Dlec_def by blast
have Dgec_cbox: "⋀K. K ∈ Dgec ⟹ ∃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: "⟦content (L ∩ {x. x ∙ i ≥ c}) ≠ 0; L ∈ 𝒟⟧ ⟹ ∃a b. L ∩ {x. x ∙ i ≥ c} = cbox a b" for L
using Dgec_def by blast
have zero_left: "⋀x y. ⟦x ∈ 𝒟; y ∈ 𝒟; x ≠ y; x ∩ {x. x ∙ i ≤ c} = y ∩ {x. x ∙ i ≤ c}⟧
⟹ content (y ∩ {x. x ∙ i ≤ c}) = 0"
by (metis division_split_left_inj [OF div] lec_is_cbox content_eq_0_interior)
have zero_right: "⋀x y. ⟦x ∈ 𝒟; y ∈ 𝒟; x ≠ y; x ∩ {x. c ≤ x ∙ i} = y ∩ {x. c ≤ x ∙ i}⟧
⟹ content (y ∩ {x. c ≤ x ∙ i}) = 0"
by (metis division_split_right_inj [OF div] gec_is_cbox content_eq_0_interior)
have "(b' ∙ i - a ∙ i) * (∑K∈Dlec. content K / interv_diff K i) ≤ content(cbox a b')"
unfolding interv_diff_def
proof (rule content_division_lemma1)
show "Dlec division_of ⋃Dlec"
unfolding division_of_def
proof (intro conjI ballI Dlec_cbox)
show "⋀K1 K2. ⟦K1 ∈ Dlec; K2 ∈ Dlec⟧ ⟹ K1 ≠ K2 ⟶ interior K1 ∩ interior K2 = {}"
by (clarsimp simp: Dlec_def) (use div in auto)
qed (use ‹finite 𝒟› Dlec_def in auto)
show "⋃Dlec ⊆ cbox a b'"
using Dlec_def div S by (auto simp: b'_def division_of_def mem_box)
show "(∀K∈Dlec. K ∩ {x. x ∙ i = a ∙ i} ≠ {}) ∨ (∀K∈Dlec. K ∩ {x. x ∙ i = b' ∙ i} ≠ {})"
using nonmt by (fastforce simp: Dlec_def b'_def i)
qed (use i Dlec_def in auto)
moreover
have "(∑K∈Dlec. content K / (interv_diff K i)) = (∑K∈(λK. K ∩ {x. x ∙ i ≤ c}) ` 𝒟. content K / interv_diff K i)"
unfolding Dlec_def using ‹finite 𝒟› by (auto simp: sum.mono_neutral_left)
moreover have "... =
(∑K∈𝒟. ((λK. content K / (interv_diff K i)) ∘ ((λK. K ∩ {x. x ∙ i ≤ c}))) K)"
by (simp add: zero_left sum.reindex_nontrivial [OF ‹finite 𝒟›])
moreover have "(b' ∙ i - a ∙ i) = (c - a ∙ i)"
by (simp add: b'_def i)
ultimately
have lec: "(c - a ∙ i) * (∑K∈𝒟. ((λK. content K / (interv_diff K i)) ∘ ((λK. K ∩ {x. x ∙ i ≤ c}))) K)
≤ content(cbox a b')"
by simp
have "(b ∙ i - a' ∙ i) * (∑K∈Dgec. content K / (interv_diff K i)) ≤ content(cbox a' b)"
unfolding interv_diff_def
proof (rule content_division_lemma1)
show "Dgec division_of ⋃Dgec"
unfolding division_of_def
proof (intro conjI ballI Dgec_cbox)
show "⋀K1 K2. ⟦K1 ∈ Dgec; K2 ∈ Dgec⟧ ⟹ K1 ≠ K2 ⟶ interior K1 ∩ interior K2 = {}"
by (clarsimp simp: Dgec_def) (use div in auto)
qed (use ‹finite 𝒟› Dgec_def in auto)
show "⋃Dgec ⊆ cbox a' b"
using Dgec_def div S by (auto simp: a'_def division_of_def mem_box)
show "(∀K∈Dgec. K ∩ {x. x ∙ i = a' ∙ i} ≠ {}) ∨ (∀K∈Dgec. K ∩ {x. x ∙ i = b ∙ i} ≠ {})"
using nonmt by (fastforce simp: Dgec_def a'_def i)
qed (use i Dgec_def in auto)
moreover
have "(∑K∈Dgec. content K / (interv_diff K i)) = (∑K∈(λK. K ∩ {x. c ≤ x ∙ i}) ` 𝒟.
content K / interv_diff K i)"
unfolding Dgec_def using ‹finite 𝒟› by (auto simp: sum.mono_neutral_left)
moreover have "… =
(∑K∈𝒟. ((λK. content K / (interv_diff K i)) ∘ ((λK. K ∩ {x. x ∙ i ≥ c}))) K)"
by (simp add: zero_right sum.reindex_nontrivial [OF ‹finite 𝒟›])
moreover have "(b ∙ i - a' ∙ i) = (b ∙ i - c)"
by (simp add: a'_def i)
ultimately
have gec: "(b ∙ i - c) * (∑K∈𝒟. ((λK. content K / (interv_diff K i)) ∘ ((λK. K ∩ {x. x ∙ i ≥ c}))) K)
≤ content(cbox a' b)"
by simp
show ?thesis
proof (cases "c = a ∙ i ∨ c = b ∙ i")
case True
then show ?thesis
proof
assume c: "c = a ∙ i"
moreover
have "(∑j∈Basis. (if j = i then a ∙ i else a ∙ j) *⇩R j) = a"
using euclidean_representation [of a] sum.cong [OF refl, of Basis "λi. (a ∙ i) *⇩R i"] by presburger
ultimately have "a' = a"
by (simp add: i a'_def cong: if_cong)
then have "content (cbox a' b) ≤ 2 * content (cbox a b)" by simp
moreover
have eq: "(∑K∈𝒟. content (K ∩ {x. a ∙ i ≤ x ∙ i}) / interv_diff (K ∩ {x. a ∙ i ≤ x ∙ i}) i)
= (∑K∈𝒟. content K / interv_diff K i)"
(is "sum ?f _ = sum ?g _")
proof (rule sum.cong [OF refl])
fix K assume "K ∈ 𝒟"
then have "a ∙ i ≤ x ∙ i" if "x ∈ K" for x
by (metis S UnionI div division_ofD(6) i mem_box(2) subsetCE that)
then have "K ∩ {x. a ∙ i ≤ x ∙ i} = K"
by blast
then show "?f K = ?g K"
by simp
qed
ultimately show ?thesis
using gec c eq interv_diff_def by auto
next
assume c: "c = b ∙ i"
moreover have "(∑j∈Basis. (if j = i then b ∙ i else b ∙ j) *⇩R j) = b"
using euclidean_representation [of b] sum.cong [OF refl, of Basis "λi. (b ∙ i) *⇩R i"] by presburger
ultimately have "b' = b"
by (simp add: i b'_def cong: if_cong)
then have "content (cbox a b') ≤ 2 * content (cbox a b)" by simp
moreover
have eq: "(∑K∈𝒟. content (K ∩ {x. x ∙ i ≤ b ∙ i}) / interv_diff (K ∩ {x. x ∙ i ≤ b ∙ i}) i)
= (∑K∈𝒟. content K / interv_diff K i)"
(is "sum ?f _ = sum ?g _")
proof (rule sum.cong [OF refl])
fix K assume "K ∈ 𝒟"
then have "x ∙ i ≤ b ∙ i" if "x ∈ K" for x
by (metis S UnionI div division_ofD(6) i mem_box(2) subsetCE that)
then have "K ∩ {x. x ∙ i ≤ b ∙ i} = K"
by blast
then show "?f K = ?g K"
by simp
qed
ultimately show ?thesis
using lec c eq interv_diff_def by auto
qed
next
case False
have prod_if: "(∏k∈Basis ∩ - {i}. f k) = (∏k∈Basis. f k) / f i" if "f i ≠ (0::real)" for f
proof -
have "f i * prod f (Basis ∩ - {i}) = prod f Basis"
using that mk_disjoint_insert [OF i]
by (metis Int_insert_left_if0 finite_Basis finite_insert le_iff_inf order_refl prod.insert subset_Compl_singleton)
then show ?thesis
by (metis nonzero_mult_div_cancel_left that)
qed
have abc: "a ∙ i < c" "c < b ∙ i"
using False assms by auto
then have "(∑K∈𝒟. ((λK. content K / (interv_diff K i)) ∘ ((λK. K ∩ {x. x ∙ i ≤ c}))) K)
≤ content(cbox a b') / (c - a ∙ i)"
"(∑K∈𝒟. ((λK. content K / (interv_diff K i)) ∘ ((λK. K ∩ {x. x ∙ i ≥ c}))) K)
≤ content(cbox a' b) / (b ∙ i - c)"
using lec gec by (simp_all add: field_split_simps)
moreover
have "(∑K∈𝒟. content K / (interv_diff K i))
≤ (∑K∈𝒟. ((λK. content K / (interv_diff K i)) ∘ ((λK. K ∩ {x. x ∙ i ≤ c}))) K) +
(∑K∈𝒟. ((λK. content K / (interv_diff K i)) ∘ ((λK. K ∩ {x. x ∙ i ≥ c}))) K)"
(is "?lhs ≤ ?rhs")
proof -
have "?lhs ≤
(∑K∈𝒟. ((λK. content K / (interv_diff K i)) ∘ ((λK. K ∩ {x. x ∙ i ≤ c}))) K +
((λK. content K / (interv_diff K i)) ∘ ((λK. K ∩ {x. x ∙ i ≥ c}))) K)"
(is "sum ?f _ ≤ sum ?g _")
proof (rule sum_mono)
fix K assume "K ∈ 𝒟"
then obtain u v where uv: "K = cbox u v"
using div by blast
obtain u' v' where uv': "cbox u v ∩ {x. x ∙ i ≤ c} = cbox u v'"
"cbox u v ∩ {x. c ≤ x ∙ i} = cbox u' v"
"⋀k. k ∈ Basis ⟹ u' ∙ k = (if k = i then max (u ∙ i) c else u ∙ k)"
"⋀k. k ∈ Basis ⟹ v' ∙ k = (if k = i then min (v ∙ i) c else v ∙ k)"
using i by (auto simp: interval_split)
have *: "⟦content (cbox u v') = 0; content (cbox u' v) = 0⟧ ⟹ content (cbox u v) = 0"
"content (cbox u' v) ≠ 0 ⟹ content (cbox u v) ≠ 0"
"content (cbox u v') ≠ 0 ⟹ content (cbox u v) ≠ 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)
have uniq: "⋀j. ⟦j ∈ Basis; ¬ u ∙ j ≤ v ∙ j⟧ ⟹ j = i"
by (metis ‹K ∈ 𝒟› box_ne_empty(1) div division_of_def uv)
show "?f K ≤ ?g K"
using i uv uv' by (auto simp add: interv_diff_def lemma0 dest: uniq * intro!: prod_nonneg)
qed
also have "... = ?rhs"
by (simp add: sum.distrib)
finally show ?thesis .
qed
moreover have "content (cbox a b') / (c - a ∙ i) = content (cbox a b) / (b ∙ i - a ∙ 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 "λf. f x" for x] prod.If_cases [of Basis "λx. x = i", simplified] prod_if field_simps)
done
moreover have "content (cbox a' b) / (b ∙ i - c) = content (cbox a b) / (b ∙ i - a ∙ 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 "λx. x = i", simplified] prod_if field_simps)
done
ultimately
have "(∑K∈𝒟. content K / (interv_diff K i)) ≤ 2 * content (cbox a b) / (b ∙ i - a ∙ i)"
by linarith
then show ?thesis
using abc interv_diff_def by (simp add: field_split_simps)
qed
qed
proposition bounded_equiintegral_over_thin_tagged_partial_division:
fixes f :: "'a::euclidean_space ⇒ 'b::euclidean_space"
assumes F: "F equiintegrable_on cbox a b" and f: "f ∈ F" and "0 < ε"
and norm_f: "⋀h x. ⟦h ∈ F; x ∈ cbox a b⟧ ⟹ norm(h x) ≤ norm(f x)"
obtains γ where "gauge γ"
"⋀c i S h. ⟦c ∈ cbox a b; i ∈ Basis; S tagged_partial_division_of cbox a b;
γ fine S; h ∈ F; ⋀x K. (x,K) ∈ S ⟹ (K ∩ {x. x ∙ i = c ∙ i} ≠ {})⟧
⟹ (∑(x,K) ∈ S. norm (integral K h)) < ε"
proof (cases "content(cbox a b) = 0")
case True
show ?thesis
proof
show "gauge (λx. ball x 1)"
by (simp add: gauge_trivial)
show "(∑(x,K) ∈ S. norm (integral K h)) < ε"
if "S tagged_partial_division_of cbox a b" "(λx. ball x 1) fine S" for S and h:: "'a ⇒ 'b"
proof -
have "(∑(x,K) ∈ S. norm (integral K h)) = 0"
using that True content_0_subset
by (fastforce simp: tagged_partial_division_of_def intro: sum.neutral)
with ‹0 < ε› 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: "⋀i. i ∈ Basis ⟹ a∙i < b∙i"
by (auto simp: content_pos_lt_eq)
obtain γ0 where "gauge γ0"
and γ0: "⋀S h. ⟦S tagged_partial_division_of cbox a b; γ0 fine S; h ∈ F⟧
⟹ (∑(x,K) ∈ S. norm (content K *⇩R h x - integral K h)) < ε/2"
proof -
obtain γ where "gauge γ"
and γ: "⋀f 𝒟. ⟦f ∈ F; 𝒟 tagged_division_of cbox a b; γ fine 𝒟⟧
⟹ norm ((∑(x,K) ∈ 𝒟. content K *⇩R f x) - integral (cbox a b) f)
< ε/(5 * (Suc DIM('b)))"
proof -
have e5: "ε/(5 * (Suc DIM('b))) > 0"
using ‹ε > 0› by auto
then show ?thesis
using F that by (auto simp: equiintegrable_on_def)
qed
show ?thesis
proof
show "gauge γ"
by (rule ‹gauge γ›)
show "(∑(x,K) ∈ S. norm (content K *⇩R h x - integral K h)) < ε/2"
if "S tagged_partial_division_of cbox a b" "γ fine S" "h ∈ F" for S h
proof -
have "(∑(x,K) ∈ S. norm (content K *⇩R h x - integral K h)) ≤ 2 * real DIM('b) * (ε/(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 γ"
by (rule ‹gauge γ›)
qed (use that ‹ε > 0› γ in auto)
also have "... < ε/2"
using ‹ε > 0› by (simp add: divide_simps)
finally show ?thesis .
qed
qed
qed
define γ where "γ ≡ λx. γ0 x ∩
ball x ((ε/8 / (norm(f x) + 1)) * (INF m∈Basis. b ∙ m - a ∙ m) / content(cbox a b))"
define interv_diff where "interv_diff ≡ λK. λi::'a. interval_upperbound K ∙ i - interval_lowerbound K ∙ i"
have "8 * content (cbox a b) + norm (f x) * (8 * content (cbox a b)) > 0" for x
by (metis add.right_neutral add_pos_pos contab_gt0 mult_pos_pos mult_zero_left norm_eq_zero zero_less_norm_iff zero_less_numeral)
then have "gauge (λx. ball x
(ε * (INF m∈Basis. b ∙ m - a ∙ m) / ((8 * norm (f x) + 8) * content (cbox a b))))"
using ‹0 < content (cbox a b)› ‹0 < ε› a_less_b
by (auto simp add: gauge_def field_split_simps add_nonneg_eq_0_iff finite_less_Inf_iff)
then have "gauge γ"
unfolding γ_def using ‹gauge γ0› gauge_Int by auto
moreover
have "(∑(x,K) ∈ S. norm (integral K h)) < ε"
if "c ∈ cbox a b" "i ∈ Basis" and S: "S tagged_partial_division_of cbox a b"
and "γ fine S" "h ∈ F" and ne: "⋀x K. (x,K) ∈ S ⟹ K ∩ {x. x ∙ i = c ∙ i} ≠ {}" for c i S h
proof -
have "cbox c b ⊆ cbox a b"
by (meson mem_box(2) order_refl subset_box(1) that(1))
have "finite S"
using S unfolding tagged_partial_division_of_def by blast
have "γ0 fine S" and fineS:
"(λx. ball x (ε * (INF m∈Basis. b ∙ m - a ∙ m) / ((8 * norm (f x) + 8) * content (cbox a b)))) fine S"
using ‹γ fine S› by (auto simp: γ_def fine_Int)
then have "(∑(x,K) ∈ S. norm (content K *⇩R h x - integral K h)) < ε/2"
by (intro γ0 that fineS)
moreover have "(∑(x,K) ∈ S. norm (integral K h) - norm (content K *⇩R h x - integral K h)) ≤ ε/2"
proof -
have "(∑(x,K) ∈ S. norm (integral K h) - norm (content K *⇩R h x - integral K h))
≤ (∑(x,K) ∈ S. norm (content K *⇩R h x))"
proof (clarify intro!: sum_mono)
fix x K
assume xK: "(x,K) ∈ S"
have "norm (integral K h) - norm (content K *⇩R h x - integral K h) ≤ norm (integral K h - (integral K h - content K *⇩R h x))"
by (metis norm_minus_commute norm_triangle_ineq2)
also have "... ≤ norm (content K *⇩R h x)"
by simp
finally show "norm (integral K h) - norm (content K *⇩R h x - integral K h) ≤ norm (content K *⇩R h x)" .
qed
also have "... ≤ (∑(x,K) ∈ S. ε/4 * (b ∙ i - a ∙ i) / content (cbox a b) * content K / interv_diff K i)"
proof (clarify intro!: sum_mono)
fix x K
assume xK: "(x,K) ∈ S"
then have x: "x ∈ cbox a b"
using S unfolding tagged_partial_division_of_def by (meson subset_iff)
show "norm (content K *⇩R h x) ≤ ε/4 * (b ∙ i - a ∙ i) / content (cbox a b) * content K / interv_diff K i"
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 ‹(x,K) ∈ S› unfolding tagged_partial_division_of_def by blast
then have u_less_v: "⋀i. i ∈ Basis ⟹ u ∙ i < v ∙ 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) ≤ (ε * (b ∙ i - a ∙ i)) / (4 * content (cbox a b) * interv_diff K i)"
proof -
have "dist x u < ε * (INF m∈Basis. b ∙ m - a ∙ m) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2"
"dist x v < ε * (INF m∈Basis. b ∙ m - a ∙ 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 "ε * (INF m∈Basis. b ∙ m - a ∙ m) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2
≤ ε * (b ∙ i - a ∙ i) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2"
proof (intro mult_left_mono divide_right_mono)
show "(INF m∈Basis. b ∙ m - a ∙ m) ≤ b ∙ i - a ∙ i"
using ‹i ∈ Basis› by (auto intro!: cInf_le_finite)
qed (use ‹0 < ε› in auto)
ultimately
have "dist x u < ε * (b ∙ i - a ∙ i) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2"
"dist x v < ε * (b ∙ i - a ∙ i) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2"
by linarith+
then have duv: "dist u v < ε * (b ∙ i - a ∙ i) / (4 * (norm (f x) + 1) * content (cbox a b))"
using dist_triangle_half_r by blast
have uvi: "¦v ∙ i - u ∙ i¦ ≤ norm (v - u)"
by (metis inner_commute inner_diff_right ‹i ∈ Basis› Basis_le_norm)
have "norm (h x) ≤ norm (f x)"
using x that by (auto simp: norm_f)
also have "... < (norm (f x) + 1)"
by simp
also have "... < ε * (b ∙ i - a ∙ i) / dist u v / (4 * content (cbox a b))"
proof -
have "0 < norm (f x) + 1"
by (simp add: add.commute add_pos_nonneg)
then show ?thesis
using duv dist_uv contab_gt0
by (simp only: mult_ac divide_simps) auto
qed
also have "... = ε * (b ∙ i - a ∙ i) / norm (v - u) / (4 * content (cbox a b))"
by (simp add: dist_norm norm_minus_commute)
also have "... ≤ ε * (b ∙ i - a ∙ i) / ¦v ∙ i - u ∙ i¦ / (4 * content (cbox a b))"
proof (intro mult_right_mono divide_left_mono divide_right_mono uvi)
show "norm (v - u) * ¦v ∙ i - u ∙ i¦ > 0"
using u_less_v [OF ‹i ∈ Basis›]
by (auto simp: less_eq_real_def zero_less_mult_iff that)
show "ε * (b ∙ i - a ∙ i) ≥ 0"
using a_less_b ‹0 < ε› ‹i ∈ Basis› by force
qed auto
also have "... = ε * (b ∙ i - a ∙ i) / (4 * content (cbox a b) * interv_diff K i)"
using uv False that(2) u_less_v interv_diff_def by fastforce
finally show ?thesis by simp
qed
with Kgt0 have "norm (content K *⇩R h x) ≤ content K * ((ε/4 * (b ∙ i - a ∙ i) / content (cbox a b)) / interv_diff K i)"
using mult_left_mono by fastforce
also have "... = ε/4 * (b ∙ i - a ∙ i) / content (cbox a b) * content K / interv_diff K i"
by (simp add: field_split_simps)
finally show ?thesis .
qed
qed
also have "... = (∑K∈snd ` S. ε/4 * (b ∙ i - a ∙ i) / content (cbox a b) * content K / interv_diff K i)"
unfolding interv_diff_def
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 "... = ε/2 * ((b ∙ i - a ∙ i) / (2 * content (cbox a b)) * (∑K∈snd ` S. content K / interv_diff K i))"
by (simp add: interv_diff_def sum_distrib_left mult.assoc)
also have "... ≤ (ε/2) * 1"
proof (rule mult_left_mono)
have "(b ∙ i - a ∙ i) * (∑K∈snd ` S. content K / interv_diff K i) ≤ 2 * content (cbox a b)"
unfolding interv_diff_def
proof (rule sum_content_area_over_thin_division)
show "snd ` S division_of ⋃(snd ` S)"
by (auto intro: S tagged_partial_division_of_Union_self division_of_tagged_division)
show "⋃(snd ` S) ⊆ cbox a b"
using S unfolding tagged_partial_division_of_def by force
show "a ∙ i ≤ c ∙ i" "c ∙ i ≤ b ∙ i"
using mem_box(2) that by blast+
qed (use