Theory Borel

Up to index of Isabelle/HOL/HOL-Probability

theory Borel
imports Measure

header {*Borel Sets*}

theory Borel
imports Measure
begin


text{*From the Hurd/Coble measure theory development, translated by Lawrence Paulson.*}

definition borel_space where
"borel_space = sigma (UNIV::real set) (range (λa::real. {x. x ≤ a}))"


definition borel_measurable where
"borel_measurable a = measurable a borel_space"


definition indicator_fn where
"indicator_fn s = (λx. if x ∈ s then 1 else (0::real))"


lemma in_borel_measurable:
"f ∈ borel_measurable M <->
sigma_algebra M ∧
(∀s ∈ sets (sigma UNIV (range (λa::real. {x. x ≤ a}))).
f -` s ∩ space M ∈ sets M)"

apply (auto simp add: borel_measurable_def measurable_def borel_space_def)
apply (metis PowD UNIV_I Un_commute sigma_algebra_sigma subset_Pow_Union subset_UNIV subset_Un_eq)
done

lemma (in sigma_algebra) borel_measurable_const:
"(λx. c) ∈ borel_measurable M"

by (auto simp add: in_borel_measurable prems)

lemma (in sigma_algebra) borel_measurable_indicator:
assumes a: "a ∈ sets M"
shows "indicator_fn a ∈ borel_measurable M"

apply (auto simp add: in_borel_measurable indicator_fn_def prems)
apply (metis Diff_eq Int_commute a compl_sets)
done

lemma Collect_eq: "{w ∈ X. f w ≤ a} = {w. f w ≤ a} ∩ X"
by (metis Collect_conj_eq Collect_mem_eq Int_commute)

lemma (in measure_space) borel_measurable_le_iff:
"f ∈ borel_measurable M = (∀a. {w ∈ space M. f w ≤ a} ∈ sets M)"

proof
assume f: "f ∈ borel_measurable M"
{ fix a
have "{w ∈ space M. f w ≤ a} ∈ sets M" using f
apply (auto simp add: in_borel_measurable sigma_def Collect_eq)
apply (drule_tac x="{x. x ≤ a}" in bspec, auto)
apply (metis equalityE rangeI subsetD sigma_sets.Basic)
done
}
thus "∀a. {w ∈ space M. f w ≤ a} ∈ sets M" by blast
next
assume "∀a. {w ∈ space M. f w ≤ a} ∈ sets M"
thus "f ∈ borel_measurable M"
apply (simp add: borel_measurable_def borel_space_def Collect_eq)
apply (rule measurable_sigma, auto)
done
qed

lemma Collect_less_le:
"{w ∈ X. f w < g w} = (\<Union>n. {w ∈ X. f w ≤ g w - inverse(real(Suc n))})"

proof auto
fix w
assume w: "f w < g w"
hence nz: "g w - f w ≠ 0"
by arith
with w have "real(Suc(natceiling(inverse(g w - f w)))) > inverse(g w - f w)"
by (metis lessI order_le_less_trans real_natceiling_ge real_of_nat_less_iff) hence "inverse(real(Suc(natceiling(inverse(g w - f w)))))
< inverse(inverse(g w - f w))"

by (metis less_iff_diff_less_0 less_imp_inverse_less linorder_neqE_linordered_idom nz positive_imp_inverse_positive order_antisym less_le w)
hence "inverse(real(Suc(natceiling(inverse(g w - f w))))) < g w - f w"
by (metis inverse_inverse_eq order_less_le_trans order_refl)
thus "∃n. f w ≤ g w - inverse(real(Suc n))" using w
by (rule_tac x="natceiling(inverse(g w - f w))" in exI, auto)
next
fix w n
assume le: "f w ≤ g w - inverse(real(Suc n))"
hence "0 < inverse(real(Suc n))"
by simp
thus "f w < g w" using le
by arith
qed


lemma (in sigma_algebra) sigma_le_less:
assumes M: "!!a::real. {w ∈ space M. f w ≤ a} ∈ sets M"
shows "{w ∈ space M. f w < a} ∈ sets M"

proof -
show ?thesis using Collect_less_le [of "space M" f "λx. a"]
by (auto simp add: countable_UN M)
qed

lemma (in sigma_algebra) sigma_less_ge:
assumes M: "!!a::real. {w ∈ space M. f w < a} ∈ sets M"
shows "{w ∈ space M. a ≤ f w} ∈ sets M"

proof -
have "{w ∈ space M. a ≤ f w} = space M - {w ∈ space M. f w < a}"
by auto
thus ?thesis using M
by auto
qed

lemma (in sigma_algebra) sigma_ge_gr:
assumes M: "!!a::real. {w ∈ space M. a ≤ f w} ∈ sets M"
shows "{w ∈ space M. a < f w} ∈ sets M"

proof -
show ?thesis using Collect_less_le [of "space M" "λx. a" f]
by (auto simp add: countable_UN le_diff_eq M)
qed

lemma (in sigma_algebra) sigma_gr_le:
assumes M: "!!a::real. {w ∈ space M. a < f w} ∈ sets M"
shows "{w ∈ space M. f w ≤ a} ∈ sets M"

proof -
have "{w ∈ space M. f w ≤ a} = space M - {w ∈ space M. a < f w}"
by auto
thus ?thesis
by (simp add: M compl_sets)
qed

lemma (in measure_space) borel_measurable_gr_iff:
"f ∈ borel_measurable M = (∀a. {w ∈ space M. a < f w} ∈ sets M)"

proof (auto simp add: borel_measurable_le_iff sigma_gr_le)
fix u
assume M: "∀a. {w ∈ space M. f w ≤ a} ∈ sets M"
have "{w ∈ space M. u < f w} = space M - {w ∈ space M. f w ≤ u}"
by auto
thus "{w ∈ space M. u < f w} ∈ sets M"
by (force simp add: compl_sets countable_UN M)
qed

lemma (in measure_space) borel_measurable_less_iff:
"f ∈ borel_measurable M = (∀a. {w ∈ space M. f w < a} ∈ sets M)"

proof (auto simp add: borel_measurable_le_iff sigma_le_less)
fix u
assume M: "∀a. {w ∈ space M. f w < a} ∈ sets M"
have "{w ∈ space M. f w ≤ u} = space M - {w ∈ space M. u < f w}"
by auto
thus "{w ∈ space M. f w ≤ u} ∈ sets M"
using Collect_less_le [of "space M" "λx. u" f]
by (force simp add: compl_sets countable_UN le_diff_eq sigma_less_ge M)
qed

lemma (in measure_space) borel_measurable_ge_iff:
"f ∈ borel_measurable M = (∀a. {w ∈ space M. a ≤ f w} ∈ sets M)"

proof (auto simp add: borel_measurable_less_iff sigma_le_less sigma_ge_gr sigma_gr_le)
fix u
assume M: "∀a. {w ∈ space M. f w < a} ∈ sets M"
have "{w ∈ space M. u ≤ f w} = space M - {w ∈ space M. f w < u}"
by auto
thus "{w ∈ space M. u ≤ f w} ∈ sets M"
by (force simp add: compl_sets countable_UN M)
qed

lemma (in measure_space) affine_borel_measurable:
assumes g: "g ∈ borel_measurable M"
shows "(λx. a + (g x) * b) ∈ borel_measurable M"

proof (cases rule: linorder_cases [of b 0])
case equal thus ?thesis
by (simp add: borel_measurable_const)
next
case less
{
fix w c
have "a + g w * b ≤ c <-> g w * b ≤ c - a"
by auto
also have "... <-> (c-a)/b ≤ g w" using less
by (metis divide_le_eq less less_asym)
finally have "a + g w * b ≤ c <-> (c-a)/b ≤ g w" .
}
hence "!!w c. a + g w * b ≤ c <-> (c-a)/b ≤ g w" .
thus ?thesis using less g
by (simp add: borel_measurable_ge_iff [of g])
(simp add: borel_measurable_le_iff)

next
case greater
hence 0: "!!x c. (g x * b ≤ c - a) <-> (g x ≤ (c - a) / b)"
by (metis mult_imp_le_div_pos le_divide_eq)
have 1: "!!x c. (a + g x * b ≤ c) <-> (g x * b ≤ c - a)"
by auto
from greater g
show ?thesis
by (simp add: borel_measurable_le_iff 0 1)
qed

definition
nat_to_rat_surj :: "nat => rat" where
"nat_to_rat_surj n = (let (i,j) = prod_decode n
in Fract (int_decode i) (int_decode j))"


lemma nat_to_rat_surj: "surj nat_to_rat_surj"
proof (auto simp add: surj_def nat_to_rat_surj_def)
fix y
show "∃x. y = (λ(i, j). Fract (int_decode i) (int_decode j)) (prod_decode x)"
proof (cases y)
case (Fract a b)
obtain i where i: "int_decode i = a" using surj_int_decode
by (metis surj_def)
obtain j where j: "int_decode j = b" using surj_int_decode
by (metis surj_def)
obtain n where n: "prod_decode n = (i,j)" using surj_prod_decode
by (metis surj_def)

from Fract i j n show ?thesis
by (metis prod.cases prod_case_split)
qed
qed

lemma rats_enumeration: "\<rat> = range (of_rat o nat_to_rat_surj)"
using nat_to_rat_surj
by (auto simp add: image_def surj_def) (metis Rats_cases)

lemma (in measure_space) borel_measurable_less_borel_measurable:
assumes f: "f ∈ borel_measurable M"
assumes g: "g ∈ borel_measurable M"
shows "{w ∈ space M. f w < g w} ∈ sets M"

proof -
have "{w ∈ space M. f w < g w} =
(\<Union>r∈\<rat>. {w ∈ space M. f w < r} ∩ {w ∈ space M. r < g w })"

by (auto simp add: Rats_dense_in_real)
thus ?thesis using f g
by (simp add: borel_measurable_less_iff [of f]
borel_measurable_gr_iff [of g])
(blast intro: gen_countable_UN [OF rats_enumeration])

qed

lemma (in measure_space) borel_measurable_leq_borel_measurable:
assumes f: "f ∈ borel_measurable M"
assumes g: "g ∈ borel_measurable M"
shows "{w ∈ space M. f w ≤ g w} ∈ sets M"

proof -
have "{w ∈ space M. f w ≤ g w} = space M - {w ∈ space M. g w < f w}"
by auto
thus ?thesis using f g
by (simp add: borel_measurable_less_borel_measurable compl_sets)
qed

lemma (in measure_space) borel_measurable_eq_borel_measurable:
assumes f: "f ∈ borel_measurable M"
assumes g: "g ∈ borel_measurable M"
shows "{w ∈ space M. f w = g w} ∈ sets M"

proof -
have "{w ∈ space M. f w = g w} =
{w ∈ space M. f w ≤ g w} ∩ {w ∈ space M. g w ≤ f w}"

by auto
thus ?thesis using f g
by (simp add: borel_measurable_leq_borel_measurable Int)
qed

lemma (in measure_space) borel_measurable_neq_borel_measurable:
assumes f: "f ∈ borel_measurable M"
assumes g: "g ∈ borel_measurable M"
shows "{w ∈ space M. f w ≠ g w} ∈ sets M"

proof -
have "{w ∈ space M. f w ≠ g w} = space M - {w ∈ space M. f w = g w}"
by auto
thus ?thesis using f g
by (simp add: borel_measurable_eq_borel_measurable compl_sets)
qed

lemma (in measure_space) borel_measurable_add_borel_measurable:
assumes f: "f ∈ borel_measurable M"
assumes g: "g ∈ borel_measurable M"
shows "(λx. f x + g x) ∈ borel_measurable M"

proof -
have 1:"!!a. {w ∈ space M. a ≤ f w + g w} = {w ∈ space M. a + (g w) * -1 ≤ f w}"
by auto
have "!!a. (λw. a + (g w) * -1) ∈ borel_measurable M"
by (rule affine_borel_measurable [OF g])
hence "!!a. {w ∈ space M. (λw. a + (g w) * -1)(w) ≤ f w} ∈ sets M" using f
by (rule borel_measurable_leq_borel_measurable)
hence "!!a. {w ∈ space M. a ≤ f w + g w} ∈ sets M"
by (simp add: 1)
thus ?thesis
by (simp add: borel_measurable_ge_iff)
qed


lemma (in measure_space) borel_measurable_square:
assumes f: "f ∈ borel_measurable M"
shows "(λx. (f x)^2) ∈ borel_measurable M"

proof -
{
fix a
have "{w ∈ space M. (f w)² ≤ a} ∈ sets M"
proof (cases rule: linorder_cases [of a 0])
case less
hence "{w ∈ space M. (f w)² ≤ a} = {}"
by auto (metis less order_le_less_trans power2_less_0)
also have "... ∈ sets M"
by (rule empty_sets)
finally show ?thesis .
next
case equal
hence "{w ∈ space M. (f w)² ≤ a} =
{w ∈ space M. f w ≤ 0} ∩ {w ∈ space M. 0 ≤ f w}"

by auto
also have "... ∈ sets M"
apply (insert f)
apply (rule Int)
apply (simp add: borel_measurable_le_iff)
apply (simp add: borel_measurable_ge_iff)
done
finally show ?thesis .
next
case greater
have "∀x. (f x ^ 2 ≤ sqrt a ^ 2) = (- sqrt a ≤ f x & f x ≤ sqrt a)"
by (metis abs_le_interval_iff abs_of_pos greater real_sqrt_abs
real_sqrt_le_iff real_sqrt_power)

hence "{w ∈ space M. (f w)² ≤ a} =
{w ∈ space M. -(sqrt a) ≤ f w} ∩ {w ∈ space M. f w ≤ sqrt a}"

using greater by auto
also have "... ∈ sets M"
apply (insert f)
apply (rule Int)
apply (simp add: borel_measurable_ge_iff)
apply (simp add: borel_measurable_le_iff)
done
finally show ?thesis .
qed
}
thus ?thesis by (auto simp add: borel_measurable_le_iff)
qed

lemma times_eq_sum_squares:
fixes x::real
shows"x*y = ((x+y)^2)/4 - ((x-y)^ 2)/4"

by (simp add: power2_eq_square ring_distribs diff_divide_distrib [symmetric])


lemma (in measure_space) borel_measurable_uminus_borel_measurable:
assumes g: "g ∈ borel_measurable M"
shows "(λx. - g x) ∈ borel_measurable M"

proof -
have "(λx. - g x) = (λx. 0 + (g x) * -1)"
by simp
also have "... ∈ borel_measurable M"
by (fast intro: affine_borel_measurable g)
finally show ?thesis .
qed

lemma (in measure_space) borel_measurable_times_borel_measurable:
assumes f: "f ∈ borel_measurable M"
assumes g: "g ∈ borel_measurable M"
shows "(λx. f x * g x) ∈ borel_measurable M"

proof -
have 1: "(λx. 0 + (f x + g x)² * inverse 4) ∈ borel_measurable M"
by (fast intro: affine_borel_measurable borel_measurable_square
borel_measurable_add_borel_measurable f g)

have "(λx. -((f x + -g x) ^ 2 * inverse 4)) =
(λx. 0 + ((f x + -g x) ^ 2 * inverse -4))"

by (simp add: minus_divide_right)
also have "... ∈ borel_measurable M"
by (fast intro: affine_borel_measurable borel_measurable_square
borel_measurable_add_borel_measurable
borel_measurable_uminus_borel_measurable f g)

finally have 2: "(λx. -((f x + -g x) ^ 2 * inverse 4)) ∈ borel_measurable M" .
show ?thesis
apply (simp add: times_eq_sum_squares diff_def)
using 1 2 apply (simp add: borel_measurable_add_borel_measurable)
done
qed

lemma (in measure_space) borel_measurable_diff_borel_measurable:
assumes f: "f ∈ borel_measurable M"
assumes g: "g ∈ borel_measurable M"
shows "(λx. f x - g x) ∈ borel_measurable M"

unfolding diff_def
by (fast intro: borel_measurable_add_borel_measurable
borel_measurable_uminus_borel_measurable f g)


lemma (in measure_space) borel_measurable_setsum_borel_measurable:
assumes s: "finite s"
shows "(!!i. i ∈ s ==> f i ∈ borel_measurable M) ==> (λx. setsum (λi. f i x) s) ∈ borel_measurable M"
using s
proof (induct s)
case empty
thus ?case
by (simp add: borel_measurable_const)
next
case (insert x s)
thus ?case
by (auto simp add: borel_measurable_add_borel_measurable)
qed

lemma (in measure_space) borel_measurable_cong:
assumes "!! w. w ∈ space M ==> f w = g w"
shows "f ∈ borel_measurable M <-> g ∈ borel_measurable M"

using assms unfolding in_borel_measurable by (simp cong: vimage_inter_cong)

lemma (in measure_space) borel_measurable_inverse:
assumes "f ∈ borel_measurable M"
shows "(λx. inverse (f x)) ∈ borel_measurable M"

unfolding borel_measurable_ge_iff
proof (safe, rule linorder_cases)
fix a :: real assume "0 < a"
{ fix w
from `0 < a` have "a ≤ inverse (f w) <-> 0 < f w ∧ f w ≤ 1 / a"
by (metis inverse_eq_divide inverse_inverse_eq le_imp_inverse_le
less_le_trans zero_less_divide_1_iff)
}
hence "{w ∈ space M. a ≤ inverse (f w)} =
{w ∈ space M. 0 < f w} ∩ {w ∈ space M. f w ≤ 1 / a}"
by auto
with Int assms[unfolded borel_measurable_gr_iff]
assms[unfolded borel_measurable_le_iff]

show "{w ∈ space M. a ≤ inverse (f w)} ∈ sets M" by simp
next
fix a :: real assume "0 = a"
{ fix w have "a ≤ inverse (f w) <-> 0 ≤ f w"
unfolding `0 = a`[symmetric] by auto }
thus "{w ∈ space M. a ≤ inverse (f w)} ∈ sets M"
using assms[unfolded borel_measurable_ge_iff] by simp
next
fix a :: real assume "a < 0"
{ fix w
from `a < 0` have "a ≤ inverse (f w) <-> f w ≤ 1 / a ∨ 0 ≤ f w"
apply (cases "0 ≤ f w")
apply (metis inverse_eq_divide linorder_not_le xt1(8) xt1(9)
zero_le_divide_1_iff)

apply (metis inverse_eq_divide inverse_inverse_eq inverse_le_imp_le_neg
linorder_not_le order_refl order_trans)

done }
hence "{w ∈ space M. a ≤ inverse (f w)} =
{w ∈ space M. f w ≤ 1 / a} ∪ {w ∈ space M. 0 ≤ f w}"
by auto
with Un assms[unfolded borel_measurable_ge_iff]
assms[unfolded borel_measurable_le_iff]

show "{w ∈ space M. a ≤ inverse (f w)} ∈ sets M" by simp
qed

lemma (in measure_space) borel_measurable_divide:
assumes "f ∈ borel_measurable M"
and "g ∈ borel_measurable M"
shows "(λx. f x / g x) ∈ borel_measurable M"

unfolding field_divide_inverse
by (rule borel_measurable_inverse borel_measurable_times_borel_measurable assms)+

lemma (in measure_space) borel_measurable_vimage:
assumes borel: "f ∈ borel_measurable M"
shows "f -` {X} ∩ space M ∈ sets M"

proof -
have "{w ∈ space M. f w = X} = {w. f w = X} ∩ space M" by auto
with borel_measurable_eq_borel_measurable[OF borel borel_measurable_const, of X]
show ?thesis unfolding vimage_def by simp
qed

section "Monotone convergence"

definition mono_convergent where
"mono_convergent u f s ≡
(∀x∈s. incseq (λn. u n x)) ∧
(∀x ∈ s. (λi. u i x) ----> f x)"


definition "upclose f g x = max (f x) (g x)"

primrec mon_upclose where
"mon_upclose 0 u = u 0" |
"mon_upclose (Suc n) u = upclose (u (Suc n)) (mon_upclose n u)"


lemma mono_convergentD:
assumes "mono_convergent u f s" and "x ∈ s"
shows "incseq (λn. u n x)" and "(λi. u i x) ----> f x"

using assms unfolding mono_convergent_def by auto

lemma mono_convergentI:
assumes "!!x. x ∈ s ==> incseq (λn. u n x)"
assumes "!!x. x ∈ s ==> (λi. u i x) ----> f x"
shows "mono_convergent u f s"

using assms unfolding mono_convergent_def by auto

lemma (in measure_space) mono_convergent_borel_measurable:
assumes u: "!!n. u n ∈ borel_measurable M"
assumes mc: "mono_convergent u f (space M)"
shows "f ∈ borel_measurable M"

proof -
{
fix a
have 1: "!!w. w ∈ space M & f w <= a <-> w ∈ space M & (∀i. u i w <= a)"
proof safe
fix w i
assume w: "w ∈ space M" and f: "f w ≤ a"
hence "u i w ≤ f w"
by (auto intro: SEQ.incseq_le
simp add: mc [unfolded mono_convergent_def])

thus "u i w ≤ a" using f
by auto
next
fix w
assume w: "w ∈ space M" and u: "∀i. u i w ≤ a"
thus "f w ≤ a"
by (metis LIMSEQ_le_const2 mc [unfolded mono_convergent_def])
qed
have "{w ∈ space M. f w ≤ a} = {w ∈ space M. (∀i. u i w <= a)}"
by (simp add: 1)
also have "... = (\<Inter>i. {w ∈ space M. u i w ≤ a})"
by auto
also have "... ∈ sets M" using u
by (auto simp add: borel_measurable_le_iff intro: countable_INT)
finally have "{w ∈ space M. f w ≤ a} ∈ sets M" .
}
thus ?thesis
by (auto simp add: borel_measurable_le_iff)
qed

lemma mono_convergent_le:
assumes "mono_convergent u f s" and "t ∈ s"
shows "u n t ≤ f t"

using mono_convergentD[OF assms] by (auto intro!: incseq_le)

lemma mon_upclose_ex:
fixes u :: "nat => 'a => ('b::linorder)"
shows "∃n ≤ m. mon_upclose m u x = u n x"

proof (induct m)
case (Suc m)
then obtain n where "n ≤ m" and *: "mon_upclose m u x = u n x" by blast
thus ?case
proof (cases "u n x ≤ u (Suc m) x")
case True with min_max.sup_absorb1 show ?thesis
by (auto simp: * upclose_def intro!: exI[of _ "Suc m"])
next
case False
with min_max.sup_absorb2 `n ≤ m` show ?thesis
by (auto simp: * upclose_def intro!: exI[of _ n] min_max.sup_absorb2)
qed
qed simp

lemma mon_upclose_all:
fixes u :: "nat => 'a => ('b::linorder)"
assumes "m ≤ n"
shows "u m x ≤ mon_upclose n u x"

using assms proof (induct n)
case 0 thus ?case by auto
next
case (Suc n)
show ?case
proof (cases "m = Suc n")
case True thus ?thesis by (simp add: upclose_def)
next
case False
hence "m ≤ n" using `m ≤ Suc n` by simp
from Suc.hyps[OF this]
show ?thesis by (auto simp: upclose_def intro!: min_max.le_supI2)
qed
qed

lemma mono_convergent_limit:
fixes f :: "'a => real"
assumes "mono_convergent u f s" and "x ∈ s" and "0 < r"
shows "∃N. ∀n≥N. f x - u n x < r"

proof -
from LIMSEQ_D[OF mono_convergentD(2)[OF assms(1,2)] `0 < r`]
obtain N where "!!n. N ≤ n ==> ¦ u n x - f x ¦ < r" by auto
with mono_convergent_le[OF assms(1,2)] `0 < r`
show ?thesis by (auto intro!: exI[of _ N])
qed

lemma mon_upclose_le_mono_convergent:
assumes mc: "!!n. mono_convergent (λm. u m n) (f n) s" and "x ∈ s"
and "incseq (λn. f n x)"
shows "mon_upclose n (u n) x <= f n x"

proof -
obtain m where *: "mon_upclose n (u n) x = u n m x" and "m ≤ n"
using mon_upclose_ex[of n "u n" x] by auto
note this(1)
also have "u n m x ≤ f m x" using mono_convergent_le[OF assms(1,2)] .
also have "... ≤ f n x" using assms(3) `m ≤ n` unfolding incseq_def by auto
finally show ?thesis .
qed

lemma mon_upclose_mono_convergent:
assumes mc_u: "!!n. mono_convergent (λm. u m n) (f n) s"
and mc_f: "mono_convergent f h s"
shows "mono_convergent (λn. mon_upclose n (u n)) h s"

proof (rule mono_convergentI)
fix x assume "x ∈ s"
show "incseq (λn. mon_upclose n (u n) x)" unfolding incseq_def
proof safe
fix m n :: nat assume "m ≤ n"
obtain i where mon: "mon_upclose m (u m) x = u m i x" and "i ≤ m"
using mon_upclose_ex[of m "u m" x] by auto
hence "i ≤ n" using `m ≤ n` by auto

note mon
also have "u m i x ≤ u n i x"
using mono_convergentD(1)[OF mc_u `x ∈ s`] `m ≤ n`
unfolding incseq_def by auto
also have "u n i x ≤ mon_upclose n (u n) x"
using mon_upclose_all[OF `i ≤ n`, of "u n" x] .
finally show "mon_upclose m (u m) x ≤ mon_upclose n (u n) x" .
qed

show "(λi. mon_upclose i (u i) x) ----> h x"
proof (rule LIMSEQ_I)
fix r :: real assume "0 < r"
hence "0 < r / 2" by auto
from mono_convergent_limit[OF mc_f `x ∈ s` this]
obtain N where f_h: "!!n. N ≤ n ==> h x - f n x < r / 2" by auto

from mono_convergent_limit[OF mc_u `x ∈ s` `0 < r / 2`]
obtain N' where u_f: "!!n. N' ≤ n ==> f N x - u n N x < r / 2" by auto

show "∃N. ∀n≥N. norm (mon_upclose n (u n) x - h x) < r"
proof (rule exI[of _ "max N N'"], safe)
fix n assume "max N N' ≤ n"
hence "N ≤ n" and "N' ≤ n" by auto
hence "u n N x ≤ mon_upclose n (u n) x"
using mon_upclose_all[of N n "u n" x] by auto
moreover
from add_strict_mono[OF u_f[OF `N' ≤ n`] f_h[OF order_refl]]
have "h x - u n N x < r" by auto
ultimately have "h x - mon_upclose n (u n) x < r" by auto
moreover
obtain i where "mon_upclose n (u n) x = u n i x"
using mon_upclose_ex[of n "u n"] by blast
with mono_convergent_le[OF mc_u `x ∈ s`, of n i]
mono_convergent_le[OF mc_f `x ∈ s`, of i]

have "mon_upclose n (u n) x ≤ h x" by auto
ultimately
show "norm (mon_upclose n (u n) x - h x) < r" by auto
qed
qed
qed

lemma mono_conv_outgrow:
assumes "incseq x" "x ----> y" "z < y"
shows "∃b. ∀ a ≥ b. z < x a"

using assms
proof -
from assms have "y - z > 0" by simp
hence A: "∃n. (∀ m ≥ n. ¦ x m + - y ¦ < y - z)" using assms
unfolding incseq_def LIMSEQ_def dist_real_def diff_def
by simp
have "∀m. x m ≤ y" using incseq_le assms by auto
hence B: "∀m. ¦ x m + - y ¦ = y - x m"
by (metis abs_if abs_minus_add_cancel less_iff_diff_less_0 linorder_not_le diff_def)
from A B show ?thesis by auto
qed

end