Theory Complex_Bounded_Linear_Function0
section ‹‹Complex_Bounded_Linear_Function0› -- Bounded Linear Function›
theory Complex_Bounded_Linear_Function0
imports
"HOL-Analysis.Bounded_Linear_Function"
Complex_Inner_Product
Complex_Euclidean_Space0
begin
unbundle cinner_syntax
lemma conorm_componentwise:
assumes "bounded_clinear f"
shows "onorm f ≤ (∑i∈CBasis. norm (f i))"
proof -
{
fix i::'a
assume "i ∈ CBasis"
hence "onorm (λx. (i ∙⇩C x) *⇩C f i) ≤ onorm (λx. (i ∙⇩C x)) * norm (f i)"
by (auto intro!: onorm_scaleC_left_lemma bounded_clinear_cinner_right)
also have "… ≤ norm i * norm (f i)"
apply (rule mult_right_mono)
apply (simp add: complex_inner_class.Cauchy_Schwarz_ineq2 onorm_bound)
by simp
finally have "onorm (λx. (i ∙⇩C x) *⇩C f i) ≤ norm (f i)" using ‹i ∈ CBasis›
by simp
} hence "onorm (λx. ∑i∈CBasis. (i ∙⇩C x) *⇩C f i) ≤ (∑i∈CBasis. norm (f i))"
by (auto intro!: order_trans[OF onorm_sum_le] bounded_clinear_scaleC_const
sum_mono bounded_clinear_cinner_right bounded_clinear.bounded_linear)
also have "(λx. ∑i∈CBasis. (i ∙⇩C x) *⇩C f i) = (λx. f (∑i∈CBasis. (i ∙⇩C x) *⇩C i))"
by (simp add: clinear.scaleC linear_sum bounded_clinear.clinear clinear.linear assms)
also have "… = f"
by (simp add: ceuclidean_representation)
finally show ?thesis .
qed
lemmas conorm_componentwise_le = order_trans[OF conorm_componentwise]
subsection ‹Intro rules for \<^term>‹bounded_linear››
lemma onorm_cinner_left:
assumes "bounded_linear r"
shows "onorm (λx. r x ∙⇩C f) ≤ onorm r * norm f"
proof (rule onorm_bound)
fix x
have "norm (r x ∙⇩C f) ≤ norm (r x) * norm f"
by (simp add: Cauchy_Schwarz_ineq2)
also have "… ≤ onorm r * norm x * norm f"
by (simp add: assms mult.commute mult_left_mono onorm)
finally show "norm (r x ∙⇩C f) ≤ onorm r * norm f * norm x"
by (simp add: ac_simps)
qed (intro mult_nonneg_nonneg norm_ge_zero onorm_pos_le assms)
lemma onorm_cinner_right:
assumes "bounded_linear r"
shows "onorm (λx. f ∙⇩C r x) ≤ norm f * onorm r"
proof (rule onorm_bound)
fix x
have "norm (f ∙⇩C r x) ≤ norm f * norm (r x)"
by (simp add: Cauchy_Schwarz_ineq2)
also have "… ≤ onorm r * norm x * norm f"
by (simp add: assms mult.commute mult_left_mono onorm)
finally show "norm (f ∙⇩C r x) ≤ norm f * onorm r * norm x"
by (simp add: ac_simps)
qed (intro mult_nonneg_nonneg norm_ge_zero onorm_pos_le assms)
lemmas [bounded_linear_intros] =
bounded_clinear_zero
bounded_clinear_add
bounded_clinear_const_mult
bounded_clinear_mult_const
bounded_clinear_scaleC_const
bounded_clinear_const_scaleC
bounded_clinear_const_scaleR
bounded_clinear_ident
bounded_clinear_sum
bounded_clinear_sub
bounded_antilinear_cinner_left_comp
bounded_clinear_cinner_right_comp
subsection ‹declaration of derivative/continuous/tendsto introduction rules for bounded linear functions›
attribute_setup bounded_clinear =
‹let val bounded_linear = Attrib.attribute \<^context> (the_single @{attributes [bounded_linear]}) in
Scan.succeed (Thm.declaration_attribute (fn thm =>
Thm.attribute_declaration bounded_linear (thm RS @{thm bounded_clinear.bounded_linear}) o
fold (fn (r, s) => Named_Theorems.add_thm s (thm RS r))
[
(@{thm bounded_clinear_compose}, \<^named_theorems>‹bounded_linear_intros›),
(@{thm bounded_clinear_o_bounded_antilinear[unfolded o_def]}, \<^named_theorems>‹bounded_linear_intros›)
]))
end›
attribute_setup bounded_antilinear =
‹let val bounded_linear = Attrib.attribute \<^context> (the_single @{attributes [bounded_linear]}) in
Scan.succeed (Thm.declaration_attribute (fn thm =>
Thm.attribute_declaration bounded_linear (thm RS @{thm bounded_antilinear.bounded_linear}) o
fold (fn (r, s) => Named_Theorems.add_thm s (thm RS r))
[
(@{thm bounded_antilinear_o_bounded_clinear[unfolded o_def]}, \<^named_theorems>‹bounded_linear_intros›),
(@{thm bounded_antilinear_o_bounded_antilinear[unfolded o_def]}, \<^named_theorems>‹bounded_linear_intros›)
]))
end›
attribute_setup bounded_cbilinear =
‹let val bounded_bilinear = Attrib.attribute \<^context> (the_single @{attributes [bounded_bilinear]}) in
Scan.succeed (Thm.declaration_attribute (fn thm =>
Thm.attribute_declaration bounded_bilinear (thm RS @{thm bounded_cbilinear.bounded_bilinear}) o
fold (fn (r, s) => Named_Theorems.add_thm s (thm RS r))
[
(@{thm bounded_clinear_compose[OF bounded_cbilinear.bounded_clinear_left]},
\<^named_theorems>‹bounded_linear_intros›),
(@{thm bounded_clinear_compose[OF bounded_cbilinear.bounded_clinear_right]},
\<^named_theorems>‹bounded_linear_intros›),
(@{thm bounded_clinear_o_bounded_antilinear[unfolded o_def, OF bounded_cbilinear.bounded_clinear_left]},
\<^named_theorems>‹bounded_linear_intros›),
(@{thm bounded_clinear_o_bounded_antilinear[unfolded o_def, OF bounded_cbilinear.bounded_clinear_right]},
\<^named_theorems>‹bounded_linear_intros›)
]))
end›
attribute_setup bounded_sesquilinear =
‹let val bounded_bilinear = Attrib.attribute \<^context> (the_single @{attributes [bounded_bilinear]}) in
Scan.succeed (Thm.declaration_attribute (fn thm =>
Thm.attribute_declaration bounded_bilinear (thm RS @{thm bounded_sesquilinear.bounded_bilinear}) o
fold (fn (r, s) => Named_Theorems.add_thm s (thm RS r))
[
(@{thm bounded_antilinear_o_bounded_clinear[unfolded o_def, OF bounded_sesquilinear.bounded_antilinear_left]},
\<^named_theorems>‹bounded_linear_intros›),
(@{thm bounded_clinear_compose[OF bounded_sesquilinear.bounded_clinear_right]},
\<^named_theorems>‹bounded_linear_intros›),
(@{thm bounded_antilinear_o_bounded_antilinear[unfolded o_def, OF bounded_sesquilinear.bounded_antilinear_left]},
\<^named_theorems>‹bounded_linear_intros›),
(@{thm bounded_clinear_o_bounded_antilinear[unfolded o_def, OF bounded_sesquilinear.bounded_clinear_right]},
\<^named_theorems>‹bounded_linear_intros›)
]))
end›
subsection ‹Type of complex bounded linear functions›
typedef (overloaded) ('a, 'b) cblinfun ("(_ ⇒⇩C⇩L /_)" [22, 21] 21) =
"{f::'a::complex_normed_vector⇒'b::complex_normed_vector. bounded_clinear f}"
morphisms cblinfun_apply CBlinfun
by (blast intro: bounded_linear_intros)
declare [[coercion
"cblinfun_apply :: ('a::complex_normed_vector ⇒⇩C⇩L'b::complex_normed_vector) ⇒ 'a ⇒ 'b"]]
lemma bounded_clinear_cblinfun_apply[bounded_linear_intros]:
"bounded_clinear g ⟹ bounded_clinear (λx. cblinfun_apply f (g x))"
by (metis cblinfun_apply mem_Collect_eq bounded_clinear_compose)
setup_lifting type_definition_cblinfun
lemma cblinfun_eqI: "(⋀i. cblinfun_apply x i = cblinfun_apply y i) ⟹ x = y"
by transfer auto
lemma bounded_clinear_CBlinfun_apply: "bounded_clinear f ⟹ cblinfun_apply (CBlinfun f) = f"
by (auto simp: CBlinfun_inverse)
subsection ‹Type class instantiations›
instantiation cblinfun :: (complex_normed_vector, complex_normed_vector) complex_normed_vector
begin
lift_definition norm_cblinfun :: "'a ⇒⇩C⇩L 'b ⇒ real" is onorm .
lift_definition minus_cblinfun :: "'a ⇒⇩C⇩L 'b ⇒ 'a ⇒⇩C⇩L 'b ⇒ 'a ⇒⇩C⇩L 'b"
is "λf g x. f x - g x"
by (rule bounded_clinear_sub)
definition dist_cblinfun :: "'a ⇒⇩C⇩L 'b ⇒ 'a ⇒⇩C⇩L 'b ⇒ real"
where "dist_cblinfun a b = norm (a - b)"
definition [code del]:
"(uniformity :: (('a ⇒⇩C⇩L 'b) × ('a ⇒⇩C⇩L 'b)) filter) = (INF e∈{0 <..}. principal {(x, y). dist x y < e})"
definition open_cblinfun :: "('a ⇒⇩C⇩L 'b) set ⇒ bool"
where [code del]: "open_cblinfun S = (∀x∈S. ∀⇩F (x', y) in uniformity. x' = x ⟶ y ∈ S)"
lift_definition uminus_cblinfun :: "'a ⇒⇩C⇩L 'b ⇒ 'a ⇒⇩C⇩L 'b" is "λf x. - f x"
by (rule bounded_clinear_minus)
lift_definition zero_cblinfun :: "'a ⇒⇩C⇩L 'b" is "λx. 0"
by (rule bounded_clinear_zero)
lift_definition plus_cblinfun :: "'a ⇒⇩C⇩L 'b ⇒ 'a ⇒⇩C⇩L 'b ⇒ 'a ⇒⇩C⇩L 'b"
is "λf g x. f x + g x"
by (metis bounded_clinear_add)
lift_definition scaleC_cblinfun::"complex ⇒ 'a ⇒⇩C⇩L 'b ⇒ 'a ⇒⇩C⇩L 'b" is "λr f x. r *⇩C f x"
by (metis bounded_clinear_compose bounded_clinear_scaleC_right)
lift_definition scaleR_cblinfun::"real ⇒ 'a ⇒⇩C⇩L 'b ⇒ 'a ⇒⇩C⇩L 'b" is "λr f x. r *⇩R f x"
by (rule bounded_clinear_const_scaleR)
definition sgn_cblinfun :: "'a ⇒⇩C⇩L 'b ⇒ 'a ⇒⇩C⇩L 'b"
where "sgn_cblinfun x = scaleC (inverse (norm x)) x"
instance
proof
fix a b c :: "'a ⇒⇩C⇩L'b" and r q :: real and s t :: complex
show ‹a + b + c = a + (b + c)›
apply transfer by auto
show ‹0 + a = a›
apply transfer by auto
show ‹a + b = b + a›
apply transfer by auto
show ‹- a + a = 0›
apply transfer by auto
show ‹a - b = a + - b›
apply transfer by auto
show scaleR_scaleC: ‹((*⇩R) r::('a ⇒⇩C⇩L 'b) ⇒ _) = (*⇩C) (complex_of_real r)› for r
apply (rule ext, transfer fixing: r) by (simp add: scaleR_scaleC)
show ‹s *⇩C (b + c) = s *⇩C b + s *⇩C c›
apply transfer by (simp add: scaleC_add_right)
show ‹(s + t) *⇩C a = s *⇩C a + t *⇩C a›
apply transfer by (simp add: scaleC_left.add)
show ‹s *⇩C t *⇩C a = (s * t) *⇩C a›
apply transfer by auto
show ‹1 *⇩C a = a›
apply transfer by auto
show ‹dist a b = norm (a - b)›
unfolding dist_cblinfun_def by simp
show ‹sgn a = (inverse (norm a)) *⇩R a›
unfolding sgn_cblinfun_def unfolding scaleR_scaleC by auto
show ‹uniformity = (INF e∈{0<..}. principal {(x, y). dist (x::('a ⇒⇩C⇩L 'b)) y < e})›
by (simp add: uniformity_cblinfun_def)
show ‹open U = (∀x∈U. ∀⇩F (x', y) in uniformity. (x'::('a ⇒⇩C⇩L 'b)) = x ⟶ y ∈ U)› for U
by (simp add: open_cblinfun_def)
show ‹(norm a = 0) = (a = 0)›
apply transfer using bounded_clinear.bounded_linear onorm_eq_0 by blast
show ‹norm (a + b) ≤ norm a + norm b›
apply transfer by (simp add: bounded_clinear.bounded_linear onorm_triangle)
show ‹norm (s *⇩C a) = cmod s * norm a›
apply transfer using onorm_scalarC by blast
show ‹norm (r *⇩R a) = ¦r¦ * norm a›
apply transfer using bounded_clinear.bounded_linear onorm_scaleR by blast
show ‹r *⇩R (a + b) = r *⇩R a + r *⇩R b›
apply transfer by (simp add: scaleR_add_right)
show ‹(r + q) *⇩R a = r *⇩R a + q *⇩R a›
apply transfer by (simp add: scaleR_add_left)
show ‹r *⇩R q *⇩R a = (r * q) *⇩R a›
apply transfer by auto
show ‹1 *⇩R a = a›
apply transfer by auto
qed
end
declare uniformity_Abort[where 'a="('a :: complex_normed_vector) ⇒⇩C⇩L ('b :: complex_normed_vector)", code]
lemma norm_cblinfun_eqI:
assumes "n ≤ norm (cblinfun_apply f x) / norm x"
assumes "⋀x. norm (cblinfun_apply f x) ≤ n * norm x"
assumes "0 ≤ n"
shows "norm f = n"
by (auto simp: norm_cblinfun_def
intro!: antisym onorm_bound assms order_trans[OF _ le_onorm] bounded_clinear.bounded_linear
bounded_linear_intros)
lemma norm_cblinfun: "norm (cblinfun_apply f x) ≤ norm f * norm x"
apply transfer by (simp add: bounded_clinear.bounded_linear onorm)
lemma norm_cblinfun_bound: "0 ≤ b ⟹ (⋀x. norm (cblinfun_apply f x) ≤ b * norm x) ⟹ norm f ≤ b"
by transfer (rule onorm_bound)
lemma bounded_cbilinear_cblinfun_apply[bounded_cbilinear]: "bounded_cbilinear cblinfun_apply"
proof
fix f g::"'a ⇒⇩C⇩L 'b" and a b::'a and r::complex
show "(f + g) a = f a + g a" "(r *⇩C f) a = r *⇩C f a"
by (transfer, simp)+
interpret bounded_clinear f for f::"'a ⇒⇩C⇩L 'b"
by (auto intro!: bounded_linear_intros)
show "f (a + b) = f a + f b" "f (r *⇩C a) = r *⇩C f a"
by (simp_all add: add scaleC)
show "∃K. ∀a b. norm (cblinfun_apply a b) ≤ norm a * norm b * K"
by (auto intro!: exI[where x=1] norm_cblinfun)
qed
interpretation cblinfun: bounded_cbilinear cblinfun_apply
by (rule bounded_cbilinear_cblinfun_apply)
lemmas bounded_clinear_apply_cblinfun[intro, simp] = cblinfun.bounded_clinear_left
declare cblinfun.zero_left [simp] cblinfun.zero_right [simp]
context bounded_cbilinear
begin
named_theorems cbilinear_simps
lemmas [cbilinear_simps] =
add_left
add_right
diff_left
diff_right
minus_left
minus_right
scaleC_left
scaleC_right
zero_left
zero_right
sum_left
sum_right
end
instance cblinfun :: (complex_normed_vector, cbanach) cbanach
proof
fix X::"nat ⇒ 'a ⇒⇩C⇩L 'b"
assume "Cauchy X"
{
fix x::'a
{
fix x::'a
assume "norm x ≤ 1"
have "Cauchy (λn. X n x)"
proof (rule CauchyI)
fix e::real
assume "0 < e"
from CauchyD[OF ‹Cauchy X› ‹0 < e›] obtain M
where M: "⋀m n. m ≥ M ⟹ n ≥ M ⟹ norm (X m - X n) < e"
by auto
show "∃M. ∀m≥M. ∀n≥M. norm (X m x - X n x) < e"
proof (safe intro!: exI[where x=M])
fix m n
assume le: "M ≤ m" "M ≤ n"
have "norm (X m x - X n x) = norm ((X m - X n) x)"
by (simp add: cblinfun.cbilinear_simps)
also have "… ≤ norm (X m - X n) * norm x"
by (rule norm_cblinfun)
also have "… ≤ norm (X m - X n) * 1"
using ‹norm x ≤ 1› norm_ge_zero by (rule mult_left_mono)
also have "… = norm (X m - X n)" by simp
also have "… < e" using le by fact
finally show "norm (X m x - X n x) < e" .
qed
qed
hence "convergent (λn. X n x)"
by (metis Cauchy_convergent_iff)
} note convergent_norm1 = this
define y where "y = x /⇩R norm x"
have y: "norm y ≤ 1" and xy: "x = norm x *⇩R y"
by (simp_all add: y_def inverse_eq_divide)
have "convergent (λn. norm x *⇩R X n y)"
by (intro bounded_bilinear.convergent[OF bounded_bilinear_scaleR] convergent_const
convergent_norm1 y)
also have "(λn. norm x *⇩R X n y) = (λn. X n x)"
by (metis cblinfun.scaleC_right scaleR_scaleC xy)
finally have "convergent (λn. X n x)" .
}
then obtain v where v: "⋀x. (λn. X n x) ⇢ v x"
unfolding convergent_def
by metis
have "Cauchy (λn. norm (X n))"
proof (rule CauchyI)
fix e::real
assume "e > 0"
from CauchyD[OF ‹Cauchy X› ‹0 < e›] obtain M
where M: "⋀m n. m ≥ M ⟹ n ≥ M ⟹ norm (X m - X n) < e"
by auto
show "∃M. ∀m≥M. ∀n≥M. norm (norm (X m) - norm (X n)) < e"
proof (safe intro!: exI[where x=M])
fix m n assume mn: "m ≥ M" "n ≥ M"
have "norm (norm (X m) - norm (X n)) ≤ norm (X m - X n)"
by (metis norm_triangle_ineq3 real_norm_def)
also have "… < e" using mn by fact
finally show "norm (norm (X m) - norm (X n)) < e" .
qed
qed
then obtain K where K: "(λn. norm (X n)) ⇢ K"
unfolding Cauchy_convergent_iff convergent_def
by metis
have "bounded_clinear v"
proof
fix x y and r::complex
from tendsto_add[OF v[of x] v [of y]] v[of "x + y", unfolded cblinfun.cbilinear_simps]
tendsto_scaleC[OF tendsto_const[of r] v[of x]] v[of "r *⇩C x", unfolded cblinfun.cbilinear_simps]
show "v (x + y) = v x + v y" "v (r *⇩C x) = r *⇩C v x"
by (metis (poly_guards_query) LIMSEQ_unique)+
show "∃K. ∀x. norm (v x) ≤ norm x * K"
proof (safe intro!: exI[where x=K])
fix x
have "norm (v x) ≤ K * norm x"
apply (rule tendsto_le[OF _ tendsto_mult[OF K tendsto_const] tendsto_norm[OF v]])
by (auto simp: norm_cblinfun)
thus "norm (v x) ≤ norm x * K"
by (simp add: ac_simps)
qed
qed
hence Bv: "⋀x. (λn. X n x) ⇢ CBlinfun v x"
by (auto simp: bounded_clinear_CBlinfun_apply v)
have "X ⇢ CBlinfun v"
proof (rule LIMSEQ_I)
fix r::real assume "r > 0"
define r' where "r' = r / 2"
have "0 < r'" "r' < r" using ‹r > 0› by (simp_all add: r'_def)
from CauchyD[OF ‹Cauchy X› ‹r' > 0›]
obtain M where M: "⋀m n. m ≥ M ⟹ n ≥ M ⟹ norm (X m - X n) < r'"
by metis
show "∃no. ∀n≥no. norm (X n - CBlinfun v) < r"
proof (safe intro!: exI[where x=M])
fix n assume n: "M ≤ n"
have "norm (X n - CBlinfun v) ≤ r'"
proof (rule norm_cblinfun_bound)
fix x
have "eventually (λm. m ≥ M) sequentially"
by (metis eventually_ge_at_top)
hence ev_le: "eventually (λm. norm (X n x - X m x) ≤ r' * norm x) sequentially"
proof eventually_elim
case (elim m)
have "norm (X n x - X m x) = norm ((X n - X m) x)"
by (simp add: cblinfun.cbilinear_simps)
also have "… ≤ norm ((X n - X m)) * norm x"
by (rule norm_cblinfun)
also have "… ≤ r' * norm x"
using M[OF n elim] by (simp add: mult_right_mono)
finally show ?case .
qed
have tendsto_v: "(λm. norm (X n x - X m x)) ⇢ norm (X n x - CBlinfun v x)"
by (auto intro!: tendsto_intros Bv)
show "norm ((X n - CBlinfun v) x) ≤ r' * norm x"
by (auto intro!: tendsto_upperbound tendsto_v ev_le simp: cblinfun.cbilinear_simps)
qed (simp add: ‹0 < r'› less_imp_le)
thus "norm (X n - CBlinfun v) < r"
by (metis ‹r' < r› le_less_trans)
qed
qed
thus "convergent X"
by (rule convergentI)
qed
subsection ‹On Euclidean Space›
lemma norm_cblinfun_ceuclidean_le:
fixes a::"'a::ceuclidean_space ⇒⇩C⇩L 'b::complex_normed_vector"
shows "norm a ≤ sum (λx. norm (a x)) CBasis"
apply (rule norm_cblinfun_bound)
apply (simp add: sum_nonneg)
apply (subst ceuclidean_representation[symmetric, where 'a='a])
apply (simp only: cblinfun.cbilinear_simps sum_distrib_right)
apply (rule order.trans[OF norm_sum sum_mono])
apply (simp add: abs_mult mult_right_mono ac_simps CBasis_le_norm)
by (metis complex_inner_class.Cauchy_Schwarz_ineq2 mult.commute mult.left_neutral mult_right_mono norm_CBasis norm_ge_zero)
lemma ctendsto_componentwise1:
fixes a::"'a::ceuclidean_space ⇒⇩C⇩L 'b::complex_normed_vector"
and b::"'c ⇒ 'a ⇒⇩C⇩L 'b"
assumes "(⋀j. j ∈ CBasis ⟹ ((λn. b n j) ⤏ a j) F)"
shows "(b ⤏ a) F"
proof -
have "⋀j. j ∈ CBasis ⟹ Zfun (λx. norm (b x j - a j)) F"
using assms unfolding tendsto_Zfun_iff Zfun_norm_iff .
hence "Zfun (λx. ∑j∈CBasis. norm (b x j - a j)) F"
by (auto intro!: Zfun_sum)
thus ?thesis
unfolding tendsto_Zfun_iff
by (rule Zfun_le)
(auto intro!: order_trans[OF norm_cblinfun_ceuclidean_le] simp: cblinfun.cbilinear_simps)
qed
lift_definition
cblinfun_of_matrix::"('b::ceuclidean_space ⇒ 'a::ceuclidean_space ⇒ complex) ⇒ 'a ⇒⇩C⇩L 'b"
is "λa x. ∑i∈CBasis. ∑j∈CBasis. ((j ∙⇩C x) * a i j) *⇩C i"
by (intro bounded_linear_intros)
lemma cblinfun_of_matrix_works:
fixes f::"'a::ceuclidean_space ⇒⇩C⇩L 'b::ceuclidean_space"
shows "cblinfun_of_matrix (λi j. i ∙⇩C (f j)) = f"
proof (transfer, rule, rule ceuclidean_eqI)
fix f::"'a ⇒ 'b" and x::'a and b::'b assume "bounded_clinear f" and b: "b ∈ CBasis"
then interpret bounded_clinear f by simp
have "(∑j∈CBasis. ∑i∈CBasis. (i ∙⇩C x * (j ∙⇩C f i)) *⇩C j) ∙⇩C b
= (∑j∈CBasis. if j = b then (∑i∈CBasis. (x ∙⇩C i * (f i ∙⇩C j))) else 0)"
using b
apply (simp add: cinner_sum_left cinner_CBasis if_distrib cong: if_cong)
by (simp add: sum.swap)
also have "… = (∑i∈CBasis. ((x ∙⇩C i) * (f i ∙⇩C b)))"
using b by (simp)
also have "… = f x ∙⇩C b"
proof -
have ‹(∑i∈CBasis. (x ∙⇩C i) * (f i ∙⇩C b)) = (∑i∈CBasis. (i ∙⇩C x) *⇩C f i) ∙⇩C b›
by (auto simp: cinner_sum_left)
also have ‹… = f x ∙⇩C b›
by (simp add: ceuclidean_representation sum[symmetric] scale[symmetric])
finally show ?thesis by -
qed
finally show "(∑j∈CBasis. ∑i∈CBasis. (i ∙⇩C x * (j ∙⇩C f i)) *⇩C j) ∙⇩C b = f x ∙⇩C b" .
qed
lemma cblinfun_of_matrix_apply:
"cblinfun_of_matrix a x = (∑i∈CBasis. ∑j∈CBasis. ((j ∙⇩C x) * a i j) *⇩C i)"
apply transfer by simp
lemma cblinfun_of_matrix_minus: "cblinfun_of_matrix x - cblinfun_of_matrix y = cblinfun_of_matrix (x - y)"
by transfer (auto simp: algebra_simps sum_subtractf)
lemma norm_cblinfun_of_matrix:
"norm (cblinfun_of_matrix a) ≤ (∑i∈CBasis. ∑j∈CBasis. cmod (a i j))"
apply (rule norm_cblinfun_bound)
apply (simp add: sum_nonneg)
apply (simp only: cblinfun_of_matrix_apply sum_distrib_right)
apply (rule order_trans[OF norm_sum sum_mono])
apply (rule order_trans[OF norm_sum sum_mono])
apply (simp add: abs_mult mult_right_mono ac_simps Basis_le_norm)
by (metis complex_inner_class.Cauchy_Schwarz_ineq2 complex_scaleC_def mult.left_neutral mult_right_mono norm_CBasis norm_ge_zero norm_scaleC)
lemma tendsto_cblinfun_of_matrix:
assumes "⋀i j. i ∈ CBasis ⟹ j ∈ CBasis ⟹ ((λn. b n i j) ⤏ a i j) F"
shows "((λn. cblinfun_of_matrix (b n)) ⤏ cblinfun_of_matrix a) F"
proof -
have "⋀i j. i ∈ CBasis ⟹ j ∈ CBasis ⟹ Zfun (λx. norm (b x i j - a i j)) F"
using assms unfolding tendsto_Zfun_iff Zfun_norm_iff .
hence "Zfun (λx. (∑i∈CBasis. ∑j∈CBasis. cmod (b x i j - a i j))) F"
by (auto intro!: Zfun_sum)
thus ?thesis
unfolding tendsto_Zfun_iff cblinfun_of_matrix_minus
by (rule Zfun_le) (auto intro!: order_trans[OF norm_cblinfun_of_matrix])
qed
lemma ctendsto_componentwise:
fixes a::"'a::ceuclidean_space ⇒⇩C⇩L 'b::ceuclidean_space"
and b::"'c ⇒ 'a ⇒⇩C⇩L 'b"
shows "(⋀i j. i ∈ CBasis ⟹ j ∈ CBasis ⟹ ((λn. b n j ∙⇩C i) ⤏ a j ∙⇩C i) F) ⟹ (b ⤏ a) F"
apply (subst cblinfun_of_matrix_works[of a, symmetric])
apply (subst cblinfun_of_matrix_works[of "b x" for x, symmetric, abs_def])
apply (rule tendsto_cblinfun_of_matrix)
apply (subst (1) cinner_commute, subst (2) cinner_commute)
by (metis lim_cnj)
lemma
continuous_cblinfun_componentwiseI:
fixes f:: "'b::t2_space ⇒ 'a::ceuclidean_space ⇒⇩C⇩L 'c::ceuclidean_space"
assumes "⋀i j. i ∈ CBasis ⟹ j ∈ CBasis ⟹ continuous F (λx. (f x) j ∙⇩C i)"
shows "continuous F f"
using assms by (auto simp: continuous_def intro!: ctendsto_componentwise)
lemma
continuous_cblinfun_componentwiseI1:
fixes f:: "'b::t2_space ⇒ 'a::ceuclidean_space ⇒⇩C⇩L 'c::complex_normed_vector"
assumes "⋀i. i ∈ CBasis ⟹ continuous F (λx. f x i)"
shows "continuous F f"
using assms by (auto simp: continuous_def intro!: ctendsto_componentwise1)
lemma
continuous_on_cblinfun_componentwise:
fixes f:: "'d::t2_space ⇒ 'e::ceuclidean_space ⇒⇩C⇩L 'f::complex_normed_vector"
assumes "⋀i. i ∈ CBasis ⟹ continuous_on s (λx. f x i)"
shows "continuous_on s f"
using assms
by (auto intro!: continuous_at_imp_continuous_on intro!: ctendsto_componentwise1
simp: continuous_on_eq_continuous_within continuous_def)
lemma bounded_antilinear_cblinfun_matrix: "bounded_antilinear (λx. (x::_⇒⇩C⇩L _) j ∙⇩C i)"
by (auto intro!: bounded_linear_intros)
lemma continuous_cblinfun_matrix:
fixes f:: "'b::t2_space ⇒ 'a::complex_normed_vector ⇒⇩C⇩L 'c::complex_inner"
assumes "continuous F f"
shows "continuous F (λx. (f x) j ∙⇩C i)"
by (rule bounded_antilinear.continuous[OF bounded_antilinear_cblinfun_matrix assms])
lemma continuous_on_cblinfun_matrix:
fixes f::"'a::t2_space ⇒ 'b::complex_normed_vector ⇒⇩C⇩L 'c::complex_inner"
assumes "continuous_on S f"
shows "continuous_on S (λx. (f x) j ∙⇩C i)"
using assms
by (auto simp: continuous_on_eq_continuous_within continuous_cblinfun_matrix)
lemma continuous_on_cblinfun_of_matrix[continuous_intros]:
assumes "⋀i j. i ∈ CBasis ⟹ j ∈ CBasis ⟹ continuous_on S (λs. g s i j)"
shows "continuous_on S (λs. cblinfun_of_matrix (g s))"
using assms
by (auto simp: continuous_on intro!: tendsto_cblinfun_of_matrix)
lemma cblinfun_euclidean_eqI: "(⋀i. i ∈ CBasis ⟹ cblinfun_apply x i = cblinfun_apply y i) ⟹ x = y"
apply (auto intro!: cblinfun_eqI)
apply (subst (2) ceuclidean_representation[symmetric, where 'a='a])
apply (subst (1) ceuclidean_representation[symmetric, where 'a='a])
by (simp add: cblinfun.cbilinear_simps)
lemma CBlinfun_eq_matrix: "bounded_clinear f ⟹ CBlinfun f = cblinfun_of_matrix (λi j. i ∙⇩C f j)"
apply (intro cblinfun_euclidean_eqI)
by (auto simp: cblinfun_of_matrix_apply bounded_clinear_CBlinfun_apply cinner_CBasis if_distrib
if_distribR sum.delta' ceuclidean_representation
cong: if_cong)
subsection ‹concrete bounded linear functions›
lemma transfer_bounded_cbilinear_bounded_clinearI:
assumes "g = (λi x. (cblinfun_apply (f i) x))"
shows "bounded_cbilinear g = bounded_clinear f"
proof
assume "bounded_cbilinear g"
then interpret bounded_cbilinear f by (simp add: assms)
show "bounded_clinear f"
proof (unfold_locales, safe intro!: cblinfun_eqI)
fix i
show "f (x + y) i = (f x + f y) i" "f (r *⇩C x) i = (r *⇩C f x) i" for r x y
by (auto intro!: cblinfun_eqI simp: cblinfun.cbilinear_simps)
from _ nonneg_bounded show "∃K. ∀x. norm (f x) ≤ norm x * K"
by (rule ex_reg) (auto intro!: onorm_bound simp: norm_cblinfun.rep_eq ac_simps)
qed
qed (auto simp: assms intro!: cblinfun.comp)
lemma transfer_bounded_cbilinear_bounded_clinear[transfer_rule]:
"(rel_fun (rel_fun (=) (pcr_cblinfun (=) (=))) (=)) bounded_cbilinear bounded_clinear"
by (auto simp: pcr_cblinfun_def cr_cblinfun_def rel_fun_def OO_def
intro!: transfer_bounded_cbilinear_bounded_clinearI)
lemma transfer_bounded_sesquilinear_bounded_antilinearI:
assumes "g = (λi x. (cblinfun_apply (f i) x))"
shows "bounded_sesquilinear g = bounded_antilinear f"
proof
assume "bounded_sesquilinear g"
then interpret bounded_sesquilinear f by (simp add: assms)
show "bounded_antilinear f"
proof (unfold_locales, safe intro!: cblinfun_eqI)
fix i
show "f (x + y) i = (f x + f y) i" "f (r *⇩C x) i = (cnj r *⇩C f x) i" for r x y
by (auto intro!: cblinfun_eqI simp: cblinfun.scaleC_left scaleC_left add_left cblinfun.add_left)
from _ real.nonneg_bounded show "∃K. ∀x. norm (f x) ≤ norm x * K"
by (rule ex_reg) (auto intro!: onorm_bound simp: norm_cblinfun.rep_eq ac_simps)
qed
next
assume "bounded_antilinear f"
then obtain K where K: ‹norm (f x) ≤ norm x * K› for x
using bounded_antilinear.bounded by blast
have ‹norm (cblinfun_apply (f a) b) ≤ norm (f a) * norm b› for a b
by (simp add: norm_cblinfun)
also have ‹… a b ≤ norm a * norm b * K› for a b
by (smt (verit, best) K mult.assoc mult.commute mult_mono' norm_ge_zero)
finally have *: ‹norm (cblinfun_apply (f a) b) ≤ norm a * norm b * K› for a b
by simp
show "bounded_sesquilinear g"
using ‹bounded_antilinear f›
apply (auto intro!: bounded_sesquilinear.intro simp: assms cblinfun.add_left cblinfun.add_right
linear_simps bounded_antilinear.bounded_linear antilinear.scaleC bounded_antilinear.antilinear
cblinfun.scaleC_left cblinfun.scaleC_right)
using * by blast
qed
lemma transfer_bounded_sesquilinear_bounded_antilinear[transfer_rule]:
"(rel_fun (rel_fun (=) (pcr_cblinfun (=) (=))) (=)) bounded_sesquilinear bounded_antilinear"
by (auto simp: pcr_cblinfun_def cr_cblinfun_def rel_fun_def OO_def
intro!: transfer_bounded_sesquilinear_bounded_antilinearI)
context bounded_cbilinear
begin
lift_definition prod_left::"'b ⇒ 'a ⇒⇩C⇩L 'c" is "(λb a. prod a b)"
by (rule bounded_clinear_left)
declare prod_left.rep_eq[simp]
lemma bounded_clinear_prod_left[bounded_clinear]: "bounded_clinear prod_left"
by transfer (rule flip)
lift_definition prod_right::"'a ⇒ 'b ⇒⇩C⇩L 'c" is "(λa b. prod a b)"
by (rule bounded_clinear_right)
declare prod_right.rep_eq[simp]
lemma bounded_clinear_prod_right[bounded_clinear]: "bounded_clinear prod_right"
by transfer (rule bounded_cbilinear_axioms)
end
lift_definition id_cblinfun::"'a::complex_normed_vector ⇒⇩C⇩L 'a" is "λx. x"
by (rule bounded_clinear_ident)
lemmas cblinfun_id_cblinfun_apply[simp] = id_cblinfun.rep_eq
lemma norm_cblinfun_id[simp]:
"norm (id_cblinfun::'a::{complex_normed_vector, not_singleton} ⇒⇩C⇩L 'a) = 1"
apply transfer
apply (rule onorm_id[internalize_sort' 'a])
apply standard[1]
by simp
lemma norm_cblinfun_id_le:
"norm (id_cblinfun::'a::complex_normed_vector ⇒⇩C⇩L 'a) ≤ 1"
by transfer (auto simp: onorm_id_le)
lift_definition cblinfun_compose::
"'a::complex_normed_vector ⇒⇩C⇩L 'b::complex_normed_vector ⇒
'c::complex_normed_vector ⇒⇩C⇩L 'a ⇒
'c ⇒⇩C⇩L 'b" (infixl "o⇩C⇩L" 67) is "(o)"
parametric comp_transfer
unfolding o_def
by (rule bounded_clinear_compose)
lemma cblinfun_apply_cblinfun_compose[simp]: "(a o⇩C⇩L b) c = a (b c)"
by (simp add: cblinfun_compose.rep_eq)
lemma norm_cblinfun_compose:
"norm (f o⇩C⇩L g) ≤ norm f * norm g"
apply transfer
by (auto intro!: onorm_compose simp: bounded_clinear.bounded_linear)
lemma bounded_cbilinear_cblinfun_compose[bounded_cbilinear]: "bounded_cbilinear (o⇩C⇩L)"
by unfold_locales
(auto intro!: cblinfun_eqI exI[where x=1] simp: cblinfun.cbilinear_simps norm_cblinfun_compose)
lemma cblinfun_compose_zero[simp]:
"blinfun_compose 0 = (λ_. 0)"
"blinfun_compose x 0 = 0"
by (auto simp: blinfun.bilinear_simps intro!: blinfun_eqI)
lemma cblinfun_bij2:
fixes f::"'a ⇒⇩C⇩L 'a::ceuclidean_space"
assumes "f o⇩C⇩L g = id_cblinfun"
shows "bij (cblinfun_apply g)"
proof (rule bijI)
show "inj g"
using assms
by (metis cblinfun_id_cblinfun_apply cblinfun_compose.rep_eq injI inj_on_imageI2)
then show "surj g"
using bounded_clinear_def cblinfun.bounded_clinear_right ceucl.linear_inj_imp_surj by blast
qed
lemma cblinfun_bij1:
fixes f::"'a ⇒⇩C⇩L 'a::ceuclidean_space"
assumes "f o⇩C⇩L g = id_cblinfun"
shows "bij (cblinfun_apply f)"
proof (rule bijI)
show "surj (cblinfun_apply f)"
by (metis assms cblinfun_apply_cblinfun_compose cblinfun_id_cblinfun_apply surjI)
then show "inj (cblinfun_apply f)"
using bounded_clinear_def cblinfun.bounded_clinear_right ceucl.linear_surjective_imp_injective by blast
qed
lift_definition cblinfun_cinner_right::"'a::complex_inner ⇒ 'a ⇒⇩C⇩L complex" is "(∙⇩C)"
by (rule bounded_clinear_cinner_right)
declare cblinfun_cinner_right.rep_eq[simp]
lemma bounded_antilinear_cblinfun_cinner_right[bounded_antilinear]: "bounded_antilinear cblinfun_cinner_right"
apply transfer by (simp add: bounded_sesquilinear_cinner)
lift_definition cblinfun_scaleC_right::"complex ⇒ 'a ⇒⇩C⇩L 'a::complex_normed_vector" is "(*⇩C)"
by (rule bounded_clinear_scaleC_right)
declare cblinfun_scaleC_right.rep_eq[simp]
lemma bounded_clinear_cblinfun_scaleC_right[bounded_clinear]: "bounded_clinear cblinfun_scaleC_right"
by transfer (rule bounded_cbilinear_scaleC)
lift_definition cblinfun_scaleC_left::"'a::complex_normed_vector ⇒ complex ⇒⇩C⇩L 'a" is "λx y. y *⇩C x"
by (rule bounded_clinear_scaleC_left)
lemmas [simp] = cblinfun_scaleC_left.rep_eq
lemma bounded_clinear_cblinfun_scaleC_left[bounded_clinear]: "bounded_clinear cblinfun_scaleC_left"
by transfer (rule bounded_cbilinear.flip[OF bounded_cbilinear_scaleC])
lift_definition cblinfun_mult_right::"'a ⇒ 'a ⇒⇩C⇩L 'a::complex_normed_algebra" is "(*)"
by (rule bounded_clinear_mult_right)
declare cblinfun_mult_right.rep_eq[simp]
lemma bounded_clinear_cblinfun_mult_right[bounded_clinear]: "bounded_clinear cblinfun_mult_right"
by transfer (rule bounded_cbilinear_mult)
lift_definition cblinfun_mult_left::"'a::complex_normed_algebra ⇒ 'a ⇒⇩C⇩L 'a" is "λx y. y * x"
by (rule bounded_clinear_mult_left)
lemmas [simp] = cblinfun_mult_left.rep_eq
lemma bounded_clinear_cblinfun_mult_left[bounded_clinear]: "bounded_clinear cblinfun_mult_left"
by transfer (rule bounded_cbilinear.flip[OF bounded_cbilinear_mult])
lemmas bounded_clinear_function_uniform_limit_intros[uniform_limit_intros] =
bounded_clinear.uniform_limit[OF bounded_clinear_apply_cblinfun]
bounded_clinear.uniform_limit[OF bounded_clinear_cblinfun_apply]
bounded_antilinear.uniform_limit[OF bounded_antilinear_cblinfun_matrix]
subsection ‹The strong operator topology on continuous linear operators›
text ‹Let ‹'a› and ‹'b› be two normed real vector spaces. Then the space of linear continuous
operators from ‹'a› to ‹'b› has a canonical norm, and therefore a canonical corresponding topology
(the type classes instantiation are given in 🗏‹Complex_Bounded_Linear_Function0.thy›).
However, there is another topology on this space, the strong operator topology, where ‹T⇩n› tends to
‹T› iff, for all ‹x› in ‹'a›, then ‹T⇩n x› tends to ‹T x›. This is precisely the product topology
where the target space is endowed with the norm topology. It is especially useful when ‹'b› is the set
of real numbers, since then this topology is compact.
We can not implement it using type classes as there is already a topology, but at least we
can define it as a topology.
Note that there is yet another (common and useful) topology on operator spaces, the weak operator
topology, defined analogously using the product topology, but where the target space is given the
weak-* topology, i.e., the pullback of the weak topology on the bidual of the space under the
canonical embedding of a space into its bidual. We do not define it there, although it could also be
defined analogously.
›
definition cstrong_operator_topology::"('a::complex_normed_vector ⇒⇩C⇩L'b::complex_normed_vector) topology"
where "cstrong_operator_topology = pullback_topology UNIV cblinfun_apply euclidean"
lemma cstrong_operator_topology_topspace:
"topspace cstrong_operator_topology = UNIV"
unfolding cstrong_operator_topology_def topspace_pullback_topology topspace_euclidean by auto
lemma cstrong_operator_topology_basis:
fixes f::"('a::complex_normed_vector ⇒⇩C⇩L'b::complex_normed_vector)" and U::"'i ⇒ 'b set" and x::"'i ⇒ 'a"
assumes "finite I" "⋀i. i ∈ I ⟹ open (U i)"
shows "openin cstrong_operator_topology {f. ∀i∈I. cblinfun_apply f (x i) ∈ U i}"
proof -
have "open {g::('a⇒'b). ∀i∈I. g (x i) ∈ U i}"
by (rule product_topology_basis'[OF assms])
moreover have "{f. ∀i∈I. cblinfun_apply f (x i) ∈ U i}
= cblinfun_apply-`{g::('a⇒'b). ∀i∈I. g (x i) ∈ U i} ∩ UNIV"
by auto
ultimately show ?thesis
unfolding cstrong_operator_topology_def by (subst openin_pullback_topology) auto
qed
lemma cstrong_operator_topology_continuous_evaluation:
"continuous_map cstrong_operator_topology euclidean (λf. cblinfun_apply f x)"
proof -
have "continuous_map cstrong_operator_topology euclidean ((λf. f x) o cblinfun_apply)"
unfolding cstrong_operator_topology_def apply (rule continuous_map_pullback)
using continuous_on_product_coordinates by fastforce
then show ?thesis unfolding comp_def by simp
qed
lemma continuous_on_cstrong_operator_topo_iff_coordinatewise:
"continuous_map T cstrong_operator_topology f
⟷ (∀x. continuous_map T euclidean (λy. cblinfun_apply (f y) x))"
proof (auto)
fix x::"'b"
assume "continuous_map T cstrong_operator_topology f"
with continuous_map_compose[OF this cstrong_operator_topology_continuous_evaluation]
have "continuous_map T euclidean ((λz. cblinfun_apply z x) o f)"
by simp
then show "continuous_map T euclidean (λy. cblinfun_apply (f y) x)"
unfolding comp_def by auto
next
assume *: "∀x. continuous_map T euclidean (λy. cblinfun_apply (f y) x)"
have "⋀i. continuous_map T euclidean (λx. cblinfun_apply (f x) i)"
using * unfolding comp_def by auto
then have "continuous_map T euclidean (cblinfun_apply o f)"
unfolding o_def
by (metis (no_types) continuous_map_componentwise_UNIV euclidean_product_topology)
show "continuous_map T cstrong_operator_topology f"
unfolding cstrong_operator_topology_def
apply (rule continuous_map_pullback')
by (auto simp add: ‹continuous_map T euclidean (cblinfun_apply o f)›)
qed
lemma cstrong_operator_topology_weaker_than_euclidean:
"continuous_map euclidean cstrong_operator_topology (λf. f)"
apply (subst continuous_on_cstrong_operator_topo_iff_coordinatewise)
by (auto simp add: linear_continuous_on continuous_at_imp_continuous_on linear_continuous_at
bounded_clinear.bounded_linear)
end