| author | immler | 
| Sat, 29 Sep 2018 16:30:44 -0400 | |
| changeset 69096 | 62a0d10386c1 | 
| parent 69064 | 5840724b1d71 | 
| child 69260 | 0a9688695a1b | 
| permissions | -rw-r--r-- | 
| 63627 | 1  | 
(* Title: HOL/Analysis/Bounded_Linear_Function.thy  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
2  | 
Author: Fabian Immler, TU München  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
3  | 
*)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
4  | 
|
| 61975 | 5  | 
section \<open>Bounded Linear Function\<close>  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
6  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
7  | 
theory Bounded_Linear_Function  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
8  | 
imports  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
9  | 
Topology_Euclidean_Space  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
10  | 
Operator_Norm  | 
| 
67685
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
11  | 
Uniform_Limit  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
12  | 
begin  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
13  | 
|
| 
67685
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
14  | 
lemma onorm_componentwise:  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
15  | 
assumes "bounded_linear f"  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
16  | 
shows "onorm f \<le> (\<Sum>i\<in>Basis. norm (f i))"  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
17  | 
proof -  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
18  | 
  {
 | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
19  | 
fix i::'a  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
20  | 
assume "i \<in> Basis"  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
21  | 
hence "onorm (\<lambda>x. (x \<bullet> i) *\<^sub>R f i) \<le> onorm (\<lambda>x. (x \<bullet> i)) * norm (f i)"  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
22  | 
by (auto intro!: onorm_scaleR_left_lemma bounded_linear_inner_left)  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
23  | 
also have "\<dots> \<le> norm i * norm (f i)"  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
24  | 
by (rule mult_right_mono)  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
25  | 
(auto simp: ac_simps Cauchy_Schwarz_ineq2 intro!: onorm_le)  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
26  | 
finally have "onorm (\<lambda>x. (x \<bullet> i) *\<^sub>R f i) \<le> norm (f i)" using \<open>i \<in> Basis\<close>  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
27  | 
by simp  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
28  | 
} hence "onorm (\<lambda>x. \<Sum>i\<in>Basis. (x \<bullet> i) *\<^sub>R f i) \<le> (\<Sum>i\<in>Basis. norm (f i))"  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
29  | 
by (auto intro!: order_trans[OF onorm_sum_le] bounded_linear_scaleR_const  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
30  | 
sum_mono bounded_linear_inner_left)  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
31  | 
also have "(\<lambda>x. \<Sum>i\<in>Basis. (x \<bullet> i) *\<^sub>R f i) = (\<lambda>x. f (\<Sum>i\<in>Basis. (x \<bullet> i) *\<^sub>R i))"  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
32  | 
by (simp add: linear_sum bounded_linear.linear assms linear_simps)  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
33  | 
also have "\<dots> = f"  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
34  | 
by (simp add: euclidean_representation)  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
35  | 
finally show ?thesis .  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
36  | 
qed  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
37  | 
|
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
38  | 
lemmas onorm_componentwise_le = order_trans[OF onorm_componentwise]  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
39  | 
|
| 68838 | 40  | 
subsection%unimportant \<open>Intro rules for @{term bounded_linear}\<close>
 | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
41  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
42  | 
named_theorems bounded_linear_intros  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
43  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
44  | 
lemma onorm_inner_left:  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
45  | 
assumes "bounded_linear r"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
46  | 
shows "onorm (\<lambda>x. r x \<bullet> f) \<le> onorm r * norm f"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
47  | 
proof (rule onorm_bound)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
48  | 
fix x  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
49  | 
have "norm (r x \<bullet> f) \<le> norm (r x) * norm f"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
50  | 
by (simp add: Cauchy_Schwarz_ineq2)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
51  | 
also have "\<dots> \<le> onorm r * norm x * norm f"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
52  | 
by (intro mult_right_mono onorm assms norm_ge_zero)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
53  | 
finally show "norm (r x \<bullet> f) \<le> onorm r * norm f * norm x"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
54  | 
by (simp add: ac_simps)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
55  | 
qed (intro mult_nonneg_nonneg norm_ge_zero onorm_pos_le assms)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
56  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
57  | 
lemma onorm_inner_right:  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
58  | 
assumes "bounded_linear r"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
59  | 
shows "onorm (\<lambda>x. f \<bullet> r x) \<le> norm f * onorm r"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
60  | 
apply (subst inner_commute)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
61  | 
apply (rule onorm_inner_left[OF assms, THEN order_trans])  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
62  | 
apply simp  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
63  | 
done  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
64  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
65  | 
lemmas [bounded_linear_intros] =  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
66  | 
bounded_linear_zero  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
67  | 
bounded_linear_add  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
68  | 
bounded_linear_const_mult  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
69  | 
bounded_linear_mult_const  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
70  | 
bounded_linear_scaleR_const  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
71  | 
bounded_linear_const_scaleR  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
72  | 
bounded_linear_ident  | 
| 64267 | 73  | 
bounded_linear_sum  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
74  | 
bounded_linear_Pair  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
75  | 
bounded_linear_sub  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
76  | 
bounded_linear_fst_comp  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
77  | 
bounded_linear_snd_comp  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
78  | 
bounded_linear_inner_left_comp  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
79  | 
bounded_linear_inner_right_comp  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
80  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
81  | 
|
| 68838 | 82  | 
subsection%unimportant \<open>declaration of derivative/continuous/tendsto introduction rules for bounded linear functions\<close>  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
83  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
84  | 
attribute_setup bounded_linear =  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
85  | 
\<open>Scan.succeed (Thm.declaration_attribute (fn thm =>  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
86  | 
fold (fn (r, s) => Named_Theorems.add_thm s (thm RS r))  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
87  | 
[  | 
| 62311 | 88  | 
        (@{thm bounded_linear.has_derivative}, @{named_theorems derivative_intros}),
 | 
89  | 
        (@{thm bounded_linear.tendsto}, @{named_theorems tendsto_intros}),
 | 
|
90  | 
        (@{thm bounded_linear.continuous}, @{named_theorems continuous_intros}),
 | 
|
91  | 
        (@{thm bounded_linear.continuous_on}, @{named_theorems continuous_intros}),
 | 
|
92  | 
        (@{thm bounded_linear.uniformly_continuous_on}, @{named_theorems continuous_intros}),
 | 
|
93  | 
        (@{thm bounded_linear_compose}, @{named_theorems bounded_linear_intros})
 | 
|
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
94  | 
]))\<close>  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
95  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
96  | 
attribute_setup bounded_bilinear =  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
97  | 
\<open>Scan.succeed (Thm.declaration_attribute (fn thm =>  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
98  | 
fold (fn (r, s) => Named_Theorems.add_thm s (thm RS r))  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
99  | 
[  | 
| 62311 | 100  | 
        (@{thm bounded_bilinear.FDERIV}, @{named_theorems derivative_intros}),
 | 
101  | 
        (@{thm bounded_bilinear.tendsto}, @{named_theorems tendsto_intros}),
 | 
|
102  | 
        (@{thm bounded_bilinear.continuous}, @{named_theorems continuous_intros}),
 | 
|
103  | 
        (@{thm bounded_bilinear.continuous_on}, @{named_theorems continuous_intros}),
 | 
|
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
104  | 
        (@{thm bounded_linear_compose[OF bounded_bilinear.bounded_linear_left]},
 | 
| 62311 | 105  | 
          @{named_theorems bounded_linear_intros}),
 | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
106  | 
        (@{thm bounded_linear_compose[OF bounded_bilinear.bounded_linear_right]},
 | 
| 62311 | 107  | 
          @{named_theorems bounded_linear_intros}),
 | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
108  | 
        (@{thm bounded_linear.uniformly_continuous_on[OF bounded_bilinear.bounded_linear_left]},
 | 
| 62311 | 109  | 
          @{named_theorems continuous_intros}),
 | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
110  | 
        (@{thm bounded_linear.uniformly_continuous_on[OF bounded_bilinear.bounded_linear_right]},
 | 
| 62311 | 111  | 
          @{named_theorems continuous_intros})
 | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
112  | 
]))\<close>  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
113  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
114  | 
|
| 68838 | 115  | 
subsection \<open>Type of bounded linear functions\<close>  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
116  | 
|
| 68838 | 117  | 
typedef%important (overloaded) ('a, 'b) blinfun ("(_ \<Rightarrow>\<^sub>L /_)" [22, 21] 21) =
 | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
118  | 
  "{f::'a::real_normed_vector\<Rightarrow>'b::real_normed_vector. bounded_linear f}"
 | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
119  | 
morphisms blinfun_apply Blinfun  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
120  | 
by (blast intro: bounded_linear_intros)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
121  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
122  | 
declare [[coercion  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
123  | 
    "blinfun_apply :: ('a::real_normed_vector \<Rightarrow>\<^sub>L'b::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> 'b"]]
 | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
124  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
125  | 
lemma bounded_linear_blinfun_apply[bounded_linear_intros]:  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
126  | 
"bounded_linear g \<Longrightarrow> bounded_linear (\<lambda>x. blinfun_apply f (g x))"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
127  | 
by (metis blinfun_apply mem_Collect_eq bounded_linear_compose)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
128  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
129  | 
setup_lifting type_definition_blinfun  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
130  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
131  | 
lemma blinfun_eqI: "(\<And>i. blinfun_apply x i = blinfun_apply y i) \<Longrightarrow> x = y"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
132  | 
by transfer auto  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
133  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
134  | 
lemma bounded_linear_Blinfun_apply: "bounded_linear f \<Longrightarrow> blinfun_apply (Blinfun f) = f"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
135  | 
by (auto simp: Blinfun_inverse)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
136  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
137  | 
|
| 68838 | 138  | 
subsection \<open>Type class instantiations\<close>  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
139  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
140  | 
instantiation blinfun :: (real_normed_vector, real_normed_vector) real_normed_vector  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
141  | 
begin  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
142  | 
|
| 68838 | 143  | 
lift_definition%important norm_blinfun :: "'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> real" is onorm .  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
144  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
145  | 
lift_definition minus_blinfun :: "'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
146  | 
is "\<lambda>f g x. f x - g x"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
147  | 
by (rule bounded_linear_sub)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
148  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
149  | 
definition dist_blinfun :: "'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> real"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
150  | 
where "dist_blinfun a b = norm (a - b)"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
151  | 
|
| 62101 | 152  | 
definition [code del]:  | 
153  | 
  "(uniformity :: (('a \<Rightarrow>\<^sub>L 'b) \<times> ('a \<Rightarrow>\<^sub>L 'b)) filter) = (INF e:{0 <..}. principal {(x, y). dist x y < e})"
 | 
|
154  | 
||
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
155  | 
definition open_blinfun :: "('a \<Rightarrow>\<^sub>L 'b) set \<Rightarrow> bool"
 | 
| 62101 | 156  | 
where [code del]: "open_blinfun S = (\<forall>x\<in>S. \<forall>\<^sub>F (x', y) in uniformity. x' = x \<longrightarrow> y \<in> S)"  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
157  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
158  | 
lift_definition uminus_blinfun :: "'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b" is "\<lambda>f x. - f x"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
159  | 
by (rule bounded_linear_minus)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
160  | 
|
| 68838 | 161  | 
lift_definition%important zero_blinfun :: "'a \<Rightarrow>\<^sub>L 'b" is "\<lambda>x. 0"  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
162  | 
by (rule bounded_linear_zero)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
163  | 
|
| 68838 | 164  | 
lift_definition%important plus_blinfun :: "'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b"  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
165  | 
is "\<lambda>f g x. f x + g x"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
166  | 
by (metis bounded_linear_add)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
167  | 
|
| 68838 | 168  | 
lift_definition%important scaleR_blinfun::"real \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b" is "\<lambda>r f x. r *\<^sub>R f x"  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
169  | 
by (metis bounded_linear_compose bounded_linear_scaleR_right)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
170  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
171  | 
definition sgn_blinfun :: "'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
172  | 
where "sgn_blinfun x = scaleR (inverse (norm x)) x"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
173  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
174  | 
instance  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
175  | 
apply standard  | 
| 62101 | 176  | 
unfolding dist_blinfun_def open_blinfun_def sgn_blinfun_def uniformity_blinfun_def  | 
177  | 
apply (rule refl | (transfer, force simp: onorm_triangle onorm_scaleR onorm_eq_0 algebra_simps))+  | 
|
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
178  | 
done  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
179  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
180  | 
end  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
181  | 
|
| 
62102
 
877463945ce9
fix code generation for uniformity: uniformity is a non-computable pure data.
 
hoelzl 
parents: 
62101 
diff
changeset
 | 
182  | 
declare uniformity_Abort[where 'a="('a :: real_normed_vector) \<Rightarrow>\<^sub>L ('b :: real_normed_vector)", code]
 | 
| 
 
877463945ce9
fix code generation for uniformity: uniformity is a non-computable pure data.
 
hoelzl 
parents: 
62101 
diff
changeset
 | 
183  | 
|
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
184  | 
lemma norm_blinfun_eqI:  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
185  | 
assumes "n \<le> norm (blinfun_apply f x) / norm x"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
186  | 
assumes "\<And>x. norm (blinfun_apply f x) \<le> n * norm x"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
187  | 
assumes "0 \<le> n"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
188  | 
shows "norm f = n"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
189  | 
by (auto simp: norm_blinfun_def  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
190  | 
intro!: antisym onorm_bound assms order_trans[OF _ le_onorm]  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
191  | 
bounded_linear_intros)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
192  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
193  | 
lemma norm_blinfun: "norm (blinfun_apply f x) \<le> norm f * norm x"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
194  | 
by transfer (rule onorm)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
195  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
196  | 
lemma norm_blinfun_bound: "0 \<le> b \<Longrightarrow> (\<And>x. norm (blinfun_apply f x) \<le> b * norm x) \<Longrightarrow> norm f \<le> b"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
197  | 
by transfer (rule onorm_bound)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
198  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
199  | 
lemma bounded_bilinear_blinfun_apply[bounded_bilinear]: "bounded_bilinear blinfun_apply"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
200  | 
proof  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
201  | 
fix f g::"'a \<Rightarrow>\<^sub>L 'b" and a b::'a and r::real  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
202  | 
show "(f + g) a = f a + g a" "(r *\<^sub>R f) a = r *\<^sub>R f a"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
203  | 
by (transfer, simp)+  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
204  | 
interpret bounded_linear f for f::"'a \<Rightarrow>\<^sub>L 'b"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
205  | 
by (auto intro!: bounded_linear_intros)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
206  | 
show "f (a + b) = f a + f b" "f (r *\<^sub>R a) = r *\<^sub>R f a"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
207  | 
by (simp_all add: add scaleR)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
208  | 
show "\<exists>K. \<forall>a b. norm (blinfun_apply a b) \<le> norm a * norm b * K"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
209  | 
by (auto intro!: exI[where x=1] norm_blinfun)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
210  | 
qed  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
211  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
212  | 
interpretation blinfun: bounded_bilinear blinfun_apply  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
213  | 
by (rule bounded_bilinear_blinfun_apply)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
214  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
215  | 
lemmas bounded_linear_apply_blinfun[intro, simp] = blinfun.bounded_linear_left  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
216  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
217  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
218  | 
context bounded_bilinear  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
219  | 
begin  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
220  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
221  | 
named_theorems bilinear_simps  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
222  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
223  | 
lemmas [bilinear_simps] =  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
224  | 
add_left  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
225  | 
add_right  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
226  | 
diff_left  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
227  | 
diff_right  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
228  | 
minus_left  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
229  | 
minus_right  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
230  | 
scaleR_left  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
231  | 
scaleR_right  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
232  | 
zero_left  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
233  | 
zero_right  | 
| 64267 | 234  | 
sum_left  | 
235  | 
sum_right  | 
|
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
236  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
237  | 
end  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
238  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
239  | 
|
| 66933 | 240  | 
instance blinfun :: (real_normed_vector, banach) banach  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
241  | 
proof  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
242  | 
fix X::"nat \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
243  | 
assume "Cauchy X"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
244  | 
  {
 | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
245  | 
fix x::'a  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
246  | 
    {
 | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
247  | 
fix x::'a  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
248  | 
assume "norm x \<le> 1"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
249  | 
have "Cauchy (\<lambda>n. X n x)"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
250  | 
proof (rule CauchyI)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
251  | 
fix e::real  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
252  | 
assume "0 < e"  | 
| 61975 | 253  | 
from CauchyD[OF \<open>Cauchy X\<close> \<open>0 < e\<close>] obtain M  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
254  | 
where M: "\<And>m n. m \<ge> M \<Longrightarrow> n \<ge> M \<Longrightarrow> norm (X m - X n) < e"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
255  | 
by auto  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
256  | 
show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m x - X n x) < e"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
257  | 
proof (safe intro!: exI[where x=M])  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
258  | 
fix m n  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
259  | 
assume le: "M \<le> m" "M \<le> n"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
260  | 
have "norm (X m x - X n x) = norm ((X m - X n) x)"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
261  | 
by (simp add: blinfun.bilinear_simps)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
262  | 
also have "\<dots> \<le> norm (X m - X n) * norm x"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
263  | 
by (rule norm_blinfun)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
264  | 
also have "\<dots> \<le> norm (X m - X n) * 1"  | 
| 61975 | 265  | 
using \<open>norm x \<le> 1\<close> norm_ge_zero by (rule mult_left_mono)  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
266  | 
also have "\<dots> = norm (X m - X n)" by simp  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
267  | 
also have "\<dots> < e" using le by fact  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
268  | 
finally show "norm (X m x - X n x) < e" .  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
269  | 
qed  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
270  | 
qed  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
271  | 
hence "convergent (\<lambda>n. X n x)"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
272  | 
by (metis Cauchy_convergent_iff)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
273  | 
} note convergent_norm1 = this  | 
| 63040 | 274  | 
define y where "y = x /\<^sub>R norm x"  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
275  | 
have y: "norm y \<le> 1" and xy: "x = norm x *\<^sub>R y"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
276  | 
by (simp_all add: y_def inverse_eq_divide)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
277  | 
have "convergent (\<lambda>n. norm x *\<^sub>R X n y)"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
278  | 
by (intro bounded_bilinear.convergent[OF bounded_bilinear_scaleR] convergent_const  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
279  | 
convergent_norm1 y)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
280  | 
also have "(\<lambda>n. norm x *\<^sub>R X n y) = (\<lambda>n. X n x)"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
281  | 
by (subst xy) (simp add: blinfun.bilinear_simps)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
282  | 
finally have "convergent (\<lambda>n. X n x)" .  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
283  | 
}  | 
| 61969 | 284  | 
then obtain v where v: "\<And>x. (\<lambda>n. X n x) \<longlonglongrightarrow> v x"  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
285  | 
unfolding convergent_def  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
286  | 
by metis  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
287  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
288  | 
have "Cauchy (\<lambda>n. norm (X n))"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
289  | 
proof (rule CauchyI)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
290  | 
fix e::real  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
291  | 
assume "e > 0"  | 
| 61975 | 292  | 
from CauchyD[OF \<open>Cauchy X\<close> \<open>0 < e\<close>] obtain M  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
293  | 
where M: "\<And>m n. m \<ge> M \<Longrightarrow> n \<ge> M \<Longrightarrow> norm (X m - X n) < e"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
294  | 
by auto  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
295  | 
show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (norm (X m) - norm (X n)) < e"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
296  | 
proof (safe intro!: exI[where x=M])  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
297  | 
fix m n assume mn: "m \<ge> M" "n \<ge> M"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
298  | 
have "norm (norm (X m) - norm (X n)) \<le> norm (X m - X n)"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
299  | 
by (metis norm_triangle_ineq3 real_norm_def)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
300  | 
also have "\<dots> < e" using mn by fact  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
301  | 
finally show "norm (norm (X m) - norm (X n)) < e" .  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
302  | 
qed  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
303  | 
qed  | 
| 61969 | 304  | 
then obtain K where K: "(\<lambda>n. norm (X n)) \<longlonglongrightarrow> K"  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
305  | 
unfolding Cauchy_convergent_iff convergent_def  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
306  | 
by metis  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
307  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
308  | 
have "bounded_linear v"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
309  | 
proof  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
310  | 
fix x y and r::real  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
311  | 
from tendsto_add[OF v[of x] v [of y]] v[of "x + y", unfolded blinfun.bilinear_simps]  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
312  | 
tendsto_scaleR[OF tendsto_const[of r] v[of x]] v[of "r *\<^sub>R x", unfolded blinfun.bilinear_simps]  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
313  | 
show "v (x + y) = v x + v y" "v (r *\<^sub>R x) = r *\<^sub>R v x"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
314  | 
by (metis (poly_guards_query) LIMSEQ_unique)+  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
315  | 
show "\<exists>K. \<forall>x. norm (v x) \<le> norm x * K"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
316  | 
proof (safe intro!: exI[where x=K])  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
317  | 
fix x  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
318  | 
have "norm (v x) \<le> K * norm x"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
319  | 
by (rule tendsto_le[OF _ tendsto_mult[OF K tendsto_const] tendsto_norm[OF v]])  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
320  | 
(auto simp: norm_blinfun)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
321  | 
thus "norm (v x) \<le> norm x * K"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
322  | 
by (simp add: ac_simps)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
323  | 
qed  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
324  | 
qed  | 
| 61969 | 325  | 
hence Bv: "\<And>x. (\<lambda>n. X n x) \<longlonglongrightarrow> Blinfun v x"  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
326  | 
by (auto simp: bounded_linear_Blinfun_apply v)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
327  | 
|
| 61969 | 328  | 
have "X \<longlonglongrightarrow> Blinfun v"  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
329  | 
proof (rule LIMSEQ_I)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
330  | 
fix r::real assume "r > 0"  | 
| 63040 | 331  | 
define r' where "r' = r / 2"  | 
| 61975 | 332  | 
have "0 < r'" "r' < r" using \<open>r > 0\<close> by (simp_all add: r'_def)  | 
333  | 
from CauchyD[OF \<open>Cauchy X\<close> \<open>r' > 0\<close>]  | 
|
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
334  | 
obtain M where M: "\<And>m n. m \<ge> M \<Longrightarrow> n \<ge> M \<Longrightarrow> norm (X m - X n) < r'"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
335  | 
by metis  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
336  | 
show "\<exists>no. \<forall>n\<ge>no. norm (X n - Blinfun v) < r"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
337  | 
proof (safe intro!: exI[where x=M])  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
338  | 
fix n assume n: "M \<le> n"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
339  | 
have "norm (X n - Blinfun v) \<le> r'"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
340  | 
proof (rule norm_blinfun_bound)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
341  | 
fix x  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
342  | 
have "eventually (\<lambda>m. m \<ge> M) sequentially"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
343  | 
by (metis eventually_ge_at_top)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
344  | 
hence ev_le: "eventually (\<lambda>m. norm (X n x - X m x) \<le> r' * norm x) sequentially"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
345  | 
proof eventually_elim  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
346  | 
case (elim m)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
347  | 
have "norm (X n x - X m x) = norm ((X n - X m) x)"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
348  | 
by (simp add: blinfun.bilinear_simps)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
349  | 
also have "\<dots> \<le> norm ((X n - X m)) * norm x"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
350  | 
by (rule norm_blinfun)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
351  | 
also have "\<dots> \<le> r' * norm x"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
352  | 
using M[OF n elim] by (simp add: mult_right_mono)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
353  | 
finally show ?case .  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
354  | 
qed  | 
| 61969 | 355  | 
have tendsto_v: "(\<lambda>m. norm (X n x - X m x)) \<longlonglongrightarrow> norm (X n x - Blinfun v x)"  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
356  | 
by (auto intro!: tendsto_intros Bv)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
357  | 
show "norm ((X n - Blinfun v) x) \<le> r' * norm x"  | 
| 
63952
 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 
paulson <lp15@cam.ac.uk> 
parents: 
63938 
diff
changeset
 | 
358  | 
by (auto intro!: tendsto_upperbound tendsto_v ev_le simp: blinfun.bilinear_simps)  | 
| 61975 | 359  | 
qed (simp add: \<open>0 < r'\<close> less_imp_le)  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
360  | 
thus "norm (X n - Blinfun v) < r"  | 
| 61975 | 361  | 
by (metis \<open>r' < r\<close> le_less_trans)  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
362  | 
qed  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
363  | 
qed  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
364  | 
thus "convergent X"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
365  | 
by (rule convergentI)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
366  | 
qed  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
367  | 
|
| 68838 | 368  | 
subsection%unimportant \<open>On Euclidean Space\<close>  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
369  | 
|
| 64267 | 370  | 
lemma Zfun_sum:  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
371  | 
assumes "finite s"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
372  | 
assumes f: "\<And>i. i \<in> s \<Longrightarrow> Zfun (f i) F"  | 
| 64267 | 373  | 
shows "Zfun (\<lambda>x. sum (\<lambda>i. f i x) s) F"  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
374  | 
using assms by induct (auto intro!: Zfun_zero Zfun_add)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
375  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
376  | 
lemma norm_blinfun_euclidean_le:  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
377  | 
fixes a::"'a::euclidean_space \<Rightarrow>\<^sub>L 'b::real_normed_vector"  | 
| 64267 | 378  | 
shows "norm a \<le> sum (\<lambda>x. norm (a x)) Basis"  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
379  | 
apply (rule norm_blinfun_bound)  | 
| 64267 | 380  | 
apply (simp add: sum_nonneg)  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
381  | 
apply (subst euclidean_representation[symmetric, where 'a='a])  | 
| 64267 | 382  | 
apply (simp only: blinfun.bilinear_simps sum_distrib_right)  | 
383  | 
apply (rule order.trans[OF norm_sum sum_mono])  | 
|
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
384  | 
apply (simp add: abs_mult mult_right_mono ac_simps Basis_le_norm)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
385  | 
done  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
386  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
387  | 
lemma tendsto_componentwise1:  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
388  | 
fixes a::"'a::euclidean_space \<Rightarrow>\<^sub>L 'b::real_normed_vector"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
389  | 
and b::"'c \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b"  | 
| 61973 | 390  | 
assumes "(\<And>j. j \<in> Basis \<Longrightarrow> ((\<lambda>n. b n j) \<longlongrightarrow> a j) F)"  | 
391  | 
shows "(b \<longlongrightarrow> a) F"  | 
|
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
392  | 
proof -  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
393  | 
have "\<And>j. j \<in> Basis \<Longrightarrow> Zfun (\<lambda>x. norm (b x j - a j)) F"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
394  | 
using assms unfolding tendsto_Zfun_iff Zfun_norm_iff .  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
395  | 
hence "Zfun (\<lambda>x. \<Sum>j\<in>Basis. norm (b x j - a j)) F"  | 
| 64267 | 396  | 
by (auto intro!: Zfun_sum)  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
397  | 
thus ?thesis  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
398  | 
unfolding tendsto_Zfun_iff  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
399  | 
by (rule Zfun_le)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
400  | 
(auto intro!: order_trans[OF norm_blinfun_euclidean_le] simp: blinfun.bilinear_simps)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
401  | 
qed  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
402  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
403  | 
lift_definition  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
404  | 
  blinfun_of_matrix::"('b::euclidean_space \<Rightarrow> 'a::euclidean_space \<Rightarrow> real) \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b"
 | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
405  | 
is "\<lambda>a x. \<Sum>i\<in>Basis. \<Sum>j\<in>Basis. ((x \<bullet> j) * a i j) *\<^sub>R i"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
406  | 
by (intro bounded_linear_intros)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
407  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
408  | 
lemma blinfun_of_matrix_works:  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
409  | 
fixes f::"'a::euclidean_space \<Rightarrow>\<^sub>L 'b::euclidean_space"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
410  | 
shows "blinfun_of_matrix (\<lambda>i j. (f j) \<bullet> i) = f"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
411  | 
proof (transfer, rule, rule euclidean_eqI)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
412  | 
fix f::"'a \<Rightarrow> 'b" and x::'a and b::'b assume "bounded_linear f" and b: "b \<in> Basis"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
413  | 
then interpret bounded_linear f by simp  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
414  | 
have "(\<Sum>j\<in>Basis. \<Sum>i\<in>Basis. (x \<bullet> i * (f i \<bullet> j)) *\<^sub>R j) \<bullet> b  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
415  | 
= (\<Sum>j\<in>Basis. if j = b then (\<Sum>i\<in>Basis. (x \<bullet> i * (f i \<bullet> j))) else 0)"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
416  | 
using b  | 
| 
66804
 
3f9bb52082c4
avoid name clashes on interpretation of abstract locales
 
haftmann 
parents: 
66447 
diff
changeset
 | 
417  | 
by (simp add: inner_sum_left inner_Basis if_distrib cong: if_cong) (simp add: sum.swap)  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
418  | 
also have "\<dots> = (\<Sum>i\<in>Basis. (x \<bullet> i * (f i \<bullet> b)))"  | 
| 64267 | 419  | 
using b by (simp add: sum.delta)  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
420  | 
also have "\<dots> = f x \<bullet> b"  | 
| 63938 | 421  | 
by (metis (mono_tags, lifting) Linear_Algebra.linear_componentwise linear_axioms)  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
422  | 
finally show "(\<Sum>j\<in>Basis. \<Sum>i\<in>Basis. (x \<bullet> i * (f i \<bullet> j)) *\<^sub>R j) \<bullet> b = f x \<bullet> b" .  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
423  | 
qed  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
424  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
425  | 
lemma blinfun_of_matrix_apply:  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
426  | 
"blinfun_of_matrix a x = (\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. ((x \<bullet> j) * a i j) *\<^sub>R i)"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
427  | 
by transfer simp  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
428  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
429  | 
lemma blinfun_of_matrix_minus: "blinfun_of_matrix x - blinfun_of_matrix y = blinfun_of_matrix (x - y)"  | 
| 64267 | 430  | 
by transfer (auto simp: algebra_simps sum_subtractf)  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
431  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
432  | 
lemma norm_blinfun_of_matrix:  | 
| 61945 | 433  | 
"norm (blinfun_of_matrix a) \<le> (\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. \<bar>a i j\<bar>)"  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
434  | 
apply (rule norm_blinfun_bound)  | 
| 64267 | 435  | 
apply (simp add: sum_nonneg)  | 
436  | 
apply (simp only: blinfun_of_matrix_apply sum_distrib_right)  | 
|
437  | 
apply (rule order_trans[OF norm_sum sum_mono])  | 
|
438  | 
apply (rule order_trans[OF norm_sum sum_mono])  | 
|
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
439  | 
apply (simp add: abs_mult mult_right_mono ac_simps Basis_le_norm)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
440  | 
done  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
441  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
442  | 
lemma tendsto_blinfun_of_matrix:  | 
| 61973 | 443  | 
assumes "\<And>i j. i \<in> Basis \<Longrightarrow> j \<in> Basis \<Longrightarrow> ((\<lambda>n. b n i j) \<longlongrightarrow> a i j) F"  | 
444  | 
shows "((\<lambda>n. blinfun_of_matrix (b n)) \<longlongrightarrow> blinfun_of_matrix a) F"  | 
|
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
445  | 
proof -  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
446  | 
have "\<And>i j. i \<in> Basis \<Longrightarrow> j \<in> Basis \<Longrightarrow> Zfun (\<lambda>x. norm (b x i j - a i j)) F"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
447  | 
using assms unfolding tendsto_Zfun_iff Zfun_norm_iff .  | 
| 61945 | 448  | 
hence "Zfun (\<lambda>x. (\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. \<bar>b x i j - a i j\<bar>)) F"  | 
| 64267 | 449  | 
by (auto intro!: Zfun_sum)  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
450  | 
thus ?thesis  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
451  | 
unfolding tendsto_Zfun_iff blinfun_of_matrix_minus  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
452  | 
by (rule Zfun_le) (auto intro!: order_trans[OF norm_blinfun_of_matrix])  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
453  | 
qed  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
454  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
455  | 
lemma tendsto_componentwise:  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
456  | 
fixes a::"'a::euclidean_space \<Rightarrow>\<^sub>L 'b::euclidean_space"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
457  | 
and b::"'c \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b"  | 
| 61973 | 458  | 
shows "(\<And>i j. i \<in> Basis \<Longrightarrow> j \<in> Basis \<Longrightarrow> ((\<lambda>n. b n j \<bullet> i) \<longlongrightarrow> a j \<bullet> i) F) \<Longrightarrow> (b \<longlongrightarrow> a) F"  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
459  | 
apply (subst blinfun_of_matrix_works[of a, symmetric])  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
460  | 
apply (subst blinfun_of_matrix_works[of "b x" for x, symmetric, abs_def])  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
461  | 
by (rule tendsto_blinfun_of_matrix)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
462  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
463  | 
lemma  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
464  | 
continuous_blinfun_componentwiseI:  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
465  | 
fixes f:: "'b::t2_space \<Rightarrow> 'a::euclidean_space \<Rightarrow>\<^sub>L 'c::euclidean_space"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
466  | 
assumes "\<And>i j. i \<in> Basis \<Longrightarrow> j \<in> Basis \<Longrightarrow> continuous F (\<lambda>x. (f x) j \<bullet> i)"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
467  | 
shows "continuous F f"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
468  | 
using assms by (auto simp: continuous_def intro!: tendsto_componentwise)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
469  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
470  | 
lemma  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
471  | 
continuous_blinfun_componentwiseI1:  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
472  | 
fixes f:: "'b::t2_space \<Rightarrow> 'a::euclidean_space \<Rightarrow>\<^sub>L 'c::real_normed_vector"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
473  | 
assumes "\<And>i. i \<in> Basis \<Longrightarrow> continuous F (\<lambda>x. f x i)"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
474  | 
shows "continuous F f"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
475  | 
using assms by (auto simp: continuous_def intro!: tendsto_componentwise1)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
476  | 
|
| 
67685
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
477  | 
lemma  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
478  | 
continuous_on_blinfun_componentwise:  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
479  | 
fixes f:: "'d::t2_space \<Rightarrow> 'e::euclidean_space \<Rightarrow>\<^sub>L 'f::real_normed_vector"  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
480  | 
assumes "\<And>i. i \<in> Basis \<Longrightarrow> continuous_on s (\<lambda>x. f x i)"  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
481  | 
shows "continuous_on s f"  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
482  | 
using assms  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
483  | 
by (auto intro!: continuous_at_imp_continuous_on intro!: tendsto_componentwise1  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
484  | 
simp: continuous_on_eq_continuous_within continuous_def)  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
485  | 
|
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
486  | 
lemma bounded_linear_blinfun_matrix: "bounded_linear (\<lambda>x. (x::_\<Rightarrow>\<^sub>L _) j \<bullet> i)"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
487  | 
by (auto intro!: bounded_linearI' bounded_linear_intros)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
488  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
489  | 
lemma continuous_blinfun_matrix:  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
490  | 
fixes f:: "'b::t2_space \<Rightarrow> 'a::real_normed_vector \<Rightarrow>\<^sub>L 'c::real_inner"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
491  | 
assumes "continuous F f"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
492  | 
shows "continuous F (\<lambda>x. (f x) j \<bullet> i)"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
493  | 
by (rule bounded_linear.continuous[OF bounded_linear_blinfun_matrix assms])  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
494  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
495  | 
lemma continuous_on_blinfun_matrix:  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
496  | 
fixes f::"'a::t2_space \<Rightarrow> 'b::real_normed_vector \<Rightarrow>\<^sub>L 'c::real_inner"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
497  | 
assumes "continuous_on S f"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
498  | 
shows "continuous_on S (\<lambda>x. (f x) j \<bullet> i)"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
499  | 
using assms  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
500  | 
by (auto simp: continuous_on_eq_continuous_within continuous_blinfun_matrix)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
501  | 
|
| 62963 | 502  | 
lemma continuous_on_blinfun_of_matrix[continuous_intros]:  | 
503  | 
assumes "\<And>i j. i \<in> Basis \<Longrightarrow> j \<in> Basis \<Longrightarrow> continuous_on S (\<lambda>s. g s i j)"  | 
|
504  | 
shows "continuous_on S (\<lambda>s. blinfun_of_matrix (g s))"  | 
|
505  | 
using assms  | 
|
506  | 
by (auto simp: continuous_on intro!: tendsto_blinfun_of_matrix)  | 
|
507  | 
||
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
508  | 
lemma mult_if_delta:  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
509  | 
"(if P then (1::'a::comm_semiring_1) else 0) * q = (if P then q else 0)"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
510  | 
by auto  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
511  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
512  | 
lemma compact_blinfun_lemma:  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
513  | 
fixes f :: "nat \<Rightarrow> 'a::euclidean_space \<Rightarrow>\<^sub>L 'b::euclidean_space"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
514  | 
assumes "bounded (range f)"  | 
| 
66447
 
a1f5c5c26fa6
Replaced subseq with strict_mono
 
eberlm <eberlm@in.tum.de> 
parents: 
66089 
diff
changeset
 | 
515  | 
shows "\<forall>d\<subseteq>Basis. \<exists>l::'a \<Rightarrow>\<^sub>L 'b. \<exists> r::nat\<Rightarrow>nat.  | 
| 
 
a1f5c5c26fa6
Replaced subseq with strict_mono
 
eberlm <eberlm@in.tum.de> 
parents: 
66089 
diff
changeset
 | 
516  | 
strict_mono r \<and> (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) i) (l i) < e) sequentially)"  | 
| 62127 | 517  | 
by (rule compact_lemma_general[where unproj = "\<lambda>e. blinfun_of_matrix (\<lambda>i j. e j \<bullet> i)"])  | 
518  | 
(auto intro!: euclidean_eqI[where 'a='b] bounded_linear_image assms  | 
|
| 64267 | 519  | 
simp: blinfun_of_matrix_works blinfun_of_matrix_apply inner_Basis mult_if_delta sum.delta'  | 
520  | 
scaleR_sum_left[symmetric])  | 
|
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
521  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
522  | 
lemma blinfun_euclidean_eqI: "(\<And>i. i \<in> Basis \<Longrightarrow> blinfun_apply x i = blinfun_apply y i) \<Longrightarrow> x = y"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
523  | 
apply (auto intro!: blinfun_eqI)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
524  | 
apply (subst (2) euclidean_representation[symmetric, where 'a='a])  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
525  | 
apply (subst (1) euclidean_representation[symmetric, where 'a='a])  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
526  | 
apply (simp add: blinfun.bilinear_simps)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
527  | 
done  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
528  | 
|
| 62951 | 529  | 
lemma Blinfun_eq_matrix: "bounded_linear f \<Longrightarrow> Blinfun f = blinfun_of_matrix (\<lambda>i j. f j \<bullet> i)"  | 
530  | 
by (intro blinfun_euclidean_eqI)  | 
|
531  | 
(auto simp: blinfun_of_matrix_apply bounded_linear_Blinfun_apply inner_Basis if_distrib  | 
|
| 
68072
 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 
immler 
parents: 
67685 
diff
changeset
 | 
532  | 
if_distribR sum.delta' euclidean_representation  | 
| 62951 | 533  | 
cong: if_cong)  | 
534  | 
||
| 67226 | 535  | 
text \<open>TODO: generalize (via \<open>compact_cball\<close>)?\<close>  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
536  | 
instance blinfun :: (euclidean_space, euclidean_space) heine_borel  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
537  | 
proof  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
538  | 
fix f :: "nat \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
539  | 
assume f: "bounded (range f)"  | 
| 
66447
 
a1f5c5c26fa6
Replaced subseq with strict_mono
 
eberlm <eberlm@in.tum.de> 
parents: 
66089 
diff
changeset
 | 
540  | 
then obtain l::"'a \<Rightarrow>\<^sub>L 'b" and r where r: "strict_mono r"  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
541  | 
and l: "\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>Basis. dist (f (r n) i) (l i) < e) sequentially"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
542  | 
using compact_blinfun_lemma [OF f] by blast  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
543  | 
  {
 | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
544  | 
fix e::real  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
545  | 
    let ?d = "real_of_nat DIM('a) * real_of_nat DIM('b)"
 | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
546  | 
assume "e > 0"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
547  | 
hence "e / ?d > 0" by (simp add: DIM_positive)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
548  | 
with l have "eventually (\<lambda>n. \<forall>i\<in>Basis. dist (f (r n) i) (l i) < e / ?d) sequentially"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
549  | 
by simp  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
550  | 
moreover  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
551  | 
    {
 | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
552  | 
fix n  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
553  | 
assume n: "\<forall>i\<in>Basis. dist (f (r n) i) (l i) < e / ?d"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
554  | 
have "norm (f (r n) - l) = norm (blinfun_of_matrix (\<lambda>i j. (f (r n) - l) j \<bullet> i))"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
555  | 
unfolding blinfun_of_matrix_works ..  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
556  | 
also note norm_blinfun_of_matrix  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
557  | 
also have "(\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. \<bar>(f (r n) - l) j \<bullet> i\<bar>) <  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
558  | 
        (\<Sum>i\<in>(Basis::'b set). e / real_of_nat DIM('b))"
 | 
| 64267 | 559  | 
proof (rule sum_strict_mono)  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
560  | 
fix i::'b assume i: "i \<in> Basis"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
561  | 
have "(\<Sum>j::'a\<in>Basis. \<bar>(f (r n) - l) j \<bullet> i\<bar>) < (\<Sum>j::'a\<in>Basis. e / ?d)"  | 
| 64267 | 562  | 
proof (rule sum_strict_mono)  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
563  | 
fix j::'a assume j: "j \<in> Basis"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
564  | 
have "\<bar>(f (r n) - l) j \<bullet> i\<bar> \<le> norm ((f (r n) - l) j)"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
565  | 
by (simp add: Basis_le_norm i)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
566  | 
also have "\<dots> < e / ?d"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
567  | 
using n i j by (auto simp: dist_norm blinfun.bilinear_simps)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
568  | 
finally show "\<bar>(f (r n) - l) j \<bullet> i\<bar> < e / ?d" by simp  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
569  | 
qed simp_all  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
570  | 
        also have "\<dots> \<le> e / real_of_nat DIM('b)"
 | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
571  | 
by simp  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
572  | 
        finally show "(\<Sum>j\<in>Basis. \<bar>(f (r n) - l) j \<bullet> i\<bar>) < e / real_of_nat DIM('b)"
 | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
573  | 
by simp  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
574  | 
qed simp_all  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
575  | 
also have "\<dots> \<le> e" by simp  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
576  | 
finally have "dist (f (r n)) l < e"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
577  | 
by (auto simp: dist_norm)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
578  | 
}  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
579  | 
ultimately have "eventually (\<lambda>n. dist (f (r n)) l < e) sequentially"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
580  | 
using eventually_elim2 by force  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
581  | 
}  | 
| 61973 | 582  | 
then have *: "((f \<circ> r) \<longlongrightarrow> l) sequentially"  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
583  | 
unfolding o_def tendsto_iff by simp  | 
| 
66447
 
a1f5c5c26fa6
Replaced subseq with strict_mono
 
eberlm <eberlm@in.tum.de> 
parents: 
66089 
diff
changeset
 | 
584  | 
with r show "\<exists>l r. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially"  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
585  | 
by auto  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
586  | 
qed  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
587  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
588  | 
|
| 68838 | 589  | 
subsection%unimportant \<open>concrete bounded linear functions\<close>  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
590  | 
|
| 61916 | 591  | 
lemma transfer_bounded_bilinear_bounded_linearI:  | 
592  | 
assumes "g = (\<lambda>i x. (blinfun_apply (f i) x))"  | 
|
593  | 
shows "bounded_bilinear g = bounded_linear f"  | 
|
594  | 
proof  | 
|
595  | 
assume "bounded_bilinear g"  | 
|
596  | 
then interpret bounded_bilinear f by (simp add: assms)  | 
|
597  | 
show "bounded_linear f"  | 
|
598  | 
proof (unfold_locales, safe intro!: blinfun_eqI)  | 
|
599  | 
fix i  | 
|
600  | 
show "f (x + y) i = (f x + f y) i" "f (r *\<^sub>R x) i = (r *\<^sub>R f x) i" for r x y  | 
|
601  | 
by (auto intro!: blinfun_eqI simp: blinfun.bilinear_simps)  | 
|
602  | 
from _ nonneg_bounded show "\<exists>K. \<forall>x. norm (f x) \<le> norm x * K"  | 
|
603  | 
by (rule ex_reg) (auto intro!: onorm_bound simp: norm_blinfun.rep_eq ac_simps)  | 
|
604  | 
qed  | 
|
605  | 
qed (auto simp: assms intro!: blinfun.comp)  | 
|
606  | 
||
607  | 
lemma transfer_bounded_bilinear_bounded_linear[transfer_rule]:  | 
|
| 67399 | 608  | 
"(rel_fun (rel_fun (=) (pcr_blinfun (=) (=))) (=)) bounded_bilinear bounded_linear"  | 
| 61916 | 609  | 
by (auto simp: pcr_blinfun_def cr_blinfun_def rel_fun_def OO_def  | 
610  | 
intro!: transfer_bounded_bilinear_bounded_linearI)  | 
|
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
611  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
612  | 
context bounded_bilinear  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
613  | 
begin  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
614  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
615  | 
lift_definition prod_left::"'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'c" is "(\<lambda>b a. prod a b)"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
616  | 
by (rule bounded_linear_left)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
617  | 
declare prod_left.rep_eq[simp]  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
618  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
619  | 
lemma bounded_linear_prod_left[bounded_linear]: "bounded_linear prod_left"  | 
| 61916 | 620  | 
by transfer (rule flip)  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
621  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
622  | 
lift_definition prod_right::"'a \<Rightarrow> 'b \<Rightarrow>\<^sub>L 'c" is "(\<lambda>a b. prod a b)"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
623  | 
by (rule bounded_linear_right)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
624  | 
declare prod_right.rep_eq[simp]  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
625  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
626  | 
lemma bounded_linear_prod_right[bounded_linear]: "bounded_linear prod_right"  | 
| 61916 | 627  | 
by transfer (rule bounded_bilinear_axioms)  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
628  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
629  | 
end  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
630  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
631  | 
lift_definition id_blinfun::"'a::real_normed_vector \<Rightarrow>\<^sub>L 'a" is "\<lambda>x. x"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
632  | 
by (rule bounded_linear_ident)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
633  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
634  | 
lemmas blinfun_apply_id_blinfun[simp] = id_blinfun.rep_eq  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
635  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
636  | 
lemma norm_blinfun_id[simp]:  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
637  | 
  "norm (id_blinfun::'a::{real_normed_vector, perfect_space} \<Rightarrow>\<^sub>L 'a) = 1"
 | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
638  | 
by transfer (auto simp: onorm_id)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
639  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
640  | 
lemma norm_blinfun_id_le:  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
641  | 
"norm (id_blinfun::'a::real_normed_vector \<Rightarrow>\<^sub>L 'a) \<le> 1"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
642  | 
by transfer (auto simp: onorm_id_le)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
643  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
644  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
645  | 
lift_definition fst_blinfun::"('a::real_normed_vector \<times> 'b::real_normed_vector) \<Rightarrow>\<^sub>L 'a" is fst
 | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
646  | 
by (rule bounded_linear_fst)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
647  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
648  | 
lemma blinfun_apply_fst_blinfun[simp]: "blinfun_apply fst_blinfun = fst"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
649  | 
by transfer (rule refl)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
650  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
651  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
652  | 
lift_definition snd_blinfun::"('a::real_normed_vector \<times> 'b::real_normed_vector) \<Rightarrow>\<^sub>L 'b" is snd
 | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
653  | 
by (rule bounded_linear_snd)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
654  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
655  | 
lemma blinfun_apply_snd_blinfun[simp]: "blinfun_apply snd_blinfun = snd"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
656  | 
by transfer (rule refl)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
657  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
658  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
659  | 
lift_definition blinfun_compose::  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
660  | 
"'a::real_normed_vector \<Rightarrow>\<^sub>L 'b::real_normed_vector \<Rightarrow>  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
661  | 
'c::real_normed_vector \<Rightarrow>\<^sub>L 'a \<Rightarrow>  | 
| 67399 | 662  | 
'c \<Rightarrow>\<^sub>L 'b" (infixl "o\<^sub>L" 55) is "(o)"  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
663  | 
parametric comp_transfer  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
664  | 
unfolding o_def  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
665  | 
by (rule bounded_linear_compose)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
666  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
667  | 
lemma blinfun_apply_blinfun_compose[simp]: "(a o\<^sub>L b) c = a (b c)"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
668  | 
by (simp add: blinfun_compose.rep_eq)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
669  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
670  | 
lemma norm_blinfun_compose:  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
671  | 
"norm (f o\<^sub>L g) \<le> norm f * norm g"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
672  | 
by transfer (rule onorm_compose)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
673  | 
|
| 67399 | 674  | 
lemma bounded_bilinear_blinfun_compose[bounded_bilinear]: "bounded_bilinear (o\<^sub>L)"  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
675  | 
by unfold_locales  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
676  | 
(auto intro!: blinfun_eqI exI[where x=1] simp: blinfun.bilinear_simps norm_blinfun_compose)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
677  | 
|
| 62951 | 678  | 
lemma blinfun_compose_zero[simp]:  | 
679  | 
"blinfun_compose 0 = (\<lambda>_. 0)"  | 
|
680  | 
"blinfun_compose x 0 = 0"  | 
|
681  | 
by (auto simp: blinfun.bilinear_simps intro!: blinfun_eqI)  | 
|
682  | 
||
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
683  | 
|
| 67399 | 684  | 
lift_definition blinfun_inner_right::"'a::real_inner \<Rightarrow> 'a \<Rightarrow>\<^sub>L real" is "(\<bullet>)"  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
685  | 
by (rule bounded_linear_inner_right)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
686  | 
declare blinfun_inner_right.rep_eq[simp]  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
687  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
688  | 
lemma bounded_linear_blinfun_inner_right[bounded_linear]: "bounded_linear blinfun_inner_right"  | 
| 61916 | 689  | 
by transfer (rule bounded_bilinear_inner)  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
690  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
691  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
692  | 
lift_definition blinfun_inner_left::"'a::real_inner \<Rightarrow> 'a \<Rightarrow>\<^sub>L real" is "\<lambda>x y. y \<bullet> x"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
693  | 
by (rule bounded_linear_inner_left)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
694  | 
declare blinfun_inner_left.rep_eq[simp]  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
695  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
696  | 
lemma bounded_linear_blinfun_inner_left[bounded_linear]: "bounded_linear blinfun_inner_left"  | 
| 61916 | 697  | 
by transfer (rule bounded_bilinear.flip[OF bounded_bilinear_inner])  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
698  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
699  | 
|
| 
69064
 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 
nipkow 
parents: 
68838 
diff
changeset
 | 
700  | 
lift_definition blinfun_scaleR_right::"real \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'a::real_normed_vector" is "(*\<^sub>R)"  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
701  | 
by (rule bounded_linear_scaleR_right)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
702  | 
declare blinfun_scaleR_right.rep_eq[simp]  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
703  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
704  | 
lemma bounded_linear_blinfun_scaleR_right[bounded_linear]: "bounded_linear blinfun_scaleR_right"  | 
| 61916 | 705  | 
by transfer (rule bounded_bilinear_scaleR)  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
706  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
707  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
708  | 
lift_definition blinfun_scaleR_left::"'a::real_normed_vector \<Rightarrow> real \<Rightarrow>\<^sub>L 'a" is "\<lambda>x y. y *\<^sub>R x"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
709  | 
by (rule bounded_linear_scaleR_left)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
710  | 
lemmas [simp] = blinfun_scaleR_left.rep_eq  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
711  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
712  | 
lemma bounded_linear_blinfun_scaleR_left[bounded_linear]: "bounded_linear blinfun_scaleR_left"  | 
| 61916 | 713  | 
by transfer (rule bounded_bilinear.flip[OF bounded_bilinear_scaleR])  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
714  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
715  | 
|
| 
69064
 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 
nipkow 
parents: 
68838 
diff
changeset
 | 
716  | 
lift_definition blinfun_mult_right::"'a \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'a::real_normed_algebra" is "(*)"  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
717  | 
by (rule bounded_linear_mult_right)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
718  | 
declare blinfun_mult_right.rep_eq[simp]  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
719  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
720  | 
lemma bounded_linear_blinfun_mult_right[bounded_linear]: "bounded_linear blinfun_mult_right"  | 
| 61916 | 721  | 
by transfer (rule bounded_bilinear_mult)  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
722  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
723  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
724  | 
lift_definition blinfun_mult_left::"'a::real_normed_algebra \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'a" is "\<lambda>x y. y * x"  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
725  | 
by (rule bounded_linear_mult_left)  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
726  | 
lemmas [simp] = blinfun_mult_left.rep_eq  | 
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
727  | 
|
| 
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
728  | 
lemma bounded_linear_blinfun_mult_left[bounded_linear]: "bounded_linear blinfun_mult_left"  | 
| 61916 | 729  | 
by transfer (rule bounded_bilinear.flip[OF bounded_bilinear_mult])  | 
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
730  | 
|
| 
67685
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
731  | 
lemmas bounded_linear_function_uniform_limit_intros[uniform_limit_intros] =  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
732  | 
bounded_linear.uniform_limit[OF bounded_linear_apply_blinfun]  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
733  | 
bounded_linear.uniform_limit[OF bounded_linear_blinfun_apply]  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
734  | 
bounded_linear.uniform_limit[OF bounded_linear_blinfun_matrix]  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
67399 
diff
changeset
 | 
735  | 
|
| 
61915
 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 
immler 
parents:  
diff
changeset
 | 
736  | 
end  |