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