src/HOL/Multivariate_Analysis/Bounded_Continuous_Function.thy
author wenzelm
Sat, 23 May 2015 17:19:37 +0200
changeset 60299 5ae2a2e74c93
parent 60017 b785d6d06430
child 60420 884f54e01427
permissions -rw-r--r--
clarified NEWS: document_files are officially required since Isabelle2014, but the absence was tolerated as legacy feature;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
     1
section {* Bounded Continuous Functions *}
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
     2
theory Bounded_Continuous_Function
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
     3
imports Integration
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
     4
begin
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
     5
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
     6
subsection{* Definition *}
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
     7
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
     8
definition "bcontfun = {f :: 'a::topological_space \<Rightarrow> 'b::metric_space. continuous_on UNIV f \<and> bounded (range f)}"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
     9
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    10
typedef ('a, 'b) bcontfun =
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    11
    "bcontfun :: ('a::topological_space \<Rightarrow> 'b::metric_space) set"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    12
  by (auto simp: bcontfun_def intro: continuous_intros simp: bounded_def)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    13
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    14
lemma bcontfunE:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    15
  assumes "f \<in> bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    16
  obtains y where "continuous_on UNIV f" "\<And>x. dist (f x) u \<le> y"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    17
  using assms unfolding bcontfun_def
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    18
  by (metis (lifting) bounded_any_center dist_commute mem_Collect_eq rangeI)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    19
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    20
lemma bcontfunE':
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    21
  assumes "f \<in> bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    22
  obtains y where "continuous_on UNIV f" "\<And>x. dist (f x) undefined \<le> y"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    23
  using assms bcontfunE
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    24
  by metis
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    25
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    26
lemma bcontfunI:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    27
  "continuous_on UNIV f \<Longrightarrow> (\<And>x. dist (f x) u \<le> b) \<Longrightarrow> f \<in> bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    28
  unfolding bcontfun_def
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    29
  by (metis (lifting, no_types) bounded_def dist_commute mem_Collect_eq rangeE)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    30
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    31
lemma bcontfunI':
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    32
  "continuous_on UNIV f \<Longrightarrow> (\<And>x. dist (f x) undefined \<le> b) \<Longrightarrow> f \<in> bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    33
  using bcontfunI by metis
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    34
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    35
lemma continuous_on_Rep_bcontfun[intro, simp]: "continuous_on T (Rep_bcontfun x)"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    36
  using Rep_bcontfun[of x]
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    37
  by (auto simp: bcontfun_def intro: continuous_on_subset)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    38
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    39
instantiation bcontfun :: (topological_space, metric_space) metric_space
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    40
begin
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    41
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    42
definition dist_bcontfun::"('a, 'b) bcontfun \<Rightarrow> ('a, 'b) bcontfun \<Rightarrow> real" where
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    43
  "dist_bcontfun f g = (SUP x. dist (Rep_bcontfun f x) (Rep_bcontfun g x))"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    44
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    45
definition
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    46
  open_bcontfun::"('a, 'b) bcontfun set \<Rightarrow> bool" where
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    47
  "open_bcontfun S = (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    48
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    49
lemma dist_bounded:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    50
  fixes f ::"('a, 'b) bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    51
  shows "dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le> dist f g"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    52
proof -
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    53
  have "Rep_bcontfun f \<in> bcontfun" using Rep_bcontfun .
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    54
  from bcontfunE'[OF this] obtain y where y:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    55
    "continuous_on UNIV (Rep_bcontfun f)"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    56
    "\<And>x. dist (Rep_bcontfun f x) undefined \<le> y"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    57
    by auto
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    58
  have "Rep_bcontfun g \<in> bcontfun" using Rep_bcontfun .
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    59
  from bcontfunE'[OF this] obtain z where z:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    60
    "continuous_on UNIV (Rep_bcontfun g)"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    61
    "\<And>x. dist (Rep_bcontfun g x) undefined \<le> z"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    62
    by auto
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    63
  show ?thesis unfolding dist_bcontfun_def
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    64
  proof (intro cSUP_upper bdd_aboveI2)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    65
    fix x
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    66
    have "dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le> dist (Rep_bcontfun f x) undefined + dist (Rep_bcontfun g x) undefined"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    67
      by (rule dist_triangle2)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    68
    also have "\<dots> \<le> y + z" using y(2)[of x] z(2)[of x] by (rule add_mono)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    69
    finally show "dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le> y + z" .
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    70
  qed simp
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    71
qed
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    72
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    73
lemma dist_bound:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    74
  fixes f ::"('a, 'b) bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    75
  assumes "\<And>x. dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le> b"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    76
  shows "dist f g \<le> b"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    77
  using assms by (auto simp: dist_bcontfun_def intro: cSUP_least)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    78
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    79
lemma dist_bounded_Abs:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    80
  fixes f g ::"'a \<Rightarrow> 'b"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    81
  assumes "f \<in> bcontfun" "g \<in> bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    82
  shows "dist (f x) (g x) \<le> dist (Abs_bcontfun f) (Abs_bcontfun g)"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    83
  by (metis Abs_bcontfun_inverse assms dist_bounded)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    84
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    85
lemma const_bcontfun: "(\<lambda>x::'a. b::'b) \<in> bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    86
  by (auto intro: bcontfunI continuous_on_const)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    87
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    88
lemma dist_fun_lt_imp_dist_val_lt:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    89
  assumes "dist f g < e"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    90
  shows "dist (Rep_bcontfun f x) (Rep_bcontfun g x) < e"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    91
  using dist_bounded assms by (rule le_less_trans)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    92
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    93
lemma dist_val_lt_imp_dist_fun_le:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    94
  assumes "\<forall>x. dist (Rep_bcontfun f x) (Rep_bcontfun g x) < e"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    95
  shows "dist f g \<le> e"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    96
unfolding dist_bcontfun_def
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    97
proof (intro cSUP_least)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    98
  fix x
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    99
  show "dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le> e"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   100
    using assms[THEN spec[where x=x]] by (simp add: dist_norm)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   101
qed (simp)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   102
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   103
instance
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   104
proof
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   105
  fix f g::"('a, 'b) bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   106
  show "dist f g = 0 \<longleftrightarrow> f = g"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   107
  proof
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   108
    have "\<And>x. dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le> dist f g" by (rule dist_bounded)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   109
    also assume "dist f g = 0"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   110
    finally  show "f = g" by (auto simp: Rep_bcontfun_inject[symmetric] Abs_bcontfun_inverse)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   111
  qed (auto simp: dist_bcontfun_def SUP_def simp del: Sup_image_eq intro!: cSup_eq)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   112
next
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   113
  fix f g h :: "('a, 'b) bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   114
  show "dist f g \<le> dist f h + dist g h"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   115
  proof (subst dist_bcontfun_def, safe intro!: cSUP_least)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   116
    fix x
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   117
    have "dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le>
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   118
      dist (Rep_bcontfun f x) (Rep_bcontfun h x) + dist (Rep_bcontfun g x) (Rep_bcontfun h x)"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   119
      by (rule dist_triangle2)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   120
    also have "dist (Rep_bcontfun f x) (Rep_bcontfun h x) \<le> dist f h" by (rule dist_bounded)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   121
    also have "dist (Rep_bcontfun g x) (Rep_bcontfun h x) \<le> dist g h" by (rule dist_bounded)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   122
    finally show "dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le> dist f h + dist g h" by simp
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   123
  qed
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   124
qed (simp add: open_bcontfun_def)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   125
end
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   126
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   127
lemma closed_Pi_bcontfun:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   128
  fixes I::"'a::metric_space set" and X::"'a \<Rightarrow> 'b::complete_space set"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   129
  assumes "\<And>i. i \<in> I \<Longrightarrow> closed (X i)"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   130
  shows "closed (Abs_bcontfun ` (Pi I X \<inter> bcontfun))"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   131
  unfolding closed_sequential_limits
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   132
proof safe
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   133
  fix f l
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   134
  assume seq: "\<forall>n. f n \<in> Abs_bcontfun ` (Pi I X \<inter> bcontfun)" and lim: "f ----> l"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   135
  have lim_fun: "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x. dist (Rep_bcontfun (f n) x) (Rep_bcontfun l x) < e"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   136
    using LIMSEQ_imp_Cauchy[OF lim, simplified Cauchy_def] metric_LIMSEQ_D[OF lim]
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   137
    by (intro uniformly_cauchy_imp_uniformly_convergent[where P="%_. True", simplified])
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   138
      (metis dist_fun_lt_imp_dist_val_lt)+
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   139
  show "l \<in> Abs_bcontfun ` (Pi I X \<inter> bcontfun)"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   140
  proof (rule, safe)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   141
    fix x assume "x \<in> I"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   142
    hence "closed (X x)" using assms by simp
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   143
    moreover have "eventually (\<lambda>xa. Rep_bcontfun (f xa) x \<in> X x) sequentially"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   144
    proof (rule always_eventually, safe)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   145
      fix i
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   146
      from seq[THEN spec, of i] `x \<in> I`
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   147
      show "Rep_bcontfun (f i) x \<in> X x"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   148
        by (auto simp: Abs_bcontfun_inverse)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   149
    qed
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   150
    moreover note sequentially_bot
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   151
    moreover have "(\<lambda>n. Rep_bcontfun (f n) x) ----> Rep_bcontfun l x"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   152
      using lim_fun by (blast intro!: metric_LIMSEQ_I)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   153
    ultimately show "Rep_bcontfun l x \<in> X x"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   154
      by (rule Lim_in_closed_set)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   155
  qed (auto simp: Rep_bcontfun Rep_bcontfun_inverse)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   156
qed
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   157
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   158
subsection {* Complete Space *}
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   159
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   160
instance bcontfun :: (metric_space, complete_space) complete_space
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   161
proof
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   162
  fix f::"nat \<Rightarrow> ('a,'b) bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   163
  assume "Cauchy f" --{* Cauchy equals uniform convergence *}
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   164
  then obtain g where limit_function:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   165
    "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x. dist (Rep_bcontfun (f n) x) (g x) < e"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   166
    using uniformly_convergent_eq_cauchy[of "\<lambda>_. True"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   167
      "\<lambda>n. Rep_bcontfun (f n)"]
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   168
    unfolding Cauchy_def by (metis dist_fun_lt_imp_dist_val_lt)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   169
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   170
  then obtain N where fg_dist: --{* for an upper bound on g *}
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   171
    "\<forall>n\<ge>N. \<forall>x. dist (g x) ( Rep_bcontfun (f n) x) < 1"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   172
    by (force simp add: dist_commute)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   173
  from bcontfunE'[OF Rep_bcontfun, of "f N"] obtain b where
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   174
    f_bound: "\<forall>x. dist (Rep_bcontfun (f N) x) undefined \<le> b" by force
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   175
  have "g \<in> bcontfun" --{* The limit function is bounded and continuous *}
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   176
  proof (intro bcontfunI)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   177
    show "continuous_on UNIV g"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   178
      using bcontfunE[OF Rep_bcontfun] limit_function
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   179
      by (intro continuous_uniform_limit[where
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   180
        f="%n. Rep_bcontfun (f n)" and F="sequentially"]) (auto
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   181
        simp add: eventually_sequentially trivial_limit_def dist_norm)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   182
  next
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   183
    fix x
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   184
    from fg_dist have "dist (g x) (Rep_bcontfun (f N) x) < 1"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   185
      by (simp add: dist_norm norm_minus_commute)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   186
    with dist_triangle[of "g x" undefined "Rep_bcontfun (f N) x"]
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   187
    show "dist (g x) undefined \<le> 1 + b" using f_bound[THEN spec, of x]
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   188
      by simp
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   189
  qed
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   190
  show "convergent f"
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59865
diff changeset
   191
  proof (rule convergentI, subst lim_sequentially, safe)
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   192
    --{* The limit function converges according to its norm *}
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   193
    fix e::real
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   194
    assume "e > 0" hence "e/2 > 0" by simp
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   195
    with limit_function[THEN spec, of"e/2"]
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   196
    have "\<exists>N. \<forall>n\<ge>N. \<forall>x. dist (Rep_bcontfun (f n) x) (g x) < e/2"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   197
      by simp
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   198
    then obtain N where N: "\<forall>n\<ge>N. \<forall>x. dist (Rep_bcontfun (f n) x) (g x) < e / 2" by auto
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   199
    show "\<exists>N. \<forall>n\<ge>N. dist (f n) (Abs_bcontfun g) < e"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   200
    proof (rule, safe)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   201
      fix n
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   202
      assume "N \<le> n"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   203
      with N show "dist (f n) (Abs_bcontfun g) < e"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   204
        using dist_val_lt_imp_dist_fun_le[of
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   205
          "f n" "Abs_bcontfun g" "e/2"]
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   206
          Abs_bcontfun_inverse[OF `g \<in> bcontfun`] `e > 0` by simp
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   207
    qed
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   208
  qed
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   209
qed
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   210
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   211
subsection{* Supremum norm for a normed vector space *}
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   212
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   213
instantiation bcontfun :: (topological_space, real_normed_vector) real_vector
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   214
begin
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   215
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   216
definition "-f = Abs_bcontfun (\<lambda>x. -(Rep_bcontfun f x))"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   217
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   218
definition "f + g = Abs_bcontfun (\<lambda>x. Rep_bcontfun f x + Rep_bcontfun g x)"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   219
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   220
definition "f - g = Abs_bcontfun (\<lambda>x. Rep_bcontfun f x - Rep_bcontfun g x)"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   221
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   222
definition "0 = Abs_bcontfun (\<lambda>x. 0)"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   223
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   224
definition "scaleR r f = Abs_bcontfun (\<lambda>x. r *\<^sub>R Rep_bcontfun f x)"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   225
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   226
lemma plus_cont:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   227
  fixes f g ::"'a \<Rightarrow> 'b"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   228
  assumes f: "f \<in> bcontfun" and g: "g \<in> bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   229
  shows "(\<lambda>x. f x + g x) \<in> bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   230
proof -
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   231
  from bcontfunE'[OF f] obtain y where "continuous_on UNIV f" "\<And>x. dist (f x) undefined \<le> y"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   232
    by auto
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   233
  moreover
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   234
  from bcontfunE'[OF g] obtain z where "continuous_on UNIV g" "\<And>x. dist (g x) undefined \<le> z"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   235
    by auto
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   236
  ultimately show ?thesis
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   237
  proof (intro bcontfunI)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   238
    fix x
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   239
    have "dist (f x + g x) 0 = dist (f x + g x) (0 + 0)" by simp
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   240
    also have "\<dots> \<le> dist (f x) 0 + dist (g x) 0" by (rule dist_triangle_add)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   241
    also have "\<dots> \<le> dist (Abs_bcontfun f) 0 + dist (Abs_bcontfun g) 0"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   242
      unfolding zero_bcontfun_def using assms
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   243
      by (auto intro!: add_mono dist_bounded_Abs const_bcontfun)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   244
    finally
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   245
    show "dist (f x + g x) 0 <= dist (Abs_bcontfun f) 0 + dist (Abs_bcontfun g) 0" .
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   246
  qed (simp add: continuous_on_add)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   247
qed
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   248
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   249
lemma Rep_bcontfun_plus[simp]: "Rep_bcontfun (f + g) x = Rep_bcontfun f x + Rep_bcontfun g x"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   250
  by (simp add: plus_bcontfun_def Abs_bcontfun_inverse plus_cont Rep_bcontfun)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   251
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   252
lemma uminus_cont:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   253
  fixes f ::"'a \<Rightarrow> 'b"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   254
  assumes "f \<in> bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   255
  shows "(\<lambda>x. - f x) \<in> bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   256
proof -
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   257
  from bcontfunE[OF assms, of 0] obtain y where "continuous_on UNIV f" "\<And>x. dist (f x) 0 \<le> y"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   258
    by auto
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   259
  thus ?thesis
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   260
  proof (intro bcontfunI)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   261
    fix x
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   262
    assume "\<And>x. dist (f x) 0 \<le> y"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   263
    thus "dist (- f x) 0 \<le> y" by (subst dist_minus[symmetric]) simp
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   264
  qed (simp add: continuous_on_minus)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   265
qed
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   266
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   267
lemma Rep_bcontfun_uminus[simp]:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   268
  "Rep_bcontfun (- f) x = - Rep_bcontfun f x"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   269
  by (simp add: uminus_bcontfun_def Abs_bcontfun_inverse uminus_cont Rep_bcontfun)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   270
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   271
lemma minus_cont:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   272
  fixes f g ::"'a \<Rightarrow> 'b"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   273
  assumes f: "f \<in> bcontfun" and g: "g \<in> bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   274
  shows "(\<lambda>x. f x - g x) \<in> bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   275
  using plus_cont [of f "- g"] assms by (simp add: uminus_cont fun_Compl_def)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   276
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   277
lemma Rep_bcontfun_minus[simp]:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   278
  "Rep_bcontfun (f - g) x = Rep_bcontfun f x - Rep_bcontfun g x"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   279
  by (simp add: minus_bcontfun_def Abs_bcontfun_inverse minus_cont Rep_bcontfun)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   280
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   281
lemma scaleR_cont:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   282
  fixes a and f::"'a \<Rightarrow> 'b"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   283
  assumes "f \<in> bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   284
  shows " (\<lambda>x. a *\<^sub>R f x) \<in> bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   285
proof -
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   286
  from bcontfunE[OF assms, of 0] obtain y where "continuous_on UNIV f" "\<And>x. dist (f x) 0 \<le> y"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   287
    by auto
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   288
  thus ?thesis
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   289
  proof (intro bcontfunI)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   290
    fix x assume "\<And>x. dist (f x) 0 \<le> y"
59554
4044f53326c9 inlined rules to free user-space from technical names
haftmann
parents: 59453
diff changeset
   291
    then show "dist (a *\<^sub>R f x) 0 \<le> abs a * y"
4044f53326c9 inlined rules to free user-space from technical names
haftmann
parents: 59453
diff changeset
   292
      by (metis norm_cmul_rule_thm norm_conv_dist)
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   293
  qed (simp add: continuous_intros)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   294
qed
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   295
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   296
lemma Rep_bcontfun_scaleR[simp]:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   297
   "Rep_bcontfun (a *\<^sub>R g) x = a *\<^sub>R Rep_bcontfun g x"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   298
  by (simp add: scaleR_bcontfun_def Abs_bcontfun_inverse scaleR_cont Rep_bcontfun)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   299
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   300
instance
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   301
proof
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   302
qed (simp_all add: plus_bcontfun_def zero_bcontfun_def minus_bcontfun_def scaleR_bcontfun_def
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   303
    Abs_bcontfun_inverse Rep_bcontfun_inverse Rep_bcontfun algebra_simps
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   304
    plus_cont const_bcontfun minus_cont scaleR_cont)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   305
end
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   306
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   307
instantiation bcontfun :: (topological_space, real_normed_vector) real_normed_vector
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   308
begin
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   309
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   310
definition norm_bcontfun::"('a, 'b) bcontfun \<Rightarrow> real" where
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   311
  "norm_bcontfun f = dist f 0"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   312
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   313
definition "sgn (f::('a,'b) bcontfun) = f /\<^sub>R norm f"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   314
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   315
instance
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   316
proof
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   317
  fix f g::"('a, 'b) bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   318
  show "dist f g = norm (f - g)"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   319
    by (simp add: norm_bcontfun_def dist_bcontfun_def zero_bcontfun_def
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   320
    Abs_bcontfun_inverse const_bcontfun norm_conv_dist[symmetric] dist_norm)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   321
  show "norm (f + g) \<le> norm f + norm g"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   322
    unfolding norm_bcontfun_def
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   323
  proof (subst dist_bcontfun_def, safe intro!: cSUP_least)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   324
    fix x
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   325
    have "dist (Rep_bcontfun (f + g) x) (Rep_bcontfun 0 x) \<le>
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   326
      dist (Rep_bcontfun f x) 0 + dist (Rep_bcontfun g x) 0"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   327
      by (metis (hide_lams, no_types) Rep_bcontfun_minus Rep_bcontfun_plus diff_0_right dist_norm
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   328
        le_less_linear less_irrefl norm_triangle_lt)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   329
    also have "dist (Rep_bcontfun f x) 0 \<le> dist f 0"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   330
      using dist_bounded[of f x 0]
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   331
      by (simp add: Abs_bcontfun_inverse const_bcontfun zero_bcontfun_def)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   332
    also have "dist (Rep_bcontfun g x) 0 \<le> dist g 0" using dist_bounded[of g x 0]
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   333
      by (simp add: Abs_bcontfun_inverse const_bcontfun zero_bcontfun_def)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   334
    finally show "dist (Rep_bcontfun (f + g) x) (Rep_bcontfun 0 x) \<le> dist f 0 + dist g 0" by simp
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   335
  qed
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   336
next
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   337
  fix a and f g:: "('a, 'b) bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   338
  show "norm (a *\<^sub>R f) = \<bar>a\<bar> * norm f"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   339
  proof -
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   340
    have "\<bar>a\<bar> * Sup (range (\<lambda>x. dist (Rep_bcontfun f x) 0)) =
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   341
      (SUP i:range (\<lambda>x. dist (Rep_bcontfun f x) 0). \<bar>a\<bar> * i)"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   342
    proof (intro continuous_at_Sup_mono bdd_aboveI2)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   343
      fix x
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   344
      show "dist (Rep_bcontfun f x) 0 \<le> norm f" using dist_bounded[of f x 0]
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   345
        by (simp add: norm_bcontfun_def norm_conv_dist Abs_bcontfun_inverse zero_bcontfun_def
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   346
          const_bcontfun)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   347
    qed (auto intro!: monoI mult_left_mono continuous_intros)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   348
    moreover
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   349
    have "range (\<lambda>x. dist (Rep_bcontfun (a *\<^sub>R f) x) 0) = 
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   350
      (\<lambda>x. \<bar>a\<bar> * x) ` (range (\<lambda>x. dist (Rep_bcontfun f x) 0))"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   351
      by (auto simp: norm_conv_dist[symmetric])
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   352
    ultimately
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   353
    show "norm (a *\<^sub>R f) = \<bar>a\<bar> * norm f"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   354
      by (simp add: norm_bcontfun_def dist_bcontfun_def norm_conv_dist Abs_bcontfun_inverse
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   355
                    zero_bcontfun_def const_bcontfun SUP_def del: Sup_image_eq)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   356
  qed
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   357
qed (auto simp: norm_bcontfun_def sgn_bcontfun_def)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   358
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   359
end
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   360
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   361
lemma bcontfun_normI:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   362
  "continuous_on UNIV f \<Longrightarrow> (\<And>x. norm (f x) \<le> b) \<Longrightarrow> f \<in> bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   363
  unfolding norm_conv_dist
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   364
  by (auto intro: bcontfunI)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   365
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   366
lemma norm_bounded:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   367
  fixes f ::"('a::topological_space, 'b::real_normed_vector) bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   368
  shows "norm (Rep_bcontfun f x) \<le> norm f"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   369
  using dist_bounded[of f x 0]
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   370
  by (simp add: norm_bcontfun_def norm_conv_dist Abs_bcontfun_inverse zero_bcontfun_def
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   371
    const_bcontfun)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   372
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   373
lemma norm_bound:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   374
  fixes f ::"('a::topological_space, 'b::real_normed_vector) bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   375
  assumes "\<And>x. norm (Rep_bcontfun f x) \<le> b"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   376
  shows "norm f \<le> b"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   377
  using dist_bound[of f 0 b] assms
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   378
  by (simp add: norm_bcontfun_def norm_conv_dist Abs_bcontfun_inverse zero_bcontfun_def
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   379
    const_bcontfun)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   380
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   381
subsection{* Continuously Extended Functions *}
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   382
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   383
definition clamp::"'a::euclidean_space \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" where
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   384
  "clamp a b x = (\<Sum>i\<in>Basis. (if x\<bullet>i < a\<bullet>i then a\<bullet>i else if x\<bullet>i \<le> b\<bullet>i then x\<bullet>i else b\<bullet>i) *\<^sub>R i)"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   385
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   386
definition ext_cont::"('a::euclidean_space \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> ('a, 'b) bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   387
  where "ext_cont f a b = Abs_bcontfun ((\<lambda>x. f (clamp a b x)))"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   388
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   389
lemma ext_cont_def':
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   390
  "ext_cont f a b = Abs_bcontfun (\<lambda>x.
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   391
    f (\<Sum>i\<in>Basis. (if x\<bullet>i < a\<bullet>i then a\<bullet>i else if x\<bullet>i \<le> b\<bullet>i then x\<bullet>i else b\<bullet>i) *\<^sub>R i))"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   392
unfolding ext_cont_def clamp_def ..
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   393
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   394
lemma clamp_in_interval:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   395
  assumes "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   396
  shows "clamp a b x \<in> cbox a b"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   397
  unfolding clamp_def
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   398
  using box_ne_empty(1)[of a b] assms by (auto simp: cbox_def)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   399
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   400
lemma dist_clamps_le_dist_args:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   401
  fixes x::"'a::euclidean_space"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   402
  assumes "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   403
  shows "dist (clamp a b y) (clamp a b x) \<le> dist y x"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   404
proof -
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   405
    from box_ne_empty(1)[of a b] assms have "(\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i)"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   406
      by (simp add: cbox_def)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   407
    hence "(\<Sum>i\<in>Basis. (dist (clamp a b y \<bullet> i) (clamp a b x \<bullet> i))\<^sup>2)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   408
        \<le> (\<Sum>i\<in>Basis. (dist (y \<bullet> i) (x \<bullet> i))\<^sup>2)"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   409
      by (auto intro!: setsum_mono
59865
8a20dd967385 rationalised and generalised some theorems concerning abs and x^2.
paulson <lp15@cam.ac.uk>
parents: 59554
diff changeset
   410
        simp add: clamp_def dist_real_def abs_le_square_iff[symmetric])
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   411
    thus ?thesis
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   412
      by (auto intro: real_sqrt_le_mono
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   413
        simp add: euclidean_dist_l2[where y=x] euclidean_dist_l2[where y="clamp a b x"] setL2_def)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   414
qed
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   415
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   416
lemma clamp_continuous_at:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   417
  fixes f::"'a::euclidean_space \<Rightarrow> 'b::metric_space"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   418
  fixes x
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   419
  assumes "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   420
  assumes f_cont: "continuous_on (cbox a b) f"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   421
  shows "continuous (at x) (\<lambda>x. f (clamp a b x))"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   422
unfolding continuous_at_eps_delta
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   423
proof (safe)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   424
  fix x::'a and e::real
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   425
  assume "0 < e"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   426
  moreover
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   427
  have "clamp a b x \<in> cbox a b" by (simp add: clamp_in_interval assms)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   428
  moreover
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   429
  note f_cont[simplified continuous_on_iff]
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   430
  ultimately
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   431
  obtain d where d: "0 < d"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   432
    "\<And>x'. x' \<in> cbox a b \<Longrightarrow> dist x' (clamp a b x) < d \<Longrightarrow> dist (f x') (f (clamp a b x)) < e"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   433
    by force
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   434
  show "\<exists>d>0. \<forall>x'. dist x' x < d \<longrightarrow>
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   435
    dist (f (clamp a b x')) (f (clamp a b x)) < e"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   436
    by (auto intro!: d clamp_in_interval assms dist_clamps_le_dist_args[THEN le_less_trans])
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   437
qed
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   438
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   439
lemma clamp_continuous_on:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   440
  fixes f::"'a::euclidean_space \<Rightarrow> 'b::metric_space"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   441
  assumes "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   442
  assumes f_cont: "continuous_on (cbox a b) f"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   443
  shows "continuous_on UNIV (\<lambda>x. f (clamp a b x))"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   444
  using assms
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   445
  by (auto intro: continuous_at_imp_continuous_on clamp_continuous_at)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   446
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   447
lemma clamp_bcontfun:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   448
  fixes f::"'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   449
  assumes "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   450
  assumes continuous: "continuous_on (cbox a b) f"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   451
  shows "(\<lambda>x. f (clamp a b x)) \<in> bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   452
proof -
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   453
  from compact_continuous_image[OF continuous compact_cbox[of a b], THEN compact_imp_bounded]
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   454
  have "bounded (f ` (cbox a b))" .
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   455
  then obtain c where f_bound: "\<forall>x\<in>f ` cbox a b. norm x \<le> c" by (auto simp add: bounded_pos)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   456
  show "(\<lambda>x. f (clamp a b x)) \<in> bcontfun"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   457
  proof (intro bcontfun_normI)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   458
    fix x
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   459
    from clamp_in_interval[OF assms(1), of x] f_bound
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   460
    show "norm (f (clamp a b x)) \<le> c" by (simp add: ext_cont_def)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   461
  qed (simp add: clamp_continuous_on assms)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   462
qed
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   463
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   464
lemma integral_clamp:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   465
  "integral {t0::real..clamp t0 t1 x} f =
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   466
    (if x < t0 then 0 else if x \<le> t1 then integral {t0..x} f else integral {t0..t1} f)"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   467
  by (auto simp: clamp_def)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   468
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   469
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   470
declare [[coercion Rep_bcontfun]]
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   471
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   472
lemma ext_cont_cancel[simp]:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   473
  fixes x a b::"'a::euclidean_space"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   474
  assumes x: "x \<in> cbox a b"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   475
  assumes "continuous_on (cbox a b) f"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   476
  shows "ext_cont f a b x = f x"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   477
  using assms
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   478
  unfolding ext_cont_def
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   479
proof (subst Abs_bcontfun_inverse[OF clamp_bcontfun])
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   480
  show "f (clamp a b x) = f x"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   481
    using x unfolding clamp_def mem_box
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   482
    by (intro arg_cong[where f=f] euclidean_eqI[where 'a='a]) (simp add: not_less)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   483
qed (auto simp: cbox_def)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   484
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   485
lemma ext_cont_cong:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   486
  assumes "t0 = s0"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   487
  assumes "t1 = s1"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   488
  assumes "\<And>t. t \<in> (cbox t0 t1) \<Longrightarrow> f t = g t"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   489
  assumes "continuous_on (cbox t0 t1) f"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   490
  assumes "continuous_on (cbox s0 s1) g"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   491
  assumes ord: "\<And>i. i \<in> Basis \<Longrightarrow> t0 \<bullet> i \<le> t1 \<bullet> i"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   492
  shows "ext_cont f t0 t1 = ext_cont g s0 s1"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   493
  unfolding assms ext_cont_def
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   494
  using assms clamp_in_interval[OF ord]
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   495
  by (subst Rep_bcontfun_inject[symmetric]) simp
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   496
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   497
end