src/HOL/Analysis/Bounded_Linear_Function.thy
author haftmann
Mon, 02 Aug 2021 10:01:06 +0000
changeset 74101 d804e93ae9ff
parent 71629 2e8f861d21d4
permissions -rw-r--r--
moved theory Bit_Operations into Main corpus
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63627
6ddb43c6b711 rename HOL-Multivariate_Analysis to HOL-Analysis.
hoelzl
parents: 63040
diff changeset
     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
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61973
diff changeset
     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
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    12
  Function_Topology
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    13
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    14
begin
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    15
67685
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67399
diff changeset
    16
lemma onorm_componentwise:
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67399
diff changeset
    17
  assumes "bounded_linear f"
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67399
diff changeset
    18
  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
    19
proof -
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67399
diff changeset
    20
  {
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67399
diff changeset
    21
    fix i::'a
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67399
diff changeset
    22
    assume "i \<in> Basis"
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67399
diff changeset
    23
    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
    24
      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
    25
    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
    26
      by (rule mult_right_mono)
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67399
diff changeset
    27
        (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
    28
    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
    29
      by simp
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67399
diff changeset
    30
  } 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
    31
    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
    32
      sum_mono bounded_linear_inner_left)
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67399
diff changeset
    33
  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
    34
    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
    35
  also have "\<dots> = f"
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67399
diff changeset
    36
    by (simp add: euclidean_representation)
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67399
diff changeset
    37
  finally show ?thesis .
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67399
diff changeset
    38
qed
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67399
diff changeset
    39
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67399
diff changeset
    40
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
    41
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70065
diff changeset
    42
subsection\<^marker>\<open>tag unimportant\<close> \<open>Intro rules for \<^term>\<open>bounded_linear\<close>\<close>
61915
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
named_theorems bounded_linear_intros
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    45
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    46
lemma onorm_inner_left:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    47
  assumes "bounded_linear r"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    48
  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
    49
proof (rule onorm_bound)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    50
  fix x
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    51
  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
    52
    by (simp add: Cauchy_Schwarz_ineq2)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    53
  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
    54
    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
    55
  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
    56
    by (simp add: ac_simps)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    57
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
    58
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    59
lemma onorm_inner_right:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    60
  assumes "bounded_linear r"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    61
  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
    62
  apply (subst inner_commute)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    63
  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
    64
  apply simp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    65
  done
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    66
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    67
lemmas [bounded_linear_intros] =
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    68
  bounded_linear_zero
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    69
  bounded_linear_add
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    70
  bounded_linear_const_mult
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    71
  bounded_linear_mult_const
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    72
  bounded_linear_scaleR_const
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    73
  bounded_linear_const_scaleR
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    74
  bounded_linear_ident
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63952
diff changeset
    75
  bounded_linear_sum
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    76
  bounded_linear_Pair
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    77
  bounded_linear_sub
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    78
  bounded_linear_fst_comp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    79
  bounded_linear_snd_comp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    80
  bounded_linear_inner_left_comp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    81
  bounded_linear_inner_right_comp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    82
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    83
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70065
diff changeset
    84
subsection\<^marker>\<open>tag unimportant\<close> \<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
    85
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    86
attribute_setup bounded_linear =
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    87
  \<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
    88
    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
    89
      [
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69260
diff changeset
    90
        (@{thm bounded_linear.has_derivative}, \<^named_theorems>\<open>derivative_intros\<close>),
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69260
diff changeset
    91
        (@{thm bounded_linear.tendsto}, \<^named_theorems>\<open>tendsto_intros\<close>),
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69260
diff changeset
    92
        (@{thm bounded_linear.continuous}, \<^named_theorems>\<open>continuous_intros\<close>),
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69260
diff changeset
    93
        (@{thm bounded_linear.continuous_on}, \<^named_theorems>\<open>continuous_intros\<close>),
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69260
diff changeset
    94
        (@{thm bounded_linear.uniformly_continuous_on}, \<^named_theorems>\<open>continuous_intros\<close>),
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69260
diff changeset
    95
        (@{thm bounded_linear_compose}, \<^named_theorems>\<open>bounded_linear_intros\<close>)
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    96
      ]))\<close>
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    97
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    98
attribute_setup bounded_bilinear =
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    99
  \<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
   100
    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
   101
      [
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69260
diff changeset
   102
        (@{thm bounded_bilinear.FDERIV}, \<^named_theorems>\<open>derivative_intros\<close>),
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69260
diff changeset
   103
        (@{thm bounded_bilinear.tendsto}, \<^named_theorems>\<open>tendsto_intros\<close>),
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69260
diff changeset
   104
        (@{thm bounded_bilinear.continuous}, \<^named_theorems>\<open>continuous_intros\<close>),
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69260
diff changeset
   105
        (@{thm bounded_bilinear.continuous_on}, \<^named_theorems>\<open>continuous_intros\<close>),
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_left]},
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69260
diff changeset
   107
          \<^named_theorems>\<open>bounded_linear_intros\<close>),
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   108
        (@{thm bounded_linear_compose[OF bounded_bilinear.bounded_linear_right]},
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69260
diff changeset
   109
          \<^named_theorems>\<open>bounded_linear_intros\<close>),
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_left]},
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69260
diff changeset
   111
          \<^named_theorems>\<open>continuous_intros\<close>),
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   112
        (@{thm bounded_linear.uniformly_continuous_on[OF bounded_bilinear.bounded_linear_right]},
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69260
diff changeset
   113
          \<^named_theorems>\<open>continuous_intros\<close>)
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   114
      ]))\<close>
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   115
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   116
68838
5e013478bced tagged some theories
immler
parents: 68072
diff changeset
   117
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
   118
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70065
diff changeset
   119
typedef\<^marker>\<open>tag important\<close> (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
   120
  "{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
   121
  morphisms blinfun_apply Blinfun
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   122
  by (blast intro: bounded_linear_intros)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   123
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   124
declare [[coercion
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   125
    "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
   126
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   127
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
   128
  "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
   129
  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
   130
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   131
setup_lifting type_definition_blinfun
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   132
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   133
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
   134
  by transfer auto
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   135
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   136
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
   137
  by (auto simp: Blinfun_inverse)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   138
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   139
68838
5e013478bced tagged some theories
immler
parents: 68072
diff changeset
   140
subsection \<open>Type class instantiations\<close>
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   141
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   142
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
   143
begin
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   144
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70065
diff changeset
   145
lift_definition\<^marker>\<open>tag important\<close> 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
   146
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   147
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
   148
  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
   149
  by (rule bounded_linear_sub)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   150
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   151
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
   152
  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
   153
62101
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61975
diff changeset
   154
definition [code del]:
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69064
diff changeset
   155
  "(uniformity :: (('a \<Rightarrow>\<^sub>L 'b) \<times> ('a \<Rightarrow>\<^sub>L 'b)) filter) = (INF e\<in>{0 <..}. principal {(x, y). dist x y < e})"
62101
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61975
diff changeset
   156
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   157
definition open_blinfun :: "('a \<Rightarrow>\<^sub>L 'b) set \<Rightarrow> bool"
62101
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61975
diff changeset
   158
  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
   159
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   160
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
   161
  by (rule bounded_linear_minus)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   162
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70065
diff changeset
   163
lift_definition\<^marker>\<open>tag important\<close> 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
   164
  by (rule bounded_linear_zero)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   165
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70065
diff changeset
   166
lift_definition\<^marker>\<open>tag important\<close> 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
   167
  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
   168
  by (metis bounded_linear_add)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   169
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70065
diff changeset
   170
lift_definition\<^marker>\<open>tag important\<close> 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
   171
  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
   172
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   173
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
   174
  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
   175
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   176
instance
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   177
  apply standard
62101
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61975
diff changeset
   178
  unfolding dist_blinfun_def open_blinfun_def sgn_blinfun_def uniformity_blinfun_def
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61975
diff changeset
   179
  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
   180
  done
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   181
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   182
end
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   183
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   184
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
   185
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   186
lemma norm_blinfun_eqI:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   187
  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
   188
  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
   189
  assumes "0 \<le> n"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   190
  shows "norm f = n"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   191
  by (auto simp: norm_blinfun_def
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   192
    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
   193
    bounded_linear_intros)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   194
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   195
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
   196
  by transfer (rule onorm)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   197
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   198
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
   199
  by transfer (rule onorm_bound)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   200
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   201
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
   202
proof
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   203
  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
   204
  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
   205
    by (transfer, simp)+
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   206
  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
   207
    by (auto intro!: bounded_linear_intros)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   208
  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
   209
    by (simp_all add: add scaleR)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   210
  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
   211
    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
   212
qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   213
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   214
interpretation blinfun: bounded_bilinear blinfun_apply
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   215
  by (rule bounded_bilinear_blinfun_apply)
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
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
   218
70999
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   219
declare blinfun.zero_left [simp] blinfun.zero_right [simp]
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   220
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   221
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   222
context bounded_bilinear
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   223
begin
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   224
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   225
named_theorems bilinear_simps
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   226
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   227
lemmas [bilinear_simps] =
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   228
  add_left
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   229
  add_right
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   230
  diff_left
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   231
  diff_right
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   232
  minus_left
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   233
  minus_right
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   234
  scaleR_left
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   235
  scaleR_right
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   236
  zero_left
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   237
  zero_right
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63952
diff changeset
   238
  sum_left
b9a1486e79be setsum -> sum
nipkow
parents: 63952
diff changeset
   239
  sum_right
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   240
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   241
end
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   242
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   243
66933
4e06b030730c generalized
immler
parents: 66827
diff changeset
   244
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
   245
proof
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   246
  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
   247
  assume "Cauchy X"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   248
  {
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   249
    fix x::'a
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   250
    {
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   251
      fix x::'a
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   252
      assume "norm x \<le> 1"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   253
      have "Cauchy (\<lambda>n. X n x)"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   254
      proof (rule CauchyI)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   255
        fix e::real
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   256
        assume "0 < e"
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61973
diff changeset
   257
        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
   258
          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
   259
          by auto
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   260
        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
   261
        proof (safe intro!: exI[where x=M])
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   262
          fix m n
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   263
          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
   264
          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
   265
            by (simp add: blinfun.bilinear_simps)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   266
          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
   267
             by (rule norm_blinfun)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   268
          also have "\<dots> \<le> norm (X m - X n) * 1"
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61973
diff changeset
   269
            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
   270
          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
   271
          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
   272
          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
   273
        qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   274
      qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   275
      hence "convergent (\<lambda>n. X n x)"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   276
        by (metis Cauchy_convergent_iff)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   277
    } note convergent_norm1 = this
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62963
diff changeset
   278
    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
   279
    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
   280
      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
   281
    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
   282
      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
   283
        convergent_norm1 y)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   284
    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
   285
      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
   286
    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
   287
  }
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   288
  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
   289
    unfolding convergent_def
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   290
    by metis
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   291
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   292
  have "Cauchy (\<lambda>n. norm (X n))"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   293
  proof (rule CauchyI)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   294
    fix e::real
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   295
    assume "e > 0"
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61973
diff changeset
   296
    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
   297
      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
   298
      by auto
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   299
    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
   300
    proof (safe intro!: exI[where x=M])
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   301
      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
   302
      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
   303
        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
   304
      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
   305
      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
   306
    qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   307
  qed
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   308
  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
   309
    unfolding Cauchy_convergent_iff convergent_def
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   310
    by metis
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   311
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   312
  have "bounded_linear v"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   313
  proof
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   314
    fix x y and r::real
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   315
    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
   316
      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
   317
    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
   318
      by (metis (poly_guards_query) LIMSEQ_unique)+
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   319
    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
   320
    proof (safe intro!: exI[where x=K])
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   321
      fix x
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   322
      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
   323
        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
   324
          (auto simp: norm_blinfun)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   325
      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
   326
        by (simp add: ac_simps)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   327
    qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   328
  qed
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   329
  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
   330
    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
   331
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   332
  have "X \<longlonglongrightarrow> Blinfun v"
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   333
  proof (rule LIMSEQ_I)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   334
    fix r::real assume "r > 0"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62963
diff changeset
   335
    define r' where "r' = r / 2"
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61973
diff changeset
   336
    have "0 < r'" "r' < r" using \<open>r > 0\<close> by (simp_all add: r'_def)
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61973
diff changeset
   337
    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
   338
    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
   339
      by metis
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   340
    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
   341
    proof (safe intro!: exI[where x=M])
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   342
      fix n assume n: "M \<le> n"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   343
      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
   344
      proof (rule norm_blinfun_bound)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   345
        fix x
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   346
        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
   347
          by (metis eventually_ge_at_top)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   348
        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
   349
        proof eventually_elim
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   350
          case (elim m)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   351
          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
   352
            by (simp add: blinfun.bilinear_simps)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   353
          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
   354
            by (rule norm_blinfun)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   355
          also have "\<dots> \<le> r' * norm x"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   356
            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
   357
          finally show ?case .
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   358
        qed
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   359
        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
   360
          by (auto intro!: tendsto_intros Bv)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   361
        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
   362
          by (auto intro!: tendsto_upperbound tendsto_v ev_le simp: blinfun.bilinear_simps)
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61973
diff changeset
   363
      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
   364
      thus "norm (X n - Blinfun v) < r"
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61973
diff changeset
   365
        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
   366
    qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   367
  qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   368
  thus "convergent X"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   369
    by (rule convergentI)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   370
qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   371
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70065
diff changeset
   372
subsection\<^marker>\<open>tag unimportant\<close> \<open>On Euclidean Space\<close>
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   373
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63952
diff changeset
   374
lemma Zfun_sum:
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   375
  assumes "finite s"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   376
  assumes f: "\<And>i. i \<in> s \<Longrightarrow> Zfun (f i) F"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63952
diff changeset
   377
  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
   378
  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
   379
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   380
lemma norm_blinfun_euclidean_le:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   381
  fixes a::"'a::euclidean_space \<Rightarrow>\<^sub>L 'b::real_normed_vector"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63952
diff changeset
   382
  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
   383
  apply (rule norm_blinfun_bound)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63952
diff changeset
   384
   apply (simp add: sum_nonneg)
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   385
  apply (subst euclidean_representation[symmetric, where 'a='a])
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63952
diff changeset
   386
  apply (simp only: blinfun.bilinear_simps sum_distrib_right)
b9a1486e79be setsum -> sum
nipkow
parents: 63952
diff changeset
   387
  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
   388
  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
   389
  done
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   390
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   391
lemma tendsto_componentwise1:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   392
  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
   393
    and b::"'c \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   394
  assumes "(\<And>j. j \<in> Basis \<Longrightarrow> ((\<lambda>n. b n j) \<longlongrightarrow> a j) F)"
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   395
  shows "(b \<longlongrightarrow> a) F"
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   396
proof -
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   397
  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
   398
    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
   399
  hence "Zfun (\<lambda>x. \<Sum>j\<in>Basis. norm (b x j - a j)) F"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63952
diff changeset
   400
    by (auto intro!: Zfun_sum)
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   401
  thus ?thesis
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   402
    unfolding tendsto_Zfun_iff
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   403
    by (rule Zfun_le)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   404
      (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
   405
qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   406
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   407
lift_definition
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   408
  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
   409
  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
   410
  by (intro bounded_linear_intros)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   411
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   412
lemma blinfun_of_matrix_works:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   413
  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
   414
  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
   415
proof (transfer, rule,  rule euclidean_eqI)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   416
  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
   417
  then interpret bounded_linear f by simp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   418
  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
   419
    = (\<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
   420
    using b
66804
3f9bb52082c4 avoid name clashes on interpretation of abstract locales
haftmann
parents: 66447
diff changeset
   421
    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
   422
  also have "\<dots> = (\<Sum>i\<in>Basis. (x \<bullet> i * (f i \<bullet> b)))"
71629
2e8f861d21d4 redunant simp rule
nipkow
parents: 71172
diff changeset
   423
    using b by (simp)
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   424
  also have "\<dots> = f x \<bullet> b"
63938
f6ce08859d4c More mainly topological results
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
   425
    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
   426
  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
   427
qed
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_apply:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   430
  "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
   431
  by transfer simp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   432
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   433
lemma blinfun_of_matrix_minus: "blinfun_of_matrix x - blinfun_of_matrix y = blinfun_of_matrix (x - y)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63952
diff changeset
   434
  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
   435
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   436
lemma norm_blinfun_of_matrix:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61916
diff changeset
   437
  "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
   438
  apply (rule norm_blinfun_bound)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63952
diff changeset
   439
   apply (simp add: sum_nonneg)
b9a1486e79be setsum -> sum
nipkow
parents: 63952
diff changeset
   440
  apply (simp only: blinfun_of_matrix_apply sum_distrib_right)
b9a1486e79be setsum -> sum
nipkow
parents: 63952
diff changeset
   441
  apply (rule order_trans[OF norm_sum sum_mono])
b9a1486e79be setsum -> sum
nipkow
parents: 63952
diff changeset
   442
  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
   443
  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
   444
  done
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   445
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   446
lemma tendsto_blinfun_of_matrix:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   447
  assumes "\<And>i j. i \<in> Basis \<Longrightarrow> j \<in> Basis \<Longrightarrow> ((\<lambda>n. b n i j) \<longlongrightarrow> a i j) F"
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   448
  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
   449
proof -
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   450
  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
   451
    using assms unfolding tendsto_Zfun_iff Zfun_norm_iff .
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61916
diff changeset
   452
  hence "Zfun (\<lambda>x. (\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. \<bar>b x i j - a i j\<bar>)) F"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63952
diff changeset
   453
    by (auto intro!: Zfun_sum)
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   454
  thus ?thesis
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   455
    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
   456
    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
   457
qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   458
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   459
lemma tendsto_componentwise:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   460
  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
   461
    and b::"'c \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   462
  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
   463
  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
   464
  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
   465
  by (rule tendsto_blinfun_of_matrix)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   466
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   467
lemma
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   468
  continuous_blinfun_componentwiseI:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   469
  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
   470
  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
   471
  shows "continuous F f"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   472
  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
   473
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   474
lemma
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   475
  continuous_blinfun_componentwiseI1:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   476
  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
   477
  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
   478
  shows "continuous F f"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   479
  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
   480
67685
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67399
diff changeset
   481
lemma
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67399
diff changeset
   482
  continuous_on_blinfun_componentwise:
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67399
diff changeset
   483
  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
   484
  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
   485
  shows "continuous_on s f"
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67399
diff changeset
   486
  using assms
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67399
diff changeset
   487
  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
   488
    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
   489
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   490
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
   491
  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
   492
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   493
lemma continuous_blinfun_matrix:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   494
  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
   495
  assumes "continuous F f"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   496
  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
   497
  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
   498
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   499
lemma continuous_on_blinfun_matrix:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   500
  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
   501
  assumes "continuous_on S f"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   502
  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
   503
  using assms
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   504
  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
   505
62963
2d5eff9c3baa added rule
immler
parents: 62951
diff changeset
   506
lemma continuous_on_blinfun_of_matrix[continuous_intros]:
2d5eff9c3baa added rule
immler
parents: 62951
diff changeset
   507
  assumes "\<And>i j. i \<in> Basis \<Longrightarrow> j \<in> Basis \<Longrightarrow> continuous_on S (\<lambda>s. g s i j)"
2d5eff9c3baa added rule
immler
parents: 62951
diff changeset
   508
  shows "continuous_on S (\<lambda>s. blinfun_of_matrix (g s))"
2d5eff9c3baa added rule
immler
parents: 62951
diff changeset
   509
  using assms
2d5eff9c3baa added rule
immler
parents: 62951
diff changeset
   510
  by (auto simp: continuous_on intro!: tendsto_blinfun_of_matrix)
2d5eff9c3baa added rule
immler
parents: 62951
diff changeset
   511
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   512
lemma mult_if_delta:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   513
  "(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
   514
  by auto
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   515
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   516
lemma compact_blinfun_lemma:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   517
  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
   518
  assumes "bounded (range f)"
66447
a1f5c5c26fa6 Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents: 66089
diff changeset
   519
  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
   520
    strict_mono r \<and> (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) i) (l i) < e) sequentially)"
62127
d8e7738bd2e9 generalized proofs
immler
parents: 62102
diff changeset
   521
  by (rule compact_lemma_general[where unproj = "\<lambda>e. blinfun_of_matrix (\<lambda>i j. e j \<bullet> i)"])
d8e7738bd2e9 generalized proofs
immler
parents: 62102
diff changeset
   522
   (auto intro!: euclidean_eqI[where 'a='b] bounded_linear_image assms
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63952
diff changeset
   523
    simp: blinfun_of_matrix_works blinfun_of_matrix_apply inner_Basis mult_if_delta sum.delta'
b9a1486e79be setsum -> sum
nipkow
parents: 63952
diff changeset
   524
      scaleR_sum_left[symmetric])
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   525
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   526
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
   527
  apply (auto intro!: blinfun_eqI)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   528
  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
   529
  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
   530
  apply (simp add: blinfun.bilinear_simps)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   531
  done
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   532
62951
f59ef58f420b added lemmas
immler
parents: 62390
diff changeset
   533
lemma Blinfun_eq_matrix: "bounded_linear f \<Longrightarrow> Blinfun f = blinfun_of_matrix (\<lambda>i j. f j \<bullet> i)"
f59ef58f420b added lemmas
immler
parents: 62390
diff changeset
   534
  by (intro blinfun_euclidean_eqI)
f59ef58f420b added lemmas
immler
parents: 62390
diff changeset
   535
     (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
   536
      if_distribR sum.delta' euclidean_representation
62951
f59ef58f420b added lemmas
immler
parents: 62390
diff changeset
   537
      cong: if_cong)
f59ef58f420b added lemmas
immler
parents: 62390
diff changeset
   538
67226
ec32cdaab97b isabelle update_cartouches -c -t;
wenzelm
parents: 66933
diff changeset
   539
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
   540
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
   541
proof
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   542
  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
   543
  assume f: "bounded (range f)"
66447
a1f5c5c26fa6 Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents: 66089
diff changeset
   544
  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
   545
    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
   546
    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
   547
  {
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   548
    fix e::real
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   549
    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
   550
    assume "e > 0"
71172
nipkow
parents: 70999
diff changeset
   551
    hence "e / ?d > 0" by (simp)
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   552
    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
   553
      by simp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   554
    moreover
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   555
    {
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   556
      fix n
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   557
      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
   558
      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
   559
        unfolding blinfun_of_matrix_works ..
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   560
      also note norm_blinfun_of_matrix
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   561
      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
   562
        (\<Sum>i\<in>(Basis::'b set). e / real_of_nat DIM('b))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63952
diff changeset
   563
      proof (rule sum_strict_mono)
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   564
        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
   565
        have "(\<Sum>j::'a\<in>Basis. \<bar>(f (r n) - l) j \<bullet> i\<bar>) < (\<Sum>j::'a\<in>Basis. e / ?d)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63952
diff changeset
   566
        proof (rule sum_strict_mono)
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   567
          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
   568
          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
   569
            by (simp add: Basis_le_norm i)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   570
          also have "\<dots> < e / ?d"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   571
            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
   572
          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
   573
        qed simp_all
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   574
        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
   575
          by simp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   576
        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
   577
          by simp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   578
      qed simp_all
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   579
      also have "\<dots> \<le> e" by simp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   580
      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
   581
        by (auto simp: dist_norm)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   582
    }
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   583
    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
   584
      using eventually_elim2 by force
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   585
  }
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   586
  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
   587
    unfolding o_def tendsto_iff by simp
66447
a1f5c5c26fa6 Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents: 66089
diff changeset
   588
  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
   589
    by auto
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   590
qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   591
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   592
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70065
diff changeset
   593
subsection\<^marker>\<open>tag unimportant\<close> \<open>concrete bounded linear functions\<close>
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   594
61916
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   595
lemma transfer_bounded_bilinear_bounded_linearI:
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   596
  assumes "g = (\<lambda>i x. (blinfun_apply (f i) x))"
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   597
  shows "bounded_bilinear g = bounded_linear f"
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   598
proof
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   599
  assume "bounded_bilinear g"
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   600
  then interpret bounded_bilinear f by (simp add: assms)
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   601
  show "bounded_linear f"
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   602
  proof (unfold_locales, safe intro!: blinfun_eqI)
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   603
    fix i
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   604
    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
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   605
      by (auto intro!: blinfun_eqI simp: blinfun.bilinear_simps)
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   606
    from _ nonneg_bounded show "\<exists>K. \<forall>x. norm (f x) \<le> norm x * K"
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   607
      by (rule ex_reg) (auto intro!: onorm_bound simp: norm_blinfun.rep_eq ac_simps)
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   608
  qed
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   609
qed (auto simp: assms intro!: blinfun.comp)
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   610
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   611
lemma transfer_bounded_bilinear_bounded_linear[transfer_rule]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67226
diff changeset
   612
  "(rel_fun (rel_fun (=) (pcr_blinfun (=) (=))) (=)) bounded_bilinear bounded_linear"
61916
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   613
  by (auto simp: pcr_blinfun_def cr_blinfun_def rel_fun_def OO_def
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   614
    intro!: transfer_bounded_bilinear_bounded_linearI)
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   615
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   616
context bounded_bilinear
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   617
begin
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
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
   620
  by (rule bounded_linear_left)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   621
declare prod_left.rep_eq[simp]
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   622
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   623
lemma bounded_linear_prod_left[bounded_linear]: "bounded_linear prod_left"
61916
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   624
  by transfer (rule flip)
61915
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
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
   627
  by (rule bounded_linear_right)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   628
declare prod_right.rep_eq[simp]
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   629
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   630
lemma bounded_linear_prod_right[bounded_linear]: "bounded_linear prod_right"
61916
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   631
  by transfer (rule bounded_bilinear_axioms)
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   632
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   633
end
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   634
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   635
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
   636
  by (rule bounded_linear_ident)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   637
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   638
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
   639
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   640
lemma norm_blinfun_id[simp]:
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, perfect_space} \<Rightarrow>\<^sub>L 'a) = 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)
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
lemma norm_blinfun_id_le:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   645
  "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
   646
  by transfer (auto simp: onorm_id_le)
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
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   649
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
   650
  by (rule bounded_linear_fst)
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
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
   653
  by transfer (rule refl)
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
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   656
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
   657
  by (rule bounded_linear_snd)
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
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
   660
  by transfer (rule refl)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   661
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   662
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   663
lift_definition blinfun_compose::
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   664
  "'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
   665
    'c::real_normed_vector \<Rightarrow>\<^sub>L 'a \<Rightarrow>
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67226
diff changeset
   666
    '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
   667
  parametric comp_transfer
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   668
  unfolding o_def
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   669
  by (rule bounded_linear_compose)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   670
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   671
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
   672
  by (simp add: blinfun_compose.rep_eq)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   673
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   674
lemma norm_blinfun_compose:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   675
  "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
   676
  by transfer (rule onorm_compose)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   677
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67226
diff changeset
   678
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
   679
  by unfold_locales
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   680
    (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
   681
62951
f59ef58f420b added lemmas
immler
parents: 62390
diff changeset
   682
lemma blinfun_compose_zero[simp]:
f59ef58f420b added lemmas
immler
parents: 62390
diff changeset
   683
  "blinfun_compose 0 = (\<lambda>_. 0)"
f59ef58f420b added lemmas
immler
parents: 62390
diff changeset
   684
  "blinfun_compose x 0 = 0"
f59ef58f420b added lemmas
immler
parents: 62390
diff changeset
   685
  by (auto simp: blinfun.bilinear_simps intro!: blinfun_eqI)
f59ef58f420b added lemmas
immler
parents: 62390
diff changeset
   686
70999
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   687
lemma blinfun_bij2:
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   688
  fixes f::"'a \<Rightarrow>\<^sub>L 'a::euclidean_space"
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   689
  assumes "f o\<^sub>L g = id_blinfun"
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   690
  shows "bij (blinfun_apply g)"
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   691
proof (rule bijI)
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   692
  show "inj g"
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   693
    using assms
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   694
    by (metis blinfun_apply_id_blinfun blinfun_compose.rep_eq injI inj_on_imageI2)
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   695
  then show "surj g"
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   696
    using blinfun.bounded_linear_right bounded_linear_def linear_inj_imp_surj by blast
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   697
qed
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   698
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   699
lemma blinfun_bij1:
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   700
  fixes f::"'a \<Rightarrow>\<^sub>L 'a::euclidean_space"
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   701
  assumes "f o\<^sub>L g = id_blinfun"
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   702
  shows "bij (blinfun_apply f)"
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   703
proof (rule bijI)
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   704
  show "surj (blinfun_apply f)"
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   705
    by (metis assms blinfun_apply_blinfun_compose blinfun_apply_id_blinfun surjI)
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   706
  then show "inj (blinfun_apply f)"
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   707
    using blinfun.bounded_linear_right bounded_linear_def linear_surj_imp_inj by blast
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   708
qed
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   709
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67226
diff changeset
   710
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
   711
  by (rule bounded_linear_inner_right)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   712
declare blinfun_inner_right.rep_eq[simp]
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   713
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   714
lemma bounded_linear_blinfun_inner_right[bounded_linear]: "bounded_linear blinfun_inner_right"
61916
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   715
  by transfer (rule bounded_bilinear_inner)
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   716
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   717
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   718
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
   719
  by (rule bounded_linear_inner_left)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   720
declare blinfun_inner_left.rep_eq[simp]
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   721
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   722
lemma bounded_linear_blinfun_inner_left[bounded_linear]: "bounded_linear blinfun_inner_left"
61916
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   723
  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
   724
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   725
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68838
diff changeset
   726
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
   727
  by (rule bounded_linear_scaleR_right)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   728
declare blinfun_scaleR_right.rep_eq[simp]
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   729
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   730
lemma bounded_linear_blinfun_scaleR_right[bounded_linear]: "bounded_linear blinfun_scaleR_right"
61916
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   731
  by transfer (rule bounded_bilinear_scaleR)
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   732
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   733
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   734
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
   735
  by (rule bounded_linear_scaleR_left)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   736
lemmas [simp] = blinfun_scaleR_left.rep_eq
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   737
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   738
lemma bounded_linear_blinfun_scaleR_left[bounded_linear]: "bounded_linear blinfun_scaleR_left"
61916
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   739
  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
   740
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   741
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68838
diff changeset
   742
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
   743
  by (rule bounded_linear_mult_right)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   744
declare blinfun_mult_right.rep_eq[simp]
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   745
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   746
lemma bounded_linear_blinfun_mult_right[bounded_linear]: "bounded_linear blinfun_mult_right"
61916
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   747
  by transfer (rule bounded_bilinear_mult)
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   748
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   749
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   750
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
   751
  by (rule bounded_linear_mult_left)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   752
lemmas [simp] = blinfun_mult_left.rep_eq
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   753
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   754
lemma bounded_linear_blinfun_mult_left[bounded_linear]: "bounded_linear blinfun_mult_left"
61916
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   755
  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
   756
67685
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67399
diff changeset
   757
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
   758
  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
   759
  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
   760
  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
   761
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   762
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   763
subsection \<open>The strong operator topology on continuous linear operators\<close>
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   764
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   765
text \<open>Let \<open>'a\<close> and \<open>'b\<close> be two normed real vector spaces. Then the space of linear continuous
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   766
operators from \<open>'a\<close> to \<open>'b\<close> has a canonical norm, and therefore a canonical corresponding topology
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   767
(the type classes instantiation are given in \<^file>\<open>Bounded_Linear_Function.thy\<close>).
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   768
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   769
However, there is another topology on this space, the strong operator topology, where \<open>T\<^sub>n\<close> tends to
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   770
\<open>T\<close> iff, for all \<open>x\<close> in \<open>'a\<close>, then \<open>T\<^sub>n x\<close> tends to \<open>T x\<close>. This is precisely the product topology
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   771
where the target space is endowed with the norm topology. It is especially useful when \<open>'b\<close> is the set
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   772
of real numbers, since then this topology is compact.
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   773
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   774
We can not implement it using type classes as there is already a topology, but at least we
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   775
can define it as a topology.
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   776
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   777
Note that there is yet another (common and useful) topology on operator spaces, the weak operator
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   778
topology, defined analogously using the product topology, but where the target space is given the
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   779
weak-* topology, i.e., the pullback of the weak topology on the bidual of the space under the
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   780
canonical embedding of a space into its bidual. We do not define it there, although it could also be
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   781
defined analogously.
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   782
\<close>
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   783
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70065
diff changeset
   784
definition\<^marker>\<open>tag important\<close> strong_operator_topology::"('a::real_normed_vector \<Rightarrow>\<^sub>L'b::real_normed_vector) topology"
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   785
where "strong_operator_topology = pullback_topology UNIV blinfun_apply euclidean"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   786
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   787
lemma strong_operator_topology_topspace:
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   788
  "topspace strong_operator_topology = UNIV"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   789
unfolding strong_operator_topology_def topspace_pullback_topology topspace_euclidean by auto
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   790
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   791
lemma strong_operator_topology_basis:
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   792
  fixes f::"('a::real_normed_vector \<Rightarrow>\<^sub>L'b::real_normed_vector)" and U::"'i \<Rightarrow> 'b set" and x::"'i \<Rightarrow> 'a"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   793
  assumes "finite I" "\<And>i. i \<in> I \<Longrightarrow> open (U i)"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   794
  shows "openin strong_operator_topology {f. \<forall>i\<in>I. blinfun_apply f (x i) \<in> U i}"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   795
proof -
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   796
  have "open {g::('a\<Rightarrow>'b). \<forall>i\<in>I. g (x i) \<in> U i}"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   797
    by (rule product_topology_basis'[OF assms])
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   798
  moreover have "{f. \<forall>i\<in>I. blinfun_apply f (x i) \<in> U i}
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   799
                = blinfun_apply-`{g::('a\<Rightarrow>'b). \<forall>i\<in>I. g (x i) \<in> U i} \<inter> UNIV"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   800
    by auto
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   801
  ultimately show ?thesis
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   802
    unfolding strong_operator_topology_def by (subst openin_pullback_topology) auto
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   803
qed
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   804
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   805
lemma strong_operator_topology_continuous_evaluation:
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   806
  "continuous_map strong_operator_topology euclidean (\<lambda>f. blinfun_apply f x)"
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   807
proof -
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   808
  have "continuous_map strong_operator_topology euclidean ((\<lambda>f. f x) o blinfun_apply)"
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   809
    unfolding strong_operator_topology_def apply (rule continuous_map_pullback)
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   810
    using continuous_on_product_coordinates by fastforce
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   811
  then show ?thesis unfolding comp_def by simp
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   812
qed
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   813
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   814
lemma continuous_on_strong_operator_topo_iff_coordinatewise:
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   815
  "continuous_map T strong_operator_topology f
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   816
    \<longleftrightarrow> (\<forall>x. continuous_map T euclidean (\<lambda>y. blinfun_apply (f y) x))"
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   817
proof (auto)
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   818
  fix x::"'b"
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   819
  assume "continuous_map T strong_operator_topology f"
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   820
  with continuous_map_compose[OF this strong_operator_topology_continuous_evaluation]
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   821
  have "continuous_map T euclidean ((\<lambda>z. blinfun_apply z x) o f)"
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   822
    by simp
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   823
  then show "continuous_map T euclidean (\<lambda>y. blinfun_apply (f y) x)"
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   824
    unfolding comp_def by auto
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   825
next
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   826
  assume *: "\<forall>x. continuous_map T euclidean (\<lambda>y. blinfun_apply (f y) x)"
70065
cc89a395b5a3 Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
   827
  have "\<And>i. continuous_map T euclidean (\<lambda>x. blinfun_apply (f x) i)"
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   828
    using * unfolding comp_def by auto
70065
cc89a395b5a3 Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
   829
  then have "continuous_map T euclidean (blinfun_apply o f)"
cc89a395b5a3 Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
   830
    unfolding o_def
cc89a395b5a3 Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
   831
    by (metis (no_types) continuous_map_componentwise_UNIV euclidean_product_topology)
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   832
  show "continuous_map T strong_operator_topology f"
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   833
    unfolding strong_operator_topology_def
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   834
    apply (rule continuous_map_pullback')
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   835
    by (auto simp add: \<open>continuous_map T euclidean (blinfun_apply o f)\<close>)
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   836
qed
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   837
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   838
lemma strong_operator_topology_weaker_than_euclidean:
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   839
  "continuous_map euclidean strong_operator_topology (\<lambda>f. f)"
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   840
  by (subst continuous_on_strong_operator_topo_iff_coordinatewise,
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   841
    auto simp add: linear_continuous_on)
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   842
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   843
end