Theory Functional_Spaces
theory Functional_Spaces
imports
"HOL-Analysis.Analysis"
"HOL-Library.Function_Algebras"
Ergodic_Theory.SG_Library_Complement
begin
section ‹Functions as a real vector space›
text ‹Many functional spaces are spaces of functions. To be able to use the following
framework, spaces of functions thus need to be endowed with a vector space
structure, coming from pointwise addition and multiplication.
Some instantiations for \verb+fun+ are already given in \verb+Function_Algebras.thy+ and
\verb+Lattices.thy+, we add several.›
text ‹\verb+minus_fun+ is already defined, in \verb+Lattices.thy+, but under the strange name
\verb+fun_Compl_def+. We restate the definition so that \verb+unfolding minus_fun_def+ works.
Same thing for \verb+minus_fun_def+. A better solution would be to have a coherent naming scheme
in \verb+Lattices.thy+.›
lemmas uminus_fun_def = fun_Compl_def
lemmas minus_fun_def = fun_diff_def
lemma fun_sum_apply:
fixes u::"'i ⇒ 'a ⇒ ('b::comm_monoid_add)"
shows "(sum u I) x = sum (λi. u i x) I"
by (induction I rule: infinite_finite_induct, auto)
instantiation "fun" :: (type, real_vector) real_vector
begin
definition scaleR_fun::"real ⇒ ('a ⇒ 'b) ⇒ 'a ⇒ 'b"
where "scaleR_fun = (λc f. (λx. c *⇩R f x))"
lemma scaleR_apply [simp, code]: "(c *⇩R f) x = c *⇩R (f x)"
by (simp add: scaleR_fun_def)
instance by (standard, auto simp add: scaleR_add_right scaleR_add_left)
end
lemmas divideR_apply = scaleR_apply
lemma [measurable]:
"0 ∈ borel_measurable M"
unfolding zero_fun_def by auto
lemma borel_measurable_const_scaleR' [measurable (raw)]:
"(f::('a ⇒ 'b::real_normed_vector)) ∈ borel_measurable M ⟹ c *⇩R f ∈ borel_measurable M"
unfolding scaleR_fun_def using borel_measurable_add by auto
lemma borel_measurable_add'[measurable (raw)]:
fixes f g :: "'a ⇒ 'b::{second_countable_topology, real_normed_vector}"
assumes f: "f ∈ borel_measurable M"
assumes g: "g ∈ borel_measurable M"
shows "f + g ∈ borel_measurable M"
unfolding plus_fun_def using assms by auto
lemma borel_measurable_uminus'[measurable (raw)]:
fixes f g :: "'a ⇒ 'b::{second_countable_topology, real_normed_vector}"
assumes f: "f ∈ borel_measurable M"
shows "-f ∈ borel_measurable M"
unfolding fun_Compl_def using assms by auto
lemma borel_measurable_diff'[measurable (raw)]:
fixes f g :: "'a ⇒ 'b::{second_countable_topology, real_normed_vector}"
assumes f: "f ∈ borel_measurable M"
assumes g: "g ∈ borel_measurable M"
shows "f - g ∈ borel_measurable M"
unfolding fun_diff_def using assms by auto
lemma borel_measurable_sum'[measurable (raw)]:
fixes f::"'i ⇒ 'a ⇒ 'b::{second_countable_topology, real_normed_vector}"
assumes "⋀i. i ∈ I ⟹ f i ∈ borel_measurable M"
shows "(∑i∈I. f i) ∈ borel_measurable M"
using borel_measurable_sum[of I f, OF assms] unfolding fun_sum_apply[symmetric] by simp
lemma zero_applied_to [simp]:
"(0::('a ⇒ ('b::real_vector))) x = 0"
unfolding zero_fun_def by simp
section ‹Quasinorms on function spaces›
text ‹A central feature of modern analysis is the use of various functional
spaces, and of results of functional analysis on them. Think for instance of
$L^p$ spaces, of Sobolev or Besov spaces, or variations around them. Here are
several relevant facts about this point of view:
\begin{itemize}
\item These spaces typically depend on one or several parameters.
This makes it difficult to play with type classes in a system without dependent
types.
\item The $L^p$ spaces are not spaces of functions (their elements are
equivalence classes of functions, where two functions are identified if they
coincide almost everywhere). However, in usual analysis proofs, one takes a
definite representative and works with it, never going to the equivalence class
point of view (which only becomes relevant when one wants to use the fact that
one has a Banach space at our disposal, to apply functional analytic tools).
\item It is important to describe how the spaces are related to each other,
with respect to inclusions or compact inclusions. For instance, one of the most
important theorems in analysis is Sobolev embedding theorem, describing when
one Sobolev space is included in another one. One also needs to be able to
take intersections or sums of Banach spaces, for instance to develop
interpolation theory.
\item Some other spaces play an important role in analysis, for instance the
weak $L^1$ space. This space only has a quasi-norm (i.e., its norm satisfies the
triangular inequality up to a fixed multiplicative constant). A general enough setting
should also encompass this kind of space. (One could argue that one should also consider more
general topologies such as Frechet spaces, to deal with Gevrey or analytic functions.
This is true, but considering quasi-norms already gives a wealth of applications).
\end{itemize}
Given these points, it seems that the most effective way of formalizing this
kind of question in Isabelle/HOL is to think of such a functional space not as
an abstract space or type, but as a subset of the space of all functions or of
all distributions. Functions that do not belong to the functional space
under consideration will then have infinite norm. Then inclusions, intersections,
and so on, become trivial to implement. Since the same object contains both the information
about the norm and the space where the norm is finite, it conforms to the customary habit in
mathematics of identifying the two of them, talking for instance about the $L^p$ space and the
$L^p$ norm.
All in all, this approach seems quite promising for ``real life analysis''.
›
subsection ‹Definition of quasinorms›
typedef (overloaded) ('a::real_vector) quasinorm = "{(C::real, N::('a ⇒ ennreal)). (C ≥ 1)
∧ (∀ x c. N (c *⇩R x) = ennreal ¦c¦ * N(x)) ∧ (∀ x y. N(x+y) ≤ C * N x + C * N y)}"
morphisms Rep_quasinorm quasinorm_of
proof
show "(1,(λx. 0)) ∈ {(C::real, N::('a ⇒ ennreal)). (C ≥ 1)
∧ (∀ x c. N (c *⇩R x) = ennreal ¦c¦ * N x) ∧ (∀ x y. N (x+y) ≤ C * N x + C * N y)}"
by auto
qed
definition eNorm::"'a quasinorm ⇒ ('a::real_vector) ⇒ ennreal"
where "eNorm N x = (snd (Rep_quasinorm N)) x"
definition defect::"('a::real_vector) quasinorm ⇒ real"
where "defect N = fst (Rep_quasinorm N)"
lemma eNorm_triangular_ineq:
"eNorm N (x + y) ≤ defect N * eNorm N x + defect N * eNorm N y"
unfolding eNorm_def defect_def using Rep_quasinorm[of N] by auto
lemma defect_ge_1:
"defect N ≥ 1"
unfolding defect_def using Rep_quasinorm[of N] by auto
lemma eNorm_cmult:
"eNorm N (c *⇩R x) = ennreal ¦c¦ * eNorm N x"
unfolding eNorm_def using Rep_quasinorm[of N] by auto
lemma eNorm_zero [simp]:
"eNorm N 0 = 0"
by (metis eNorm_cmult abs_zero ennreal_0 mult_zero_left real_vector.scale_zero_left)
lemma eNorm_uminus [simp]:
"eNorm N (-x) = eNorm N x"
using eNorm_cmult[of N "-1" x] by auto
lemma eNorm_sum:
"eNorm N (∑i∈{..<n}. u i) ≤ (∑i∈{..<n}. (defect N)^(Suc i) * eNorm N (u i))"
proof (cases "n=0")
case True
then show ?thesis by simp
next
case False
then obtain m where "n = Suc m" using not0_implies_Suc by blast
have "⋀v. eNorm N (∑i∈{..n}. v i) ≤ (∑i∈{..<n}. (defect N)^(Suc i) * eNorm N (v i)) + (defect N)^n * eNorm N (v n)" for n
proof (induction n)
case 0
then show ?case by simp
next
case (Suc n)
have *: "(defect N)^(Suc n) = (defect N)^n * ennreal(defect N)"
by (metis defect_ge_1 ennreal_le_iff ennreal_neg ennreal_power less_le not_less not_one_le_zero semiring_normalization_rules(28))
fix v::"nat ⇒ 'a"
define w where "w = (λi. if i = n then v n + v (Suc n) else v i)"
have "(∑i∈{..Suc n}. v i) = (∑i∈{..<n}. v i) + v n + v (Suc n)"
using lessThan_Suc_atMost sum.lessThan_Suc by auto
also have "... = (∑i∈{..<n}. w i) + w n" unfolding w_def by auto
finally have "(∑i∈{..Suc n}. v i) = (∑i∈{..n}. w i)"
by (metis lessThan_Suc_atMost sum.lessThan_Suc)
then have "eNorm N (∑i∈{..Suc n}. v i) = eNorm N (∑i∈{..n}. w i)" by simp
also have "... ≤ (∑i∈{..<n}. (defect N)^(Suc i) * eNorm N (w i)) + (defect N)^n * eNorm N (w n)"
using Suc.IH by auto
also have "... = (∑i∈{..<n}. (defect N)^(Suc i) * eNorm N (v i)) + (defect N)^n * eNorm N (v n + v (Suc n))"
unfolding w_def by auto
also have "... ≤ (∑i∈{..<n}. (defect N)^(Suc i) * eNorm N (v i)) +
(defect N)^n * (defect N * eNorm N (v n) + defect N * eNorm N (v (Suc n)))"
by (rule add_mono, simp, rule mult_left_mono, auto simp add: eNorm_triangular_ineq)
also have "... = (∑i∈{..<n}. (defect N)^(Suc i) * eNorm N (v i))
+ (defect N)^(Suc n) * eNorm N (v n) + (defect N)^(Suc n) * eNorm N (v (Suc n))"
unfolding * by (simp add: distrib_left semiring_normalization_rules(18))
also have "... = (∑i∈{..<Suc n}. (defect N)^(Suc i) * eNorm N (v i)) + (defect N)^(Suc n) * eNorm N (v (Suc n))"
by auto
finally show "eNorm N (∑i∈{..Suc n}. v i)
≤ (∑i<Suc n. ennreal (defect N ^ Suc i) * eNorm N (v i)) + ennreal (defect N ^ Suc n) * eNorm N (v (Suc n)) "
by simp
qed
then have "eNorm N (∑i∈{..<Suc m}. u i)
≤ (∑i∈{..<m}. (defect N)^(Suc i) * eNorm N (u i)) + (defect N)^m * eNorm N (u m)"
using lessThan_Suc_atMost by auto
also have "... ≤ (∑i∈{..<m}. (defect N)^(Suc i) * eNorm N (u i)) + (defect N)^(Suc m) * eNorm N (u m)"
apply (rule add_mono, auto intro!: mult_right_mono ennreal_leI)
using defect_ge_1 by (metis atMost_iff le_less lessThan_Suc_atMost lessThan_iff power_Suc power_increasing)
also have "... = (∑i∈{..<Suc m}. (defect N)^(Suc i) * eNorm N (u i))"
by auto
finally show "eNorm N (∑i∈{..<n}. u i) ≤ (∑i<n. ennreal (defect N ^ Suc i) * eNorm N (u i))"
unfolding ‹n = Suc m› by auto
qed
text ‹Quasinorms are often defined by taking a meaningful formula on a vector subspace,
and then extending by infinity elsewhere. Let us show that this results in a quasinorm on the
whole space.›
definition quasinorm_on::"('a set) ⇒ real ⇒ (('a::real_vector) ⇒ ennreal) ⇒ bool"
where "quasinorm_on F C N = (
(∀x y. (x ∈ F ∧ y ∈ F) ⟶ (x + y ∈ F) ∧ N (x+y) ≤ C * N x + C * N y)
∧ (∀c x. x ∈ F ⟶ c *⇩R x ∈ F ∧ N(c *⇩R x) = ¦c¦ * N x)
∧ C ≥ 1 ∧ 0 ∈ F)"
lemma quasinorm_of:
fixes N::"('a::real_vector) ⇒ ennreal" and C::real
assumes "quasinorm_on UNIV C N"
shows "eNorm (quasinorm_of (C,N)) x = N x"
"defect (quasinorm_of (C,N)) = C"
using assms unfolding eNorm_def defect_def quasinorm_on_def by (auto simp add: quasinorm_of_inverse)
lemma quasinorm_onI:
fixes N::"('a::real_vector) ⇒ ennreal" and C::real and F::"'a set"
assumes "⋀x y. x ∈ F ⟹ y ∈ F ⟹ x + y ∈ F"
"⋀x y. x ∈ F ⟹ y ∈ F ⟹ N (x + y) ≤ C * N x + C * N y"
"⋀c x. c ≠ 0 ⟹ x ∈ F ⟹ c *⇩R x ∈ F"
"⋀c x. c ≠ 0 ⟹ x ∈ F ⟹ N (c *⇩R x) ≤ ennreal ¦c¦ * N x"
"0 ∈ F" "N(0) = 0" "C ≥ 1"
shows "quasinorm_on F C N"
proof -
have "N(c *⇩R x) = ennreal ¦c¦ * N x" if "x ∈ F" for c x
proof (cases "c = 0")
case True
then show ?thesis using ‹N 0 = 0› by auto
next
case False
have "N((1/c) *⇩R (c *⇩R x)) ≤ ennreal (abs (1/c)) * N (c *⇩R x)"
apply (rule ‹⋀c x. c ≠ 0 ⟹ x ∈ F ⟹ N(c *⇩R x) ≤ ennreal ¦c¦ * N x›) using False assms that by auto
then have "N x ≤ ennreal (abs (1/c)) * N (c *⇩R x)" using False by auto
then have "ennreal ¦c¦ * N x ≤ ennreal ¦c¦ * ennreal (abs (1/c)) * N (c *⇩R x)"
by (simp add: mult.assoc mult_left_mono)
also have "... = N (c *⇩R x)" using ennreal_mult' abs_mult False
by (metis abs_ge_zero abs_one comm_monoid_mult_class.mult_1 ennreal_1 eq_divide_eq_1 field_class.field_divide_inverse)
finally show ?thesis
using ‹⋀c x. c ≠ 0 ⟹ x ∈ F ⟹ N(c *⇩R x) ≤ ennreal ¦c¦ * N x›[OF False ‹x ∈ F›] by auto
qed
then show ?thesis
unfolding quasinorm_on_def using assms by (auto, metis real_vector.scale_zero_left)
qed
lemma extend_quasinorm:
assumes "quasinorm_on F C N"
shows "quasinorm_on UNIV C (λx. if x ∈ F then N x else ∞)"
proof -
have *: "(if x + y ∈ F then N (x + y) else ∞)
≤ ennreal C * (if x ∈ F then N x else ∞) + ennreal C * (if y ∈ F then N y else ∞)" for x y
proof (cases "x ∈ F ∧ y ∈ F")
case True
then show ?thesis using assms unfolding quasinorm_on_def by auto
next
case False
moreover have "C ≥ 1" using assms unfolding quasinorm_on_def by auto
ultimately have *: "ennreal C * (if x ∈ F then N x else ∞) + ennreal C * (if y ∈ F then N y else ∞) = ∞"
using ennreal_mult_eq_top_iff by auto
show ?thesis by (simp add: *)
qed
show ?thesis
apply (rule quasinorm_onI)
using assms * unfolding quasinorm_on_def apply (auto simp add: ennreal_top_mult mult.commute)
by (metis abs_zero ennreal_0 mult_zero_right real_vector.scale_zero_right)
qed
subsection ‹The space and the zero space of a quasinorm›
text ‹The space of a quasinorm is the vector subspace where it is meaningful, i.e., finite.›
definition space⇩N::"('a::real_vector) quasinorm ⇒ 'a set"
where "space⇩N N = {f. eNorm N f < ∞}"
lemma spaceN_iff:
"x ∈ space⇩N N ⟷ eNorm N x < ∞"
unfolding space⇩N_def by simp
lemma spaceN_cmult [simp]:
assumes "x ∈ space⇩N N"
shows "c *⇩R x ∈ space⇩N N"
using assms unfolding spaceN_iff using eNorm_cmult[of N c x] by (simp add: ennreal_mult_less_top)
lemma spaceN_add [simp]:
assumes "x ∈ space⇩N N" "y ∈ space⇩N N"
shows "x + y ∈ space⇩N N"
proof -
have "eNorm N x < ∞" "eNorm N y < ∞" using assms unfolding space⇩N_def by auto
then have "defect N * eNorm N x + defect N * eNorm N y < ∞"
by (simp add: ennreal_mult_less_top)
then show ?thesis
unfolding space⇩N_def using eNorm_triangular_ineq[of N x y] le_less_trans by blast
qed
lemma spaceN_diff [simp]:
assumes "x ∈ space⇩N N" "y ∈ space⇩N N"
shows "x - y ∈ space⇩N N"
using spaceN_add[OF assms(1) spaceN_cmult[OF assms(2), of "-1"]] by auto
lemma spaceN_contains_zero [simp]:
"0 ∈ space⇩N N"
unfolding space⇩N_def by auto
lemma spaceN_sum [simp]:
assumes "⋀i. i ∈ I ⟹ x i ∈ space⇩N N"
shows "(∑i∈I. x i) ∈ space⇩N N"
using assms by (induction I rule: infinite_finite_induct, auto)
text ‹The zero space of a quasinorm is the vector subspace of vectors with zero norm. If one wants
to get a true metric space, one should quotient the space by the zero space.›
definition zero_space⇩N::"('a::real_vector) quasinorm ⇒ 'a set"
where "zero_space⇩N N = {f. eNorm N f = 0}"
lemma zero_spaceN_iff:
"x ∈ zero_space⇩N N ⟷ eNorm N x = 0"
unfolding zero_space⇩N_def by simp
lemma zero_spaceN_cmult:
assumes "x ∈ zero_space⇩N N"
shows "c *⇩R x ∈ zero_space⇩N N"
using assms unfolding zero_spaceN_iff using eNorm_cmult[of N c x] by simp
lemma zero_spaceN_add:
assumes "x ∈ zero_space⇩N N" "y ∈ zero_space⇩N N"
shows "x + y ∈ zero_space⇩N N"
proof -
have "eNorm N x = 0" "eNorm N y = 0" using assms unfolding zero_space⇩N_def by auto
then have "defect N * eNorm N x + defect N * eNorm N y = 0" by auto
then show ?thesis
unfolding zero_spaceN_iff using eNorm_triangular_ineq[of N x y] by auto
qed
lemma zero_spaceN_diff:
assumes "x ∈ zero_space⇩N N" "y ∈ zero_space⇩N N"
shows "x - y ∈ zero_space⇩N N"
using zero_spaceN_add[OF assms(1) zero_spaceN_cmult[OF assms(2), of "-1"]] by auto
lemma zero_spaceN_subset_spaceN:
"zero_space⇩N N ⊆ space⇩N N"
by (simp add: spaceN_iff zero_spaceN_iff subset_eq)
text ‹On the space, the norms are finite. Hence, it is much more convenient to work there with
a real valued version of the norm. We use Norm with a capital N to distinguish it from norms
in a (type class) banach space.›
definition Norm::"'a quasinorm ⇒ ('a::real_vector) ⇒ real"
where "Norm N x = enn2real (eNorm N x)"
lemma Norm_nonneg [simp]:
"Norm N x ≥ 0"
unfolding Norm_def by auto
lemma Norm_zero [simp]:
"Norm N 0 = 0"
unfolding Norm_def by auto
lemma Norm_uminus [simp]:
"Norm N (-x) = Norm N x"
unfolding Norm_def by auto
lemma eNorm_Norm:
assumes "x ∈ space⇩N N"
shows "eNorm N x = ennreal (Norm N x)"
using assms unfolding Norm_def by (simp add: spaceN_iff)
lemma eNorm_Norm':
assumes "x ∉ space⇩N N"
shows "Norm N x = 0"
using assms unfolding Norm_def apply (auto simp add: spaceN_iff)
using top.not_eq_extremum by fastforce
lemma Norm_cmult:
"Norm N (c *⇩R x) = abs c * Norm N x"
unfolding Norm_def unfolding eNorm_cmult by (simp add: enn2real_mult)
lemma Norm_triangular_ineq:
assumes "x ∈ space⇩N N"
shows "Norm N (x + y) ≤ defect N * Norm N x + defect N * Norm N y"
proof (cases "y ∈ space⇩N N")
case True
have *: "defect N * Norm N x + defect N * Norm N y ≥ 1 * 0 + 1 * 0"
apply (rule add_mono) by (rule mult_mono'[OF defect_ge_1 Norm_nonneg], simp, simp)+
have "ennreal (Norm N (x + y)) = eNorm N (x+y)"
using eNorm_Norm[OF spaceN_add[OF assms True]] by auto
also have "... ≤ defect N * eNorm N x + defect N * eNorm N y"
using eNorm_triangular_ineq[of N x y] by auto
also have "... = defect N * ennreal(Norm N x) + defect N * ennreal(Norm N y)"
using eNorm_Norm assms True by metis
also have "... = ennreal(defect N * Norm N x + defect N * Norm N y)"
using ennreal_mult ennreal_plus Norm_nonneg defect_ge_1
by (metis (no_types, opaque_lifting) ennreal_eq_0_iff less_le ennreal_ge_1 ennreal_mult' le_less_linear not_one_le_zero semiring_normalization_rules(34))
finally show ?thesis
apply (subst ennreal_le_iff[symmetric]) using * by auto
next
case False
have "x + y ∉ space⇩N N"
proof (rule ccontr)
assume "¬ (x + y ∉ space⇩N N)"
then have "x + y ∈ space⇩N N" by simp
have "y ∈ space⇩N N" using spaceN_diff[OF ‹x + y ∈ space⇩N N› assms] by auto
then show False using False by simp
qed
then have "Norm N (x+y) = 0" unfolding Norm_def using spaceN_iff top.not_eq_extremum by force
moreover have "defect N * Norm N x + defect N * Norm N y ≥ 1 * 0 + 1 * 0"
apply (rule add_mono) by (rule mult_mono'[OF defect_ge_1 Norm_nonneg], simp, simp)+
ultimately show ?thesis by simp
qed
lemma Norm_triangular_ineq_diff:
assumes "x ∈ space⇩N N"
shows "Norm N (x - y) ≤ defect N * Norm N x + defect N * Norm N y"
using Norm_triangular_ineq[OF assms, of "-y"] by auto
lemma zero_spaceN_iff':
"x ∈ zero_space⇩N N ⟷ (x ∈ space⇩N N ∧ Norm N x = 0)"
using eNorm_Norm unfolding space⇩N_def zero_space⇩N_def by (auto simp add: Norm_def, fastforce)
lemma Norm_sum:
assumes "⋀i. i < n ⟹ u i ∈ space⇩N N"
shows "Norm N (∑i∈{..<n}. u i) ≤ (∑i∈{..<n}. (defect N)^(Suc i) * Norm N (u i))"
proof -
have *: "0 ≤ defect N * defect N ^ i * Norm N (u i)" for i
by (meson Norm_nonneg defect_ge_1 dual_order.trans linear mult_nonneg_nonneg not_one_le_zero zero_le_power)
have "ennreal (Norm N (∑i∈{..<n}. u i)) = eNorm N (∑i∈{..<n}. u i)"
apply (rule eNorm_Norm[symmetric], rule spaceN_sum) using assms by auto
also have "... ≤ (∑i∈{..<n}. (defect N)^(Suc i) * eNorm N (u i))"
using eNorm_sum by simp
also have "... = (∑i∈{..<n}. (defect N)^(Suc i) * ennreal (Norm N (u i)))"
using eNorm_Norm[OF assms] by auto
also have "... = (∑i∈{..<n}. ennreal((defect N)^(Suc i) * Norm N (u i)))"
by (subst ennreal_mult'', auto)
also have "... = ennreal (∑i∈{..<n}. (defect N)^(Suc i) * Norm N (u i))"
by (auto intro!: sum_ennreal simp add: *)
finally have **: "ennreal (Norm N (∑i∈{..<n}. u i)) ≤ ennreal (∑i∈{..<n}. (defect N)^(Suc i) * Norm N (u i))"
by simp
show ?thesis
apply (subst ennreal_le_iff[symmetric], rule sum_nonneg) using * ** by auto
qed
subsection ‹An example: the ambient norm in a normed vector space›
definition N_of_norm::"'a::real_normed_vector quasinorm"
where "N_of_norm = quasinorm_of (1, λf. norm f)"
lemma N_of_norm:
"eNorm N_of_norm f = ennreal (norm f)"
"Norm N_of_norm f = norm f"
"defect (N_of_norm) = 1"
proof -
have *: "quasinorm_on UNIV 1 (λf. norm f)"
by (rule quasinorm_onI, auto simp add: ennreal_mult', metis ennreal_leI ennreal_plus norm_imp_pos_and_ge norm_triangle_ineq)
show "eNorm N_of_norm f = ennreal (norm f)"
"defect (N_of_norm) = 1"
unfolding N_of_norm_def using quasinorm_of[OF *] by auto
then show "Norm N_of_norm f = norm f" unfolding Norm_def by auto
qed
lemma N_of_norm_space [simp]:
"space⇩N N_of_norm = UNIV"
unfolding space⇩N_def apply auto unfolding N_of_norm(1) by auto
lemma N_of_norm_zero_space [simp]:
"zero_space⇩N N_of_norm = {0}"
unfolding zero_space⇩N_def apply auto unfolding N_of_norm(1) by auto
subsection ‹An example: the space of bounded continuous functions from a topological space to a normed
real vector space›
text ‹The Banach space of bounded continuous functions is defined in
\verb+Bounded_Continuous_Function.thy+, as a type \verb+bcontfun+. We import very quickly the
results proved in this file to the current framework.›
definition bcontfun⇩N::"('a::topological_space ⇒ 'b::real_normed_vector) quasinorm"
where "bcontfun⇩N = quasinorm_of (1, λf. if f ∈ bcontfun then norm(Bcontfun f) else (∞::ennreal))"
lemma bcontfun⇩N:
fixes f::"('a::topological_space ⇒ 'b::real_normed_vector)"
shows "eNorm bcontfun⇩N f = (if f ∈ bcontfun then norm(Bcontfun f) else (∞::ennreal))"
"Norm bcontfun⇩N f = (if f ∈ bcontfun then norm(Bcontfun f) else 0)"
"defect (bcontfun⇩N::(('a ⇒ 'b) quasinorm)) = 1"
proof -
have *: "quasinorm_on bcontfun 1 (λ(f::('a ⇒ 'b)). norm(Bcontfun f))"
proof (rule quasinorm_onI, auto)
fix f g::"'a ⇒ 'b" assume H: "f ∈ bcontfun" "g ∈ bcontfun"
then show "f + g ∈ bcontfun" unfolding plus_fun_def by (simp add: plus_cont)
have *: "Bcontfun(f + g) = Bcontfun f + Bcontfun g"
using H
by (auto simp: eq_onp_def plus_fun_def bcontfun_def intro!: plus_bcontfun.abs_eq[symmetric])
show "ennreal (norm (Bcontfun (f + g))) ≤ ennreal (norm (Bcontfun f)) + ennreal (norm (Bcontfun g))"
unfolding * using ennreal_leI[OF norm_triangle_ineq] by auto
next
fix c::real and f::"'a ⇒ 'b" assume H: "f ∈ bcontfun"
then show "c *⇩R f ∈ bcontfun" unfolding scaleR_fun_def by (simp add: scaleR_cont)
have *: "Bcontfun(c *⇩R f) = c *⇩R Bcontfun f"
using H
by (auto simp: eq_onp_def scaleR_fun_def bcontfun_def intro!: scaleR_bcontfun.abs_eq[symmetric])
show "ennreal (norm (Bcontfun (c *⇩R f))) ≤ ennreal ¦c¦ * ennreal (norm (Bcontfun f))"
unfolding * by (simp add: ennreal_mult'')
next
show "(0::'a⇒'b) ∈ bcontfun" "Bcontfun 0 = 0"
unfolding zero_fun_def zero_bcontfun_def by (auto simp add: const_bcontfun)
qed
have **: "quasinorm_on UNIV 1 (λ(f::'a⇒'b). if f ∈ bcontfun then norm(Bcontfun f) else (∞::ennreal))"
by (rule extend_quasinorm[OF *])
show "eNorm bcontfun⇩N f = (if f ∈ bcontfun then norm(Bcontfun f) else (∞::ennreal))"
"defect (bcontfun⇩N::('a ⇒ 'b) quasinorm) = 1"
using quasinorm_of[OF **] unfolding bcontfun⇩N_def by auto
then show "Norm bcontfun⇩N f = (if f ∈ bcontfun then norm(Bcontfun f) else 0)"
unfolding Norm_def by auto
qed
lemma bcontfun⇩N_space:
"space⇩N bcontfun⇩N = bcontfun"
using bcontfun⇩N(1) by (metis (no_types, lifting) Collect_cong bcontfun_def enn2real_top ennreal_0
ennreal_enn2real ennreal_less_top ennreal_zero_neq_top infinity_ennreal_def mem_Collect_eq space⇩N_def)
lemma bcontfun⇩N_zero_space:
"zero_space⇩N bcontfun⇩N = {0}"
apply (auto simp add: zero_spaceN_iff)
by (metis Bcontfun_inject bcontfun⇩N(1) eNorm_zero ennreal_eq_zero_iff ennreal_zero_neq_top infinity_ennreal_def norm_eq_zero norm_imp_pos_and_ge)
lemma bcontfun⇩ND:
assumes "f ∈ space⇩N bcontfun⇩N"
shows "continuous_on UNIV f"
"⋀x. norm(f x) ≤ Norm bcontfun⇩N f"
proof-
have "f ∈ bcontfun" using assms unfolding bcontfun⇩N_space by simp
then show "continuous_on UNIV f" unfolding bcontfun_def by auto
show "⋀x. norm(f x) ≤ Norm bcontfun⇩N f"
using norm_bounded bcontfun⇩N(2) ‹f ∈ bcontfun› by (metis Bcontfun_inverse)
qed
lemma bcontfun⇩NI:
assumes "continuous_on UNIV f"
"⋀x. norm(f x) ≤ C"
shows "f ∈ space⇩N bcontfun⇩N"
"Norm bcontfun⇩N f ≤ C"
proof -
have "f ∈ bcontfun" using assms bcontfun_normI by blast
then show "f ∈ space⇩N bcontfun⇩N" unfolding bcontfun⇩N_space by simp
show "Norm bcontfun⇩N f ≤ C" unfolding bcontfun⇩N(2) using ‹f ∈ bcontfun› apply auto
using assms(2) by (metis apply_bcontfun_cases apply_bcontfun_inverse norm_bound)
qed
subsection ‹Continuous inclusions between functional spaces›
text ‹Continuous inclusions between functional spaces are now defined›
instantiation quasinorm:: (real_vector) preorder
begin
definition less_eq_quasinorm::"'a quasinorm ⇒ 'a quasinorm ⇒ bool"
where "less_eq_quasinorm N1 N2 = (∃C≥(0::real). ∀f. eNorm N2 f ≤ C * eNorm N1 f)"
definition less_quasinorm::"'a quasinorm ⇒ 'a quasinorm ⇒ bool"
where "less_quasinorm N1 N2 = (less_eq N1 N2 ∧ (¬ less_eq N2 N1))"
instance proof -
have E: "N ≤ N" for N::"'a quasinorm"
unfolding less_eq_quasinorm_def by (rule exI[of _ 1], auto)
have T: "N1 ≤ N3" if "N1 ≤ N2" "N2 ≤ N3" for N1 N2 N3::"'a quasinorm"
proof -
obtain C C' where *: "⋀f. eNorm N2 f ≤ ennreal C * eNorm N1 f"
"⋀f. eNorm N3 f ≤ ennreal C' * eNorm N2 f"
"C ≥ 0" "C' ≥ 0"
using ‹N1 ≤ N2› ‹N2 ≤ N3› unfolding less_eq_quasinorm_def by metis
{
fix f
have "eNorm N3 f ≤ ennreal C' * ennreal C * eNorm N1 f"
by (metis *(1)[of f] *(2)[of f] mult.commute mult.left_commute mult_left_mono order_trans zero_le)
also have "... = ennreal(C' * C) * eNorm N1 f"
using ‹C ≥ 0› ‹C' ≥ 0› ennreal_mult by auto
finally have "eNorm N3 f ≤ ennreal(C' * C) * eNorm N1 f" by simp
}
then show ?thesis
unfolding less_eq_quasinorm_def using ‹C ≥ 0› ‹C' ≥ 0› zero_le_mult_iff by auto
qed
show "OFCLASS('a quasinorm, preorder_class)"
apply standard
unfolding less_quasinorm_def apply simp
using E apply fast
using T apply fast
done
qed
end
abbreviation quasinorm_subset :: "('a::real_vector) quasinorm ⇒ 'a quasinorm ⇒ bool"
where "quasinorm_subset ≡ less"
abbreviation quasinorm_subset_eq :: "('a::real_vector) quasinorm ⇒ 'a quasinorm ⇒ bool"
where "quasinorm_subset_eq ≡ less_eq"
notation
quasinorm_subset ("'(⊂⇩N')") and
quasinorm_subset ("(_/ ⊂⇩N _)" [51, 51] 50) and
quasinorm_subset_eq ("'(⊆⇩N')") and
quasinorm_subset_eq ("(_/ ⊆⇩N _)" [51, 51] 50)
lemma quasinorm_subsetD:
assumes "N1 ⊆⇩N N2"
shows "∃C≥(0::real). ∀f. eNorm N2 f ≤ C * eNorm N1 f"
using assms unfolding less_eq_quasinorm_def by auto
lemma quasinorm_subsetI:
assumes "⋀f. f ∈ space⇩N N1 ⟹ eNorm N2 f ≤ ennreal C * eNorm N1 f"
shows "N1 ⊆⇩N N2"
proof -
have "eNorm N2 f ≤ ennreal (max C 1) * eNorm N1 f" for f
proof (cases "f ∈ space⇩N N1")
case True
then show ?thesis using assms[OF ‹f ∈ space⇩N N1›]
by (metis (no_types, opaque_lifting) dual_order.trans ennreal_leI max.cobounded2 max.commute
mult.commute ordered_comm_semiring_class.comm_mult_left_mono zero_le)
next
case False
then show ?thesis using spaceN_iff
by (metis ennreal_ge_1 ennreal_mult_less_top infinity_ennreal_def max.cobounded1
max.commute not_le not_one_le_zero top.not_eq_extremum)
qed
then show ?thesis unfolding less_eq_quasinorm_def
by (metis ennreal_max_0' max.cobounded2)
qed
lemma quasinorm_subsetI':
assumes "⋀f. f ∈ space⇩N N1 ⟹ f ∈ space⇩N N2"
"⋀f. f ∈ space⇩N N1 ⟹ Norm N2 f ≤ C * Norm N1 f"
shows "N1 ⊆⇩N N2"
proof (rule quasinorm_subsetI)
fix f assume "f ∈ space⇩N N1"
then have "f ∈ space⇩N N2" using assms(1) by simp
then have "eNorm N2 f = ennreal(Norm N2 f)" using eNorm_Norm by auto
also have "... ≤ ennreal(C * Norm N1 f)"
using assms(2)[OF ‹f ∈ space⇩N N1›] ennreal_leI by blast
also have "... = ennreal C * ennreal(Norm N1 f)"
using ennreal_mult'' by auto
also have "... = ennreal C * eNorm N1 f"
using eNorm_Norm[OF ‹f ∈ space⇩N N1›] by auto
finally show "eNorm N2 f ≤ ennreal C * eNorm N1 f"
by simp
qed
lemma quasinorm_subset_space:
assumes "N1 ⊆⇩N N2"
shows "space⇩N N1 ⊆ space⇩N N2"
using assms unfolding space⇩N_def less_eq_quasinorm_def
by (auto, metis ennreal_mult_eq_top_iff ennreal_neq_top less_le top.extremum_strict top.not_eq_extremum)
lemma quasinorm_subset_Norm_eNorm:
assumes "f ∈ space⇩N N1 ⟹ Norm N2 f ≤ C * Norm N1 f"
"N1 ⊆⇩N N2"
"C > 0"
shows "eNorm N2 f ≤ ennreal C * eNorm N1 f"
proof (cases "f ∈ space⇩N N1")
case True
then have "f ∈ space⇩N N2" using quasinorm_subset_space[OF ‹N1 ⊆⇩N N2›] by auto
then show ?thesis
using eNorm_Norm[OF True] eNorm_Norm assms(1)[OF True] by (metis Norm_nonneg ennreal_leI ennreal_mult'')
next
case False
then show ?thesis using ‹C > 0›
by (metis ennreal_eq_zero_iff ennreal_mult_eq_top_iff infinity_ennreal_def less_imp_le neq_top_trans not_le spaceN_iff)
qed
lemma quasinorm_subset_zero_space:
assumes "N1 ⊆⇩N N2"
shows "zero_space⇩N N1 ⊆ zero_space⇩N N2"
using assms unfolding zero_space⇩N_def less_eq_quasinorm_def
by (auto, metis le_zero_eq mult_zero_right)
text ‹We would like to define the equivalence relation associated to the above order, i.e., the
equivalence between norms. This is not equality, so we do not have a true order, but nevertheless
this is handy, and not standard in a preorder in Isabelle. The file Library/Preorder.thy defines
such an equivalence relation, but including it breaks some proofs so we go the naive way.›
definition quasinorm_equivalent::"('a::real_vector) quasinorm ⇒ 'a quasinorm ⇒ bool" (infix "=⇩N" 60)
where "quasinorm_equivalent N1 N2 = ((N1 ⊆⇩N N2) ∧ (N2 ⊆⇩N N1))"
lemma quasinorm_equivalent_sym [sym]:
assumes "N1 =⇩N N2"
shows "N2 =⇩N N1"
using assms unfolding quasinorm_equivalent_def by auto
lemma quasinorm_equivalent_trans [trans]:
assumes "N1 =⇩N N2" "N2 =⇩N N3"
shows "N1 =⇩N N3"
using assms order_trans unfolding quasinorm_equivalent_def by blast
subsection ‹The intersection and the sum of two functional spaces›
text ‹In this paragraph, we define the intersection and the sum of two functional spaces.
In terms of the order introduced above, this corresponds to the minimum and the maximum.
More important, these are the first two examples of interpolation spaces between two
functional spaces, and they are central as all the other ones are built using them.›
definition quasinorm_intersection::"('a::real_vector) quasinorm ⇒ 'a quasinorm ⇒ 'a quasinorm" (infix "∩⇩N" 70)
where "quasinorm_intersection N1 N2 = quasinorm_of (max (defect N1) (defect N2), λf. eNorm N1 f + eNorm N2 f)"
lemma quasinorm_intersection:
"eNorm (N1 ∩⇩N N2) f = eNorm N1 f + eNorm N2 f"
"defect (N1 ∩⇩N N2) = max (defect N1) (defect N2)"
proof -
have T: "eNorm N1 (x + y) + eNorm N2 (x + y) ≤
ennreal (max (defect N1) (defect N2)) * (eNorm N1 x + eNorm N2 x) + ennreal (max (defect N1) (defect N2)) * (eNorm N1 y + eNorm N2 y)" for x y
proof -
have "eNorm N1 (x + y) ≤ ennreal (max (defect N1) (defect N2)) * eNorm N1 x + ennreal (max (defect N1) (defect N2)) * eNorm N1 y"
using eNorm_triangular_ineq[of N1 x y] by (metis (no_types) max_def distrib_left ennreal_leI mult_right_mono order_trans zero_le)
moreover have "eNorm N2 (x + y) ≤ ennreal (max (defect N1) (defect N2)) * eNorm N2 x + ennreal (max (defect N1) (defect N2)) * eNorm N2 y"
using eNorm_triangular_ineq[of N2 x y] by (metis (no_types) max_def max.commute distrib_left ennreal_leI mult_right_mono order_trans zero_le)
ultimately have "eNorm N1 (x + y) + eNorm N2 (x + y) ≤ ennreal (max (defect N1) (defect N2)) * (eNorm N1 x + eNorm N1 y + (eNorm N2 x + eNorm N2 y))"
by (simp add: add_mono_thms_linordered_semiring(1) distrib_left)
then show ?thesis
by (simp add: ab_semigroup_add_class.add_ac(1) add.left_commute distrib_left)
qed
have H: "eNorm N1 (c *⇩R x) + eNorm N2 (c *⇩R x) ≤ ennreal ¦c¦ * (eNorm N1 x + eNorm N2 x)" for c x
by (simp add: eNorm_cmult[of N1 c x] eNorm_cmult[of N2 c x] distrib_left)
have *: "quasinorm_on UNIV (max (defect N1) (defect N2)) (λf. eNorm N1 f + eNorm N2 f)"
apply (rule quasinorm_onI) using T H defect_ge_1[of N1] defect_ge_1[of N2] by auto
show "defect (N1 ∩⇩N N2) = max (defect N1) (defect N2)"
"eNorm (N1 ∩⇩N N2) f = eNorm N1 f + eNorm N2 f"
unfolding quasinorm_intersection_def using quasinorm_of[OF *] by auto
qed
lemma quasinorm_intersection_commute:
"N1 ∩⇩N N2 = N2 ∩⇩N N1"
unfolding quasinorm_intersection_def max.commute[of "defect N1"] add.commute[of "eNorm N1 _"] by simp
lemma quasinorm_intersection_space:
"space⇩N (N1 ∩⇩N N2) = space⇩N N1 ∩ space⇩N N2"
apply auto unfolding quasinorm_intersection(1) spaceN_iff by auto
lemma quasinorm_intersection_zero_space:
"zero_space⇩N (N1 ∩⇩N N2) = zero_space⇩N N1 ∩ zero_space⇩N N2"
apply auto unfolding quasinorm_intersection(1) zero_spaceN_iff by (auto simp add: add_eq_0_iff_both_eq_0)
lemma quasinorm_intersection_subset:
"N1 ∩⇩N N2 ⊆⇩N N1" "N1 ∩⇩N N2 ⊆⇩N N2"
by (rule quasinorm_subsetI[of _ _ 1], auto simp add: quasinorm_intersection(1))+
lemma quasinorm_intersection_minimum:
assumes "N ⊆⇩N N1" "N ⊆⇩N N2"
shows "N ⊆⇩N N1 ∩⇩N N2"
proof -
obtain C1 C2::real where *: "⋀f. eNorm N1 f ≤ C1 * eNorm N f"
"⋀f. eNorm N2 f ≤ C2 * eNorm N f"
"C1 ≥ 0" "C2 ≥ 0"
using quasinorm_subsetD[OF assms(1)] quasinorm_subsetD[OF assms(2)] by blast
have **: "eNorm (N1 ∩⇩N N2) f ≤ (C1 + C2) * eNorm N f" for f
unfolding quasinorm_intersection(1) using add_mono[OF *(1) *(2)] by (simp add: distrib_right *)
show ?thesis
apply (rule quasinorm_subsetI) using ** by auto
qed
lemma quasinorm_intersection_assoc:
"(N1 ∩⇩N N2) ∩⇩N N3 =⇩N N1 ∩⇩N (N2 ∩⇩N N3)"
unfolding quasinorm_equivalent_def by (meson order_trans quasinorm_intersection_minimum quasinorm_intersection_subset)
definition quasinorm_sum::"('a::real_vector) quasinorm ⇒ 'a quasinorm ⇒ 'a quasinorm" (infix "+⇩N" 70)
where "quasinorm_sum N1 N2 = quasinorm_of (max (defect N1) (defect N2), λf. Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2})"
lemma quasinorm_sum:
"eNorm (N1 +⇩N N2) f = Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}"
"defect (N1 +⇩N N2) = max (defect N1) (defect N2)"
proof -
define N where "N = (λf. Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2})"
have T: "N (f+g) ≤
ennreal (max (defect N1) (defect N2)) * N f + ennreal (max (defect N1) (defect N2)) * N g" for f g
proof -
have "∃u. (∀n. u n ∈ {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}) ∧ u ⇢ Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}"
by (rule Inf_as_limit, auto, rule exI[of _ "f"], rule exI[of _ 0], auto)
then obtain uf where uf: "⋀n. uf n ∈ {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}"
"uf ⇢ Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}"
by blast
have "∃f1 f2. ∀n. uf n = eNorm N1 (f1 n) + eNorm N2 (f2 n) ∧ f = f1 n + f2 n"
using uf(1) by (subst choice_iff[symmetric])+ blast
then obtain f1 f2 where F: "⋀n. uf n = eNorm N1 (f1 n) + eNorm N2 (f2 n)" "⋀n. f = f1 n + f2 n"
by blast
have "∃u. (∀n. u n ∈ {eNorm N1 g1 + eNorm N2 g2| g1 g2. g = g1 + g2}) ∧ u ⇢ Inf {eNorm N1 g1 + eNorm N2 g2| g1 g2. g = g1 + g2}"
by (rule Inf_as_limit, auto, rule exI[of _ "g"], rule exI[of _ 0], auto)
then obtain ug where ug: "⋀n. ug n ∈ {eNorm N1 g1 + eNorm N2 g2| g1 g2. g = g1 + g2}"
"ug ⇢ Inf {eNorm N1 g1 + eNorm N2 g2| g1 g2. g = g1 + g2}"
by blast
have "∃g1 g2. ∀n. ug n = eNorm N1 (g1 n) + eNorm N2 (g2 n) ∧ g = g1 n + g2 n"
using ug(1) by (subst choice_iff[symmetric])+ blast
then obtain g1 g2 where G: "⋀n. ug n = eNorm N1 (g1 n) + eNorm N2 (g2 n)" "⋀n. g = g1 n + g2 n"
by blast
define h1 where "h1 = (λn. f1 n + g1 n)"
define h2 where "h2 = (λn. f2 n + g2 n)"
have *: "f + g = h1 n + h2 n" for n
unfolding h1_def h2_def using F(2) G(2) by (auto simp add: algebra_simps)
have "N (f+g) ≤ ennreal (max (defect N1) (defect N2)) * (uf n + ug n)" for n
proof -
have "N (f+g) ≤ eNorm N1 (h1 n) + eNorm N2 (h2 n)"
unfolding N_def apply (rule Inf_lower, auto, rule exI[of _ "h1 n"], rule exI[of _ "h2 n"])
using * by auto
also have "... ≤ ennreal (defect N1) * eNorm N1 (f1 n) + ennreal (defect N1) * eNorm N1 (g1 n)
+ (ennreal (defect N2) * eNorm N2 (f2 n) + ennreal (defect N2) * eNorm N2 (g2 n))"
unfolding h1_def h2_def apply (rule add_mono) using eNorm_triangular_ineq by auto
also have "... ≤ (ennreal (max (defect N1) (defect N2)) * eNorm N1 (f1 n) + ennreal (max (defect N1) (defect N2)) * eNorm N1 (g1 n))
+ (ennreal (max (defect N1) (defect N2)) * eNorm N2 (f2 n) + ennreal (max (defect N1) (defect N2)) * eNorm N2 (g2 n))"
by (auto intro!: add_mono mult_mono ennreal_leI)
also have "... = ennreal (max (defect N1) (defect N2)) * (uf n + ug n)"
unfolding F(1) G(1) by (auto simp add: algebra_simps)
finally show ?thesis by simp
qed
moreover have "... ⇢ ennreal (max (defect N1) (defect N2)) * (N f + N g)"
unfolding N_def by (auto intro!: tendsto_intros simp add: uf(2) ug(2))
ultimately have "N (f+g) ≤ ennreal (max (defect N1) (defect N2)) * (N f + N g)"
using LIMSEQ_le_const by blast
then show ?thesis by (auto simp add: algebra_simps)
qed
have H: "N (c *⇩R f) ≤ ennreal ¦c¦ * N f" for c f
proof -
have "∃u. (∀n. u n ∈ {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}) ∧ u ⇢ Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}"
by (rule Inf_as_limit, auto, rule exI[of _ "f"], rule exI[of _ 0], auto)
then obtain uf where uf: "⋀n. uf n ∈ {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}"
"uf ⇢ Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}"
by blast
have "∃f1 f2. ∀n. uf n = eNorm N1 (f1 n) + eNorm N2 (f2 n) ∧ f = f1 n + f2 n"
using uf(1) by (subst choice_iff[symmetric])+ blast
then obtain f1 f2 where F: "⋀n. uf n = eNorm N1 (f1 n) + eNorm N2 (f2 n)" "⋀n. f = f1 n + f2 n"
by blast
have "N (c *⇩R f) ≤ ¦c¦ * uf n" for n
proof -
have "N (c *⇩R f) ≤ eNorm N1 (c *⇩R f1 n) + eNorm N2 (c *⇩R f2 n)"
unfolding N_def apply (rule Inf_lower, auto, rule exI[of _ "c *⇩R f1 n"], rule exI[of _ "c *⇩R f2 n"])
using F(2)[of n] scaleR_add_right by auto
also have "... = ¦c¦ * (eNorm N1 (f1 n) + eNorm N2 (f2 n))"
by (auto simp add: algebra_simps eNorm_cmult)
finally show ?thesis using F(1) by simp
qed
moreover have "... ⇢ ¦c¦ * N f"
unfolding N_def by (auto intro!: tendsto_intros simp add: uf(2))
ultimately show ?thesis
using LIMSEQ_le_const by blast
qed
have "Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. 0 = f1 + f2} ≤ 0"
by (rule Inf_lower, auto, rule exI[of _ 0], auto)
then have Z: "Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. 0 = f1 + f2} = 0"
by auto
have *: "quasinorm_on UNIV (max (defect N1) (defect N2)) (λf. Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2})"
apply (rule quasinorm_onI) using T H Z defect_ge_1[of N1] defect_ge_1[of N2] unfolding N_def by auto
show "defect (N1 +⇩N N2) = max (defect N1) (defect N2)"
"eNorm (N1 +⇩N N2) f = Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}"
unfolding quasinorm_sum_def using quasinorm_of[OF *] by auto
qed
lemma quasinorm_sum_limit:
"∃f1 f2. (∀n. f = f1 n + f2 n) ∧ (λn. eNorm N1 (f1 n) + eNorm N2 (f2 n)) ⇢ eNorm (N1 +⇩N N2) f"
proof -
have "∃u. (∀n. u n ∈ {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}) ∧ u ⇢ Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}"
by (rule Inf_as_limit, auto, rule exI[of _ "f"], rule exI[of _ 0], auto)
then obtain uf where uf: "⋀n. uf n ∈ {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}"
"uf ⇢ Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}"
by blast
have "∃f1 f2. ∀n. uf n = eNorm N1 (f1 n) + eNorm N2 (f2 n) ∧ f = f1 n + f2 n"
using uf(1) by (subst choice_iff[symmetric])+ blast
then obtain f1 f2 where F: "⋀n. uf n = eNorm N1 (f1 n) + eNorm N2 (f2 n)" "⋀n. f = f1 n + f2 n"
by blast
have "(λn. eNorm N1 (f1 n) + eNorm N2 (f2 n)) ⇢ eNorm (N1 +⇩N N2) f"
using F(1) uf(2) unfolding quasinorm_sum(1) by presburger
then show ?thesis using F(2) by auto
qed
lemma quasinorm_sum_space:
"space⇩N (N1 +⇩N N2) = {f + g|f g. f ∈ space⇩N N1 ∧ g ∈ space⇩N N2}"
proof (auto)
fix x assume "x ∈ space⇩N (N1 +⇩N N2)"
then have "Inf {eNorm N1 f + eNorm N2 g| f g. x = f + g} < ∞"
unfolding quasinorm_sum(1) spaceN_iff.
then have "∃z ∈ {eNorm N1 f + eNorm N2 g| f g. x = f + g}. z < ∞"
by (simp add: Inf_less_iff)
then show "∃f g. x = f + g ∧ f ∈ space⇩N N1 ∧ g ∈ space⇩N N2"
using spaceN_iff by force
next
fix f g assume H: "f ∈ space⇩N N1" "g ∈ space⇩N N2"
have "Inf {eNorm N1 u + eNorm N2 v| u v. f + g = u + v} ≤ eNorm N1 f + eNorm N2 g"
by (rule Inf_lower, auto)
also have "... < ∞" using spaceN_iff H by auto
finally show "f + g ∈ space⇩N (N1 +⇩N N2)"
unfolding spaceN_iff quasinorm_sum(1).
qed
lemma quasinorm_sum_zerospace:
"{f + g |f g. f ∈ zero_space⇩N N1 ∧ g ∈ zero_space⇩N N2} ⊆ zero_space⇩N (N1 +⇩N N2)"
proof (auto, unfold zero_spaceN_iff)
fix f g assume H: "eNorm N1 f = 0" "eNorm N2 g = 0"
have "Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f + g = f1 + f2} ≤ 0"
by (rule Inf_lower, auto, rule exI[of _ f], auto simp add: H)
then show "eNorm (N1 +⇩N N2) (f + g) = 0" unfolding quasinorm_sum(1) by auto
qed
lemma quasinorm_sum_subset:
"N1 ⊆⇩N N1 +⇩N N2" "N2 ⊆⇩N N1 +⇩N N2"
by (rule quasinorm_subsetI[of _ _ 1], auto simp add: quasinorm_sum(1), rule Inf_lower, auto,
metis add.commute add.left_neutral eNorm_zero)+
lemma quasinorm_sum_maximum:
assumes "N1 ⊆⇩N N" "N2 ⊆⇩N N"
shows "N1 +⇩N N2 ⊆⇩N N"
proof -
obtain C1 C2::real where *: "⋀f. eNorm N f ≤ C1 * eNorm N1 f"
"⋀f. eNorm N f ≤ C2 * eNorm N2 f"
"C1 ≥ 0" "C2 ≥ 0"
using quasinorm_subsetD[OF assms(1)] quasinorm_subsetD[OF assms(2)] by blast
have **: "eNorm N f ≤ (defect N * max C1 C2) * eNorm (N1 +⇩N N2) f" for f
proof -
obtain f1 f2 where F: "⋀n. f = f1 n + f2 n"
"(λn. eNorm N1 (f1 n) + eNorm N2 (f2 n)) ⇢ eNorm (N1 +⇩N N2) f"
using quasinorm_sum_limit by blast
have "eNorm N f ≤ ennreal (defect N * max C1 C2) * (eNorm N1 (f1 n) + eNorm N2 (f2 n))" for n
proof -
have "eNorm N f ≤ ennreal(defect N) * eNorm N (f1 n) + ennreal(defect N) * eNorm N (f2 n)"
unfolding ‹f = f1 n + f2 n› using eNorm_triangular_ineq by auto
also have "... ≤ ennreal(defect N) * (C1 * eNorm N1 (f1 n)) + ennreal(defect N) * (C2 * eNorm N2 (f2 n))"
apply (rule add_mono) by (rule mult_mono, simp, simp add: *, simp, simp)+
also have "... ≤ ennreal(defect N) * (max C1 C2 * eNorm N1 (f1 n)) + ennreal(defect N) * (max C1 C2 * eNorm N2 (f2 n))"
by (auto intro!:add_mono mult_mono ennreal_leI)
also have "... = ennreal (defect N * max C1 C2) * (eNorm N1 (f1 n) + eNorm N2 (f2 n))"
apply (subst ennreal_mult') using defect_ge_1 order_trans zero_le_one apply blast
by (auto simp add: algebra_simps)
finally show ?thesis by simp
qed
moreover have "... ⇢ (defect N * max C1 C2) * eNorm (N1 +⇩N N2) f"
by (auto intro!:tendsto_intros F(2))
ultimately show ?thesis
using LIMSEQ_le_const by blast
qed
then show ?thesis
using quasinorm_subsetI by force
qed
lemma quasinorm_sum_assoc:
"(N1 +⇩N N2) +⇩N N3 =⇩N N1 +⇩N (N2 +⇩N N3)"
unfolding quasinorm_equivalent_def by (meson order_trans quasinorm_sum_maximum quasinorm_sum_subset)
subsection ‹Topology›
definition topology⇩N::"('a::real_vector) quasinorm ⇒ 'a topology"
where "topology⇩N N = topology (λU. ∀x∈U. ∃e>0. ∀y. eNorm N (y-x) < e ⟶ y ∈ U)"
lemma istopology_topology⇩N:
"istopology (λU. ∀x∈U. ∃e>0. ∀y. eNorm N (y-x) < e ⟶ y ∈ U)"
unfolding istopology_def by (auto, metis dual_order.strict_trans less_linear, meson)
lemma openin_topology⇩N:
"openin (topology⇩N N) U ⟷ (∀x∈U. ∃e>0. ∀y. eNorm N (y-x) < e ⟶ y ∈ U)"
unfolding topology⇩N_def using istopology_topology⇩N[of N] by (simp add: topology_inverse')
lemma openin_topology⇩N_I:
assumes "⋀x. x ∈ U ⟹ ∃e>0. ∀y. eNorm N (y-x) < e ⟶ y ∈ U"
shows "openin (topology⇩N N) U"
using assms unfolding openin_topology⇩N by auto
lemma openin_topology⇩N_D:
assumes "openin (topology⇩N N) U"
"x ∈ U"
shows "∃e>0. ∀y. eNorm N (y-x) < e ⟶ y ∈ U"
using assms unfolding openin_topology⇩N by auto
text ‹One should then use this topology to define limits and so on. This is not something
specific to quasinorms, but to all topologies defined in this way, not using type classes.
However, there is no such body of material (yet?) in Isabelle-HOL, where topology is
essentially done with type classes. So, we do not go any further for now.
One exception is the notion of completeness, as it is so important in functional analysis.
We give a naive definition, which will be sufficient for the proof of completeness
of several spaces. Usually, the most convenient criterion to prove completeness of
a normed vector space is in terms of converging series. This criterion
is the only nontrivial thing we prove here. We will apply it to prove the
completeness of $L^p$ spaces.›
definition cauchy_ine⇩N::"('a::real_vector) quasinorm ⇒ (nat ⇒ 'a) ⇒ bool"
where "cauchy_ine⇩N N u = (∀e>0. ∃M. ∀n≥M. ∀m≥M. eNorm N (u n - u m) < e)"
definition tendsto_ine⇩N::"('a::real_vector) quasinorm ⇒ (nat ⇒ 'a) ⇒ 'a => bool"
where "tendsto_ine⇩N N u x = (λn. eNorm N (u n - x)) ⇢ 0"
definition complete⇩N::"('a::real_vector) quasinorm ⇒ bool"
where "complete⇩N N = (∀u. cauchy_ine⇩N N u ⟶ (∃x. tendsto_ine⇩N N u x))"
text ‹The above definitions are in terms of eNorms, but usually the nice definitions
only make sense on the space of the norm, and are expressed in terms of Norms. We formulate
the same definitions with norms, they will be more convenient for the proofs.›
definition cauchy_in⇩N::"('a::real_vector) quasinorm ⇒ (nat ⇒ 'a) ⇒ bool"
where "cauchy_in⇩N N u = (∀e>0. ∃M. ∀n≥M. ∀m≥M. Norm N (u n - u m) < e)"
definition tendsto_in⇩N::"('a::real_vector) quasinorm ⇒ (nat ⇒ 'a) ⇒ 'a => bool"
where "tendsto_in⇩N N u x = (λn. Norm N (u n - x)) ⇢ 0"
lemma cauchy_ine⇩N_I:
assumes "⋀e. e > 0 ⟹ (∃M. ∀n≥M. ∀m≥M. eNorm N (u n - u m) < e)"
shows "cauchy_ine⇩N N u"
using assms unfolding cauchy_ine⇩N_def by auto
lemma cauchy_in⇩N_I:
assumes "⋀e. e > 0 ⟹ (∃M. ∀n≥M. ∀m≥M. Norm N (u n - u m) < e)"
shows "cauchy_in⇩N N u"
using assms unfolding cauchy_in⇩N_def by auto
lemma cauchy_ine_in:
assumes "⋀n. u n ∈ space⇩N N"
shows "cauchy_ine⇩N N u ⟷ cauchy_in⇩N N u"
proof
assume "cauchy_in⇩N N u"
show "cauchy_ine⇩N N u"
proof (rule cauchy_ine⇩N_I)
fix e::ennreal assume "e > 0"
define e2 where "e2 = min e 1"
then obtain r where "e2 = ennreal r" "r > 0" unfolding e2_def using ‹e > 0›
by (metis ennreal_eq_1 ennreal_less_zero_iff le_ennreal_iff le_numeral_extra(1) min_def zero_less_one)
then obtain M where *: "∀n≥M. ∀m≥M. Norm N (u n - u m) < r"
using ‹cauchy_in⇩N N u› ‹r > 0› unfolding cauchy_in⇩N_def by auto
then have "∀n≥M. ∀m≥M. eNorm N (u n - u m) < r"
by (auto simp add: assms eNorm_Norm ‹0 < r› ennreal_lessI)
then have "∀n≥M. ∀m≥M. eNorm N (u n - u m) < e"
unfolding ‹e2 = ennreal r›[symmetric] e2_def by auto
then show "∃M. ∀n≥M. ∀m≥M. eNorm N (u n - u m) < e"
by auto
qed
next
assume "cauchy_ine⇩N N u"
show "cauchy_in⇩N N u"
proof (rule cauchy_in⇩N_I)
fix e::real assume "e > 0"
then obtain M where *: "∀n≥M. ∀m≥M. eNorm N (u n - u m) < e"
using ‹cauchy_ine⇩N N u› ‹e > 0› ennreal_less_zero_iff unfolding cauchy_ine⇩N_def by blast
then have "∀n≥M. ∀m≥M. Norm N (u n - u m) < e"
by (auto, metis Norm_def ‹0 < e› eNorm_Norm eNorm_Norm' enn2real_nonneg ennreal_less_iff)
then show "∃M. ∀n≥M. ∀m≥M. Norm N (u n - u m) < e"
by auto
qed
qed
lemma tendsto_ine_in:
assumes "⋀n. u n ∈ space⇩N N" "x ∈ space⇩N N"
shows "tendsto_ine⇩N N u x ⟷ tendsto_in⇩N N u x"
proof -
have *: "eNorm N (u n - x) = Norm N (u n - x)" for n
using assms eNorm_Norm spaceN_diff by blast
show ?thesis unfolding tendsto_in⇩N_def tendsto_ine⇩N_def *
apply (auto)
apply (metis (full_types) Norm_nonneg ennreal_0 eventually_sequentiallyI order_refl tendsto_ennreal_iff)
using tendsto_ennrealI by fastforce
qed
lemma complete⇩N_I:
assumes "⋀u. cauchy_in⇩N N u ⟹ (∀n. u n ∈ space⇩N N) ⟹ (∃x∈ space⇩N N. tendsto_in⇩N N u x)"
shows "complete⇩N N"
proof -
have "∃x. tendsto_ine⇩N N u x" if "cauchy_ine⇩N N u" for u
proof -
obtain M::nat where *: "⋀n m. n ≥ M ⟹ m ≥ M ⟹ eNorm N (u n - u m) < 1"
using ‹cauchy_ine⇩N N u› ennreal_zero_less_one unfolding cauchy_ine⇩N_def by presburger
define v where "v = (λn. u (n+M) - u M)"
have "eNorm N (v n) < 1" for n unfolding v_def using * by auto
then have "v n ∈ space⇩N N" for n using spaceN_iff[of _ N]
by (metis dual_order.strict_trans ennreal_1 ennreal_less_top infinity_ennreal_def)
have "cauchy_ine⇩N N v"
proof (rule cauchy_ine⇩N_I)
fix e::ennreal assume "e > 0"
then obtain P::nat where *: "⋀n m. n ≥ P ⟹ m ≥ P ⟹ eNorm N (u n - u m) < e"
using ‹cauchy_ine⇩N N u› unfolding cauchy_ine⇩N_def by presburger
have "eNorm N (v n - v m) < e" if "n ≥ P" "m ≥ P" for m n
unfolding v_def by (auto, rule *, insert that, auto)
then show "∃M. ∀n≥M. ∀m≥M. eNorm N (v n - v m) < e" by auto
qed
then have "cauchy_in⇩N N v" using cauchy_ine_in[OF ‹⋀n. v n ∈ space⇩N N›] by auto
then obtain y where "tendsto_in⇩N N v y" "y ∈ space⇩N N"
using assms ‹⋀n. v n ∈ space⇩N N› by auto
then have *: "tendsto_ine⇩N N v y"
using tendsto_ine_in ‹⋀n. v n ∈ space⇩N N› by auto
have "tendsto_ine⇩N N u (y + u M)"
unfolding tendsto_ine⇩N_def apply (rule LIMSEQ_offset[of _ M])
using * unfolding v_def tendsto_ine⇩N_def by (auto simp add: algebra_simps)
then show ?thesis by auto
qed
then show ?thesis unfolding complete⇩N_def by auto
qed
lemma cauchy_tendsto_in_subseq:
assumes "⋀n. u n ∈ space⇩N N"
"cauchy_in⇩N N u"
"strict_mono r"
"tendsto_in⇩N N (u o r) x"
shows "tendsto_in⇩N N u x"
proof -
have "∃M. ∀n≥M. Norm N (u n - x) < e" if "e > 0" for e
proof -
define f where "f = e / (2 * defect N)"
have "f > 0" unfolding f_def using ‹e > 0› defect_ge_1[of N] by (auto simp add: divide_simps)
obtain M1 where M1: "⋀m n. m ≥ M1 ⟹ n ≥ M1 ⟹ Norm N (u n - u m) < f"
using ‹cauchy_in⇩N N u› unfolding cauchy_in⇩N_def using ‹f > 0› by meson
obtain M2 where M2: "⋀n. n ≥ M2 ⟹ Norm N ((u o r) n - x) < f"
using ‹tendsto_in⇩N N (u o r) x› ‹f > 0› unfolding tendsto_in⇩N_def order_tendsto_iff eventually_sequentially by blast
define M where "M = max M1 M2"
have "Norm N (u n - x) < e" if "n ≥ M" for n
proof -
have "Norm N (u n - x) = Norm N ((u n - u (r M)) + (u (r M) - x))" by auto
also have "... ≤ defect N * Norm N (u n - u (r M)) + defect N * Norm N (u (r M) - x)"
apply (rule Norm_triangular_ineq) using ‹⋀n. u n ∈ space⇩N N› by simp
also have "... < defect N * f + defect N * f"
apply (auto intro!: add_strict_mono mult_mono simp only:)
using defect_ge_1[of N] ‹n ≥ M› seq_suble[OF ‹strict_mono r›, of M] M1 M2 o_def unfolding M_def by auto
finally show ?thesis
unfolding f_def using ‹e > 0› defect_ge_1[of N] by (auto simp add: divide_simps)
qed
then show ?thesis by auto
qed
then show ?thesis
unfolding tendsto_in⇩N_def order_tendsto_iff eventually_sequentially using Norm_nonneg less_le_trans by blast
qed
proposition complete⇩N_I':
assumes "⋀n. c n > 0"
"⋀u. (∀n. u n ∈ space⇩N N) ⟹ (∀n. Norm N (u n) ≤ c n) ⟹ ∃x∈ space⇩N N. tendsto_in⇩N N (λn. (∑i∈{0..<n}. u i)) x"
shows "complete⇩N N"
proof (rule complete⇩N_I)
fix v assume "cauchy_in⇩N N v" "∀n. v n ∈ space⇩N N"
have *: "∃y. (∀m≥y. ∀p≥y. Norm N (v m - v p) < c (Suc n)) ∧ x < y" if "∀m≥x. ∀p≥x. Norm N (v m - v p) < c n" for x n
proof -
obtain M where i: "∀m≥M. ∀p≥M. Norm N (v m - v p) < c (Suc n)"
using ‹cauchy_in⇩N N v› ‹c (Suc n) > 0› unfolding cauchy_in⇩N_def by (meson zero_less_power)
then show ?thesis
apply (intro exI[of _ "max M (x+1)"]) by auto
qed
have "∃r. ∀n. (∀m≥r n. ∀p≥r n. Norm N (v m - v p) < c n) ∧ r n < r (Suc n)"
apply (intro dependent_nat_choice) using ‹cauchy_in⇩N N v› ‹⋀n. c n > 0› * unfolding cauchy_in⇩N_def by auto
then obtain r where r: "strict_mono r" "⋀n. ∀m≥r n. ∀p≥r n. Norm N (v m - v p) < c n"
by (auto simp: strict_mono_Suc_iff)
define u where "u = (λn. v (r (Suc n)) - v (r n))"
have "u n ∈ space⇩N N" for n
unfolding u_def using ‹∀n. v n ∈ space⇩N N› by simp
moreover have "Norm N (u n) ≤ c n" for n
unfolding u_def using r by (simp add: less_imp_le strict_mono_def)
ultimately obtain y where y: "y ∈ space⇩N N" "tendsto_in⇩N N (λn. (∑i∈{0..<n}. u i)) y"
using assms(2) by blast
define x where "x = y + v (r 0)"
have "x ∈ space⇩N N"
unfolding x_def using ‹y ∈ space⇩N N› ‹∀n. v n ∈ space⇩N N› by simp
have "Norm N (v (r n) - x) = Norm N ((∑i∈{0..<n}. u i) - y)" for n
proof -
have "v (r n) = (∑i∈{0..<n}. u i) + v (r 0)" for n
unfolding u_def by (induct n, auto)
then show ?thesis unfolding x_def by (metis add_diff_cancel_right)
qed
then have "(λn. Norm N (v (r n) - x)) ⇢ 0"
using y(2) unfolding tendsto_in⇩N_def by auto
then have "tendsto_in⇩N N (v o r) x"
unfolding tendsto_in⇩N_def comp_def by force
then have "tendsto_in⇩N N v x"
using ‹∀n. v n ∈ space⇩N N›
by (intro cauchy_tendsto_in_subseq[OF _ ‹cauchy_in⇩N N v› ‹strict_mono r›], auto)
then show "∃x∈space⇩N N. tendsto_in⇩N N v x"
using ‹x ∈ space⇩N N› by blast
qed
text ‹Next, we show when the two examples of norms we have introduced before, the ambient norm
in a Banach space, and the norm on bounded continuous functions, are complete. We just have to
translate in our setting the already known completeness of these spaces.›
lemma complete_N_of_norm:
"complete⇩N (N_of_norm::'a::banach quasinorm)"
proof (rule complete⇩N_I)
fix u::"nat ⇒ 'a" assume "cauchy_in⇩N N_of_norm u"
then have "Cauchy u" unfolding Cauchy_def cauchy_in⇩N_def N_of_norm(2) by (simp add: dist_norm)
then obtain x where "u ⇢ x" using convergent_eq_Cauchy by blast
then have "tendsto_in⇩N N_of_norm u x" unfolding tendsto_in⇩N_def N_of_norm(2)
using Lim_null tendsto_norm_zero_iff by fastforce
moreover have "x ∈ space⇩N N_of_norm" by auto
ultimately show "∃x∈space⇩N N_of_norm. tendsto_in⇩N N_of_norm u x" by auto
qed
text ‹In the next statement, the assumption that \verb+'a+ is a metric space is not necessary,
a topological space would be enough, but a statement about uniform convergence is not available
in this setting.
TODO: fix it.
›
lemma complete_bcontfunN:
"complete⇩N (bcontfun⇩N::('a::metric_space ⇒ 'b::banach) quasinorm)"
proof (rule complete⇩N_I)
fix u::"nat ⇒ ('a ⇒ 'b)" assume H: "cauchy_in⇩N bcontfun⇩N u" "∀n. u n ∈ space⇩N bcontfun⇩N"
then have H2: "u n ∈ bcontfun" for n using bcontfun⇩N_space by auto
then have **: "Bcontfun(u n - u m) = Bcontfun (u n) - Bcontfun (u m)" for m n
unfolding minus_fun_def minus_bcontfun_def by (simp add: Bcontfun_inverse)
have *: "Norm bcontfun⇩N (u n - u m) = norm (Bcontfun (u n - u m))" for n m
unfolding bcontfun⇩N(2) using H(2) bcontfun⇩N_space by auto
have "Cauchy (λn. Bcontfun (u n))"
using H(1) unfolding Cauchy_def cauchy_in⇩N_def dist_norm * ** by simp
then obtain v where v: "(λn. Bcontfun (u n)) ⇢ v"
using convergent_eq_Cauchy by blast
have v_space: "apply_bcontfun v ∈ space⇩N bcontfun⇩N" unfolding bcontfun⇩N_space by (simp add: apply_bcontfun)
have ***: "Norm bcontfun⇩N (u n - v) = norm(Bcontfun (u n) - v)" for n
proof -
have "Norm bcontfun⇩N (u n - v) = norm (Bcontfun(u n - v))"
unfolding bcontfun⇩N(2) using H(2) bcontfun⇩N_space v_space by auto
moreover have "Bcontfun(u n - v) = Bcontfun (u n) - v"
unfolding minus_fun_def minus_bcontfun_def by (simp add: Bcontfun_inverse H2)
ultimately show ?thesis by simp
qed
have "tendsto_in⇩N bcontfun⇩N u v"
unfolding tendsto_in⇩N_def *** using v Lim_null tendsto_norm_zero_iff by fastforce
then show "∃v∈space⇩N bcontfun⇩N. tendsto_in⇩N bcontfun⇩N u v" using v_space by auto
qed
end