src/HOL/Multivariate_Analysis/Bounded_Linear_Function.thy
author wenzelm
Mon, 28 Dec 2015 01:28:28 +0100
changeset 61945 1135b8de26c3
parent 61916 7950ae6d3266
child 61969 e01015e49041
permissions -rw-r--r--
more symbols;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
     1
(*  Title:      HOL/Multivariate_Analysis/Bounded_Linear_Function.thy
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
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
     5
section {* Bounded Linear Function *}
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
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    11
begin
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    12
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    13
subsection {* Intro rules for @{term bounded_linear} *}
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    14
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    15
named_theorems bounded_linear_intros
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    16
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    17
lemma onorm_inner_left:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    18
  assumes "bounded_linear r"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    19
  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
    20
proof (rule onorm_bound)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    21
  fix x
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    22
  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
    23
    by (simp add: Cauchy_Schwarz_ineq2)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    24
  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
    25
    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
    26
  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
    27
    by (simp add: ac_simps)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    28
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
    29
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    30
lemma onorm_inner_right:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    31
  assumes "bounded_linear r"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    32
  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
    33
  apply (subst inner_commute)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    34
  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
    35
  apply simp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    36
  done
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    37
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    38
lemmas [bounded_linear_intros] =
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    39
  bounded_linear_zero
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    40
  bounded_linear_add
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    41
  bounded_linear_const_mult
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    42
  bounded_linear_mult_const
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    43
  bounded_linear_scaleR_const
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    44
  bounded_linear_const_scaleR
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    45
  bounded_linear_ident
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    46
  bounded_linear_setsum
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    47
  bounded_linear_Pair
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    48
  bounded_linear_sub
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    49
  bounded_linear_fst_comp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    50
  bounded_linear_snd_comp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    51
  bounded_linear_inner_left_comp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    52
  bounded_linear_inner_right_comp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    53
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    54
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    55
subsection \<open>declaration of derivative/continuous/tendsto introduction rules for bounded linear functions\<close>
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    56
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    57
attribute_setup bounded_linear =
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    58
  \<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
    59
    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
    60
      [
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    61
        (@{thm bounded_linear.has_derivative}, "Deriv.derivative_intros"),
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    62
        (@{thm bounded_linear.tendsto}, "Topological_Spaces.tendsto_intros"),
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    63
        (@{thm bounded_linear.continuous}, "Topological_Spaces.continuous_intros"),
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    64
        (@{thm bounded_linear.continuous_on}, "Topological_Spaces.continuous_intros"),
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    65
        (@{thm bounded_linear.uniformly_continuous_on}, "Topological_Spaces.continuous_intros"),
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    66
        (@{thm bounded_linear_compose}, "Bounded_Linear_Function.bounded_linear_intros")
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    67
      ]))\<close>
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    68
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    69
attribute_setup bounded_bilinear =
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    70
  \<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
    71
    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
    72
      [
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    73
        (@{thm bounded_bilinear.FDERIV}, "Deriv.derivative_intros"),
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    74
        (@{thm bounded_bilinear.tendsto}, "Topological_Spaces.tendsto_intros"),
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    75
        (@{thm bounded_bilinear.continuous}, "Topological_Spaces.continuous_intros"),
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    76
        (@{thm bounded_bilinear.continuous_on}, "Topological_Spaces.continuous_intros"),
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    77
        (@{thm bounded_linear_compose[OF bounded_bilinear.bounded_linear_left]},
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    78
          "Bounded_Linear_Function.bounded_linear_intros"),
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    79
        (@{thm bounded_linear_compose[OF bounded_bilinear.bounded_linear_right]},
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    80
          "Bounded_Linear_Function.bounded_linear_intros"),
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    81
        (@{thm bounded_linear.uniformly_continuous_on[OF bounded_bilinear.bounded_linear_left]},
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    82
          "Topological_Spaces.continuous_intros"),
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    83
        (@{thm bounded_linear.uniformly_continuous_on[OF bounded_bilinear.bounded_linear_right]},
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    84
          "Topological_Spaces.continuous_intros")
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    85
      ]))\<close>
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    86
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    87
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    88
subsection \<open>type of bounded linear functions\<close>
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    89
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    90
typedef (overloaded) ('a, 'b) blinfun ("(_ \<Rightarrow>\<^sub>L /_)" [22, 21] 21) =
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    91
  "{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
    92
  morphisms blinfun_apply Blinfun
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    93
  by (blast intro: bounded_linear_intros)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    94
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    95
declare [[coercion
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    96
    "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
    97
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
    98
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
    99
  "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
   100
  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
   101
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   102
setup_lifting type_definition_blinfun
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   103
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   104
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
   105
  by transfer auto
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   106
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   107
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
   108
  by (auto simp: Blinfun_inverse)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   109
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   110
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   111
subsection \<open>type class instantiations\<close>
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   112
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   113
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
   114
begin
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
lift_definition norm_blinfun :: "'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> real" is onorm .
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   117
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   118
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
   119
  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
   120
  by (rule bounded_linear_sub)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   121
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   122
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
   123
  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
   124
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   125
definition open_blinfun :: "('a \<Rightarrow>\<^sub>L 'b) set \<Rightarrow> bool"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   126
  where "open_blinfun S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   127
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   128
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
   129
  by (rule bounded_linear_minus)
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
lift_definition zero_blinfun :: "'a \<Rightarrow>\<^sub>L 'b" is "\<lambda>x. 0"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   132
  by (rule bounded_linear_zero)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   133
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   134
lift_definition plus_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
   135
  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
   136
  by (metis bounded_linear_add)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   137
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   138
lift_definition 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"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   139
  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
   140
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   141
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
   142
  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
   143
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   144
instance
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   145
  apply standard
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   146
  unfolding dist_blinfun_def open_blinfun_def sgn_blinfun_def
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   147
  apply (transfer, force simp: onorm_triangle onorm_scaleR onorm_eq_0 algebra_simps)+
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   148
  done
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   149
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   150
end
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   151
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   152
lemma norm_blinfun_eqI:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   153
  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
   154
  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
   155
  assumes "0 \<le> n"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   156
  shows "norm f = n"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   157
  by (auto simp: norm_blinfun_def
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   158
    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
   159
    bounded_linear_intros)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   160
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   161
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
   162
  by transfer (rule onorm)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   163
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   164
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
   165
  by transfer (rule onorm_bound)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   166
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   167
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
   168
proof
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   169
  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
   170
  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
   171
    by (transfer, simp)+
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   172
  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
   173
    by (auto intro!: bounded_linear_intros)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   174
  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
   175
    by (simp_all add: add scaleR)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   176
  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
   177
    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
   178
qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   179
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   180
interpretation blinfun: bounded_bilinear blinfun_apply
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   181
  by (rule bounded_bilinear_blinfun_apply)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   182
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   183
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
   184
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   185
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   186
context bounded_bilinear
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   187
begin
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   188
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   189
named_theorems bilinear_simps
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   190
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   191
lemmas [bilinear_simps] =
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   192
  add_left
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   193
  add_right
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   194
  diff_left
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   195
  diff_right
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   196
  minus_left
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   197
  minus_right
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   198
  scaleR_left
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   199
  scaleR_right
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   200
  zero_left
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   201
  zero_right
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   202
  setsum_left
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   203
  setsum_right
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   204
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   205
end
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   206
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   207
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   208
instance blinfun :: (banach, banach) banach
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   209
proof
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   210
  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
   211
  assume "Cauchy X"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   212
  {
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   213
    fix x::'a
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   214
    {
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   215
      fix x::'a
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   216
      assume "norm x \<le> 1"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   217
      have "Cauchy (\<lambda>n. X n x)"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   218
      proof (rule CauchyI)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   219
        fix e::real
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   220
        assume "0 < e"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   221
        from CauchyD[OF `Cauchy X` `0 < e`] obtain M
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   222
          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
   223
          by auto
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   224
        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
   225
        proof (safe intro!: exI[where x=M])
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   226
          fix m n
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   227
          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
   228
          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
   229
            by (simp add: blinfun.bilinear_simps)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   230
          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
   231
             by (rule norm_blinfun)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   232
          also have "\<dots> \<le> norm (X m - X n) * 1"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   233
            using `norm x \<le> 1` norm_ge_zero by (rule mult_left_mono)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   234
          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
   235
          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
   236
          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
   237
        qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   238
      qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   239
      hence "convergent (\<lambda>n. X n x)"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   240
        by (metis Cauchy_convergent_iff)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   241
    } note convergent_norm1 = this
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   242
    def y \<equiv> "x /\<^sub>R norm x"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   243
    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
   244
      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
   245
    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
   246
      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
   247
        convergent_norm1 y)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   248
    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
   249
      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
   250
    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
   251
  }
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   252
  then obtain v where v: "\<And>x. (\<lambda>n. X n x) ----> v x"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   253
    unfolding convergent_def
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   254
    by metis
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   255
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   256
  have "Cauchy (\<lambda>n. norm (X n))"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   257
  proof (rule CauchyI)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   258
    fix e::real
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   259
    assume "e > 0"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   260
    from CauchyD[OF `Cauchy X` `0 < e`] obtain M
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   261
      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
   262
      by auto
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   263
    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
   264
    proof (safe intro!: exI[where x=M])
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   265
      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
   266
      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
   267
        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
   268
      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
   269
      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
   270
    qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   271
  qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   272
  then obtain K where K: "(\<lambda>n. norm (X n)) ----> K"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   273
    unfolding Cauchy_convergent_iff convergent_def
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   274
    by metis
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   275
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   276
  have "bounded_linear v"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   277
  proof
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   278
    fix x y and r::real
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   279
    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
   280
      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
   281
    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
   282
      by (metis (poly_guards_query) LIMSEQ_unique)+
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   283
    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
   284
    proof (safe intro!: exI[where x=K])
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   285
      fix x
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   286
      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
   287
        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
   288
          (auto simp: norm_blinfun)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   289
      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
   290
        by (simp add: ac_simps)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   291
    qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   292
  qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   293
  hence Bv: "\<And>x. (\<lambda>n. X n x) ----> Blinfun v x"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   294
    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
   295
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   296
  have "X ----> Blinfun v"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   297
  proof (rule LIMSEQ_I)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   298
    fix r::real assume "r > 0"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   299
    def r' \<equiv> "r / 2"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   300
    have "0 < r'" "r' < r" using `r > 0` by (simp_all add: r'_def)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   301
    from CauchyD[OF `Cauchy X` `r' > 0`]
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   302
    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
   303
      by metis
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   304
    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
   305
    proof (safe intro!: exI[where x=M])
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   306
      fix n assume n: "M \<le> n"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   307
      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
   308
      proof (rule norm_blinfun_bound)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   309
        fix x
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   310
        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
   311
          by (metis eventually_ge_at_top)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   312
        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
   313
        proof eventually_elim
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   314
          case (elim m)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   315
          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
   316
            by (simp add: blinfun.bilinear_simps)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   317
          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
   318
            by (rule norm_blinfun)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   319
          also have "\<dots> \<le> r' * norm x"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   320
            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
   321
          finally show ?case .
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   322
        qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   323
        have tendsto_v: "(\<lambda>m. norm (X n x - X m x)) ----> norm (X n x - Blinfun v x)"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   324
          by (auto intro!: tendsto_intros Bv)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   325
        show "norm ((X n - Blinfun v) x) \<le> r' * norm x"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   326
          by (auto intro!: tendsto_ge_const tendsto_v ev_le simp: blinfun.bilinear_simps)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   327
      qed (simp add: `0 < r'` less_imp_le)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   328
      thus "norm (X n - Blinfun v) < r"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   329
        by (metis `r' < r` le_less_trans)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   330
    qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   331
  qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   332
  thus "convergent X"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   333
    by (rule convergentI)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   334
qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   335
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   336
subsection {* On Euclidean Space *}
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   337
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   338
lemma Zfun_setsum:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   339
  assumes "finite s"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   340
  assumes f: "\<And>i. i \<in> s \<Longrightarrow> Zfun (f i) F"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   341
  shows "Zfun (\<lambda>x. setsum (\<lambda>i. f i x) s) F"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   342
  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
   343
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   344
lemma norm_blinfun_euclidean_le:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   345
  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
   346
  shows "norm a \<le> setsum (\<lambda>x. norm (a x)) Basis"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   347
  apply (rule norm_blinfun_bound)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   348
   apply (simp add: setsum_nonneg)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   349
  apply (subst euclidean_representation[symmetric, where 'a='a])
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   350
  apply (simp only: blinfun.bilinear_simps setsum_left_distrib)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   351
  apply (rule order.trans[OF norm_setsum setsum_mono])
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   352
  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
   353
  done
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   354
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   355
lemma tendsto_componentwise1:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   356
  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
   357
    and b::"'c \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   358
  assumes "(\<And>j. j \<in> Basis \<Longrightarrow> ((\<lambda>n. b n j) ---> a j) F)"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   359
  shows "(b ---> a) F"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   360
proof -
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   361
  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
   362
    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
   363
  hence "Zfun (\<lambda>x. \<Sum>j\<in>Basis. norm (b x j - a j)) F"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   364
    by (auto intro!: Zfun_setsum)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   365
  thus ?thesis
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   366
    unfolding tendsto_Zfun_iff
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   367
    by (rule Zfun_le)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   368
      (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
   369
qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   370
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   371
lift_definition
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   372
  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
   373
  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
   374
  by (intro bounded_linear_intros)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   375
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   376
lemma blinfun_of_matrix_works:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   377
  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
   378
  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
   379
proof (transfer, rule,  rule euclidean_eqI)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   380
  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
   381
  then interpret bounded_linear f by simp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   382
  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
   383
    = (\<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
   384
    using b
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   385
    by (auto simp add: algebra_simps inner_setsum_left inner_Basis split: split_if intro!: setsum.cong)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   386
  also have "\<dots> = (\<Sum>i\<in>Basis. (x \<bullet> i * (f i \<bullet> b)))"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   387
    using b by (simp add: setsum.delta)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   388
  also have "\<dots> = f x \<bullet> b"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   389
    by (subst linear_componentwise[symmetric]) (unfold_locales, rule)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   390
  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
   391
qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   392
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   393
lemma blinfun_of_matrix_apply:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   394
  "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
   395
  by transfer simp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   396
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   397
lemma blinfun_of_matrix_minus: "blinfun_of_matrix x - blinfun_of_matrix y = blinfun_of_matrix (x - y)"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   398
  by transfer (auto simp: algebra_simps setsum_subtractf)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   399
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   400
lemma norm_blinfun_of_matrix:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61916
diff changeset
   401
  "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
   402
  apply (rule norm_blinfun_bound)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   403
   apply (simp add: setsum_nonneg)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   404
  apply (simp only: blinfun_of_matrix_apply setsum_left_distrib)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   405
  apply (rule order_trans[OF norm_setsum setsum_mono])
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   406
  apply (rule order_trans[OF norm_setsum setsum_mono])
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   407
  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
   408
  done
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   409
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   410
lemma tendsto_blinfun_of_matrix:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   411
  assumes "\<And>i j. i \<in> Basis \<Longrightarrow> j \<in> Basis \<Longrightarrow> ((\<lambda>n. b n i j) ---> a i j) F"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   412
  shows "((\<lambda>n. blinfun_of_matrix (b n)) ---> blinfun_of_matrix a) F"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   413
proof -
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   414
  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
   415
    using assms unfolding tendsto_Zfun_iff Zfun_norm_iff .
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61916
diff changeset
   416
  hence "Zfun (\<lambda>x. (\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. \<bar>b x i j - a i j\<bar>)) F"
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   417
    by (auto intro!: Zfun_setsum)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   418
  thus ?thesis
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   419
    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
   420
    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
   421
qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   422
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   423
lemma tendsto_componentwise:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   424
  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
   425
    and b::"'c \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   426
  shows "(\<And>i j. i \<in> Basis \<Longrightarrow> j \<in> Basis \<Longrightarrow> ((\<lambda>n. b n j \<bullet> i) ---> a j \<bullet> i) F) \<Longrightarrow> (b ---> a) F"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   427
  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
   428
  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
   429
  by (rule tendsto_blinfun_of_matrix)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   430
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   431
lemma
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   432
  continuous_blinfun_componentwiseI:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   433
  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
   434
  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
   435
  shows "continuous F f"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   436
  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
   437
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   438
lemma
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   439
  continuous_blinfun_componentwiseI1:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   440
  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
   441
  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
   442
  shows "continuous F f"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   443
  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
   444
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   445
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
   446
  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
   447
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   448
lemma continuous_blinfun_matrix:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   449
  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
   450
  assumes "continuous F f"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   451
  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
   452
  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
   453
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   454
lemma continuous_on_blinfun_matrix:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   455
  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
   456
  assumes "continuous_on S f"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   457
  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
   458
  using assms
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   459
  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
   460
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   461
lemma mult_if_delta:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   462
  "(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
   463
  by auto
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   464
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   465
text {* TODO: generalize this and @{thm compact_lemma}?! *}
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   466
lemma compact_blinfun_lemma:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   467
  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
   468
  assumes "bounded (range f)"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   469
  shows "\<forall>d\<subseteq>Basis. \<exists>l::'a \<Rightarrow>\<^sub>L 'b. \<exists> r.
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   470
    subseq r \<and> (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. 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
   471
proof safe
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   472
  fix d :: "'a set"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   473
  assume d: "d \<subseteq> Basis"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   474
  with finite_Basis have "finite d"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   475
    by (blast intro: finite_subset)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   476
  from this d show "\<exists>l::'a \<Rightarrow>\<^sub>L 'b. \<exists>r. subseq r \<and>
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   477
    (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. 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
   478
  proof (induct d)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   479
    case empty
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   480
    then show ?case
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   481
      unfolding subseq_def by auto
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   482
  next
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   483
    case (insert k d)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   484
    have k[intro]: "k \<in> Basis"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   485
      using insert by auto
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   486
    have s': "bounded ((\<lambda>x. blinfun_apply x k) ` range f)"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   487
      using `bounded (range f)`
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   488
      by (auto intro!: bounded_linear_image bounded_linear_intros)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   489
    obtain l1::"'a \<Rightarrow>\<^sub>L 'b" and r1 where r1: "subseq r1"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   490
      and lr1: "\<forall>e > 0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) i) (l1 i) < e) sequentially"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   491
      using insert(3) using insert(4) by auto
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   492
    have f': "\<forall>n. f (r1 n) k \<in> (\<lambda>x. blinfun_apply x k) ` range f"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   493
      by simp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   494
    have "bounded (range (\<lambda>i. f (r1 i) k))"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   495
      by (metis (lifting) bounded_subset f' image_subsetI s')
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   496
    then obtain l2 r2
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   497
      where r2: "subseq r2"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   498
      and lr2: "((\<lambda>i. f (r1 (r2 i)) k) ---> l2) sequentially"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   499
      using bounded_imp_convergent_subsequence[of "\<lambda>i. f (r1 i) k"]
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   500
      by (auto simp: o_def)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   501
    def r \<equiv> "r1 \<circ> r2"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   502
    have r:"subseq r"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   503
      using r1 and r2 unfolding r_def o_def subseq_def by auto
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   504
    moreover
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   505
    def l \<equiv> "blinfun_of_matrix (\<lambda>i j. if j = k then l2 \<bullet> i else l1 j \<bullet> i)::'a \<Rightarrow>\<^sub>L 'b"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   506
    {
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   507
      fix e::real
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   508
      assume "e > 0"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   509
      from lr1 `e > 0` have N1: "eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n)  i) (l1  i) < e) sequentially"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   510
        by blast
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   511
      from lr2 `e > 0` have N2:"eventually (\<lambda>n. dist (f (r1 (r2 n))  k) l2 < e) sequentially"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   512
        by (rule tendstoD)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   513
      from r2 N1 have N1': "eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 (r2 n))  i) (l1  i) < e) sequentially"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   514
        by (rule eventually_subseq)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   515
      have l2: "l k = l2"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   516
        using insert.prems
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   517
        by (auto simp: blinfun_of_matrix.rep_eq inner_Basis l_def mult_if_delta
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   518
          scaleR_setsum_left[symmetric] setsum.delta' intro!: euclidean_eqI[where 'a='b])
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   519
      {
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   520
        fix i assume "i \<in> d"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   521
        with insert have "i \<in> Basis" "i \<noteq> k" by auto
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   522
        hence l1: "l i = (l1 i)"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   523
          by (auto simp: blinfun_of_matrix.rep_eq inner_Basis l_def mult_if_delta
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   524
            scaleR_setsum_left[symmetric] setsum.delta' intro!: euclidean_eqI[where 'a='b])
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   525
      } note l1 = this
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   526
      have "eventually (\<lambda>n. \<forall>i\<in>(insert k d). 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
   527
        using N1' N2
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   528
        by eventually_elim (insert insert.prems, auto simp: r_def o_def l1 l2)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   529
    }
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   530
    ultimately show ?case by auto
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   531
  qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   532
qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   533
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   534
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
   535
  apply (auto intro!: blinfun_eqI)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   536
  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
   537
  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
   538
  apply (simp add: blinfun.bilinear_simps)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   539
  done
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   540
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   541
text {* TODO: generalize (via @{thm compact_cball})? *}
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   542
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
   543
proof
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   544
  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
   545
  assume f: "bounded (range f)"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   546
  then obtain l::"'a \<Rightarrow>\<^sub>L 'b" and r where r: "subseq r"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   547
    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
   548
    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
   549
  {
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   550
    fix e::real
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   551
    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
   552
    assume "e > 0"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   553
    hence "e / ?d > 0" by (simp add: DIM_positive)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   554
    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
   555
      by simp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   556
    moreover
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   557
    {
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   558
      fix n
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   559
      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
   560
      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
   561
        unfolding blinfun_of_matrix_works ..
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   562
      also note norm_blinfun_of_matrix
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   563
      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
   564
        (\<Sum>i\<in>(Basis::'b set). e / real_of_nat DIM('b))"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   565
      proof (rule setsum_strict_mono)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   566
        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
   567
        have "(\<Sum>j::'a\<in>Basis. \<bar>(f (r n) - l) j \<bullet> i\<bar>) < (\<Sum>j::'a\<in>Basis. e / ?d)"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   568
        proof (rule setsum_strict_mono)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   569
          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
   570
          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
   571
            by (simp add: Basis_le_norm i)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   572
          also have "\<dots> < e / ?d"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   573
            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
   574
          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
   575
        qed simp_all
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   576
        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
   577
          by simp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   578
        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
   579
          by simp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   580
      qed simp_all
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   581
      also have "\<dots> \<le> e" by simp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   582
      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
   583
        by (auto simp: dist_norm)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   584
    }
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   585
    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
   586
      using eventually_elim2 by force
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   587
  }
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   588
  then have *: "((f \<circ> r) ---> l) sequentially"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   589
    unfolding o_def tendsto_iff by simp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   590
  with r show "\<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   591
    by auto
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   592
qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   593
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   594
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   595
subsection \<open>concrete bounded linear functions\<close>
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   596
61916
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   597
lemma transfer_bounded_bilinear_bounded_linearI:
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   598
  assumes "g = (\<lambda>i x. (blinfun_apply (f i) x))"
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   599
  shows "bounded_bilinear g = bounded_linear f"
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   600
proof
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   601
  assume "bounded_bilinear g"
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   602
  then interpret bounded_bilinear f by (simp add: assms)
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   603
  show "bounded_linear f"
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   604
  proof (unfold_locales, safe intro!: blinfun_eqI)
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   605
    fix i
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   606
    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
   607
      by (auto intro!: blinfun_eqI simp: blinfun.bilinear_simps)
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   608
    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
   609
      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
   610
  qed
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   611
qed (auto simp: assms intro!: blinfun.comp)
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   612
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   613
lemma transfer_bounded_bilinear_bounded_linear[transfer_rule]:
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   614
  "(rel_fun (rel_fun op = (pcr_blinfun op = op =)) op =) bounded_bilinear bounded_linear"
7950ae6d3266 transfer rule for bounded_linear of blinfun
immler
parents: 61915
diff changeset
   615
  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
   616
    intro!: transfer_bounded_bilinear_bounded_linearI)
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   617
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   618
context bounded_bilinear
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   619
begin
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   620
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   621
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
   622
  by (rule bounded_linear_left)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   623
declare prod_left.rep_eq[simp]
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   624
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   625
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
   626
  by transfer (rule flip)
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   627
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   628
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
   629
  by (rule bounded_linear_right)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   630
declare prod_right.rep_eq[simp]
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   631
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   632
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
   633
  by transfer (rule bounded_bilinear_axioms)
61915
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
end
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   636
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   637
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
   638
  by (rule bounded_linear_ident)
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
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
   641
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   642
lemma norm_blinfun_id[simp]:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   643
  "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
   644
  by transfer (auto simp: onorm_id)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   645
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   646
lemma norm_blinfun_id_le:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   647
  "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
   648
  by transfer (auto simp: onorm_id_le)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   649
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   650
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   651
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
   652
  by (rule bounded_linear_fst)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   653
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   654
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
   655
  by transfer (rule refl)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   656
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   657
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   658
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
   659
  by (rule bounded_linear_snd)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   660
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   661
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
   662
  by transfer (rule refl)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   663
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   664
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   665
lift_definition blinfun_compose::
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   666
  "'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
   667
    'c::real_normed_vector \<Rightarrow>\<^sub>L 'a \<Rightarrow>
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   668
    'c \<Rightarrow>\<^sub>L 'b" (infixl "o\<^sub>L" 55) is "op o"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   669
  parametric comp_transfer
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   670
  unfolding o_def
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   671
  by (rule bounded_linear_compose)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   672
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   673
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
   674
  by (simp add: blinfun_compose.rep_eq)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   675
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   676
lemma norm_blinfun_compose:
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   677
  "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
   678
  by transfer (rule onorm_compose)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   679
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   680
lemma bounded_bilinear_blinfun_compose[bounded_bilinear]: "bounded_bilinear op o\<^sub>L"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   681
  by unfold_locales
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   682
    (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
   683
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   684
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   685
lift_definition blinfun_inner_right::"'a::real_inner \<Rightarrow> 'a \<Rightarrow>\<^sub>L real" is "op \<bullet>"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   686
  by (rule bounded_linear_inner_right)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   687
declare blinfun_inner_right.rep_eq[simp]
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   688
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   689
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
   690
  by transfer (rule bounded_bilinear_inner)
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   691
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   692
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   693
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
   694
  by (rule bounded_linear_inner_left)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   695
declare blinfun_inner_left.rep_eq[simp]
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   696
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   697
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
   698
  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
   699
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   700
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   701
lift_definition blinfun_scaleR_right::"real \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'a::real_normed_vector" is "op *\<^sub>R"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   702
  by (rule bounded_linear_scaleR_right)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   703
declare blinfun_scaleR_right.rep_eq[simp]
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   704
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   705
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
   706
  by transfer (rule bounded_bilinear_scaleR)
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   707
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   708
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   709
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
   710
  by (rule bounded_linear_scaleR_left)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   711
lemmas [simp] = blinfun_scaleR_left.rep_eq
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   712
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   713
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
   714
  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
   715
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
lift_definition blinfun_mult_right::"'a \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'a::real_normed_algebra" is "op *"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   718
  by (rule bounded_linear_mult_right)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   719
declare blinfun_mult_right.rep_eq[simp]
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   720
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   721
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
   722
  by transfer (rule bounded_bilinear_mult)
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   723
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   724
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   725
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
   726
  by (rule bounded_linear_mult_left)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   727
lemmas [simp] = blinfun_mult_left.rep_eq
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   728
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   729
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
   730
  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
   731
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
diff changeset
   732
end