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