src/HOL/Analysis/Bounded_Continuous_Function.thy
author immler
Fri, 10 Mar 2017 23:16:40 +0100
changeset 65204 d23eded35a33
parent 65036 ab7e11730ad8
child 65205 f435640193b6
permissions -rw-r--r--
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
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
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
     4
  imports
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
     5
    Henstock_Kurzweil_Integration
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
     6
begin
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
     7
60421
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
     8
subsection \<open>Definition\<close>
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
     9
60421
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
    10
definition bcontfun :: "('a::topological_space \<Rightarrow> 'b::metric_space) set"
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
    11
  where "bcontfun = {f. continuous_on UNIV f \<and> bounded (range f)}"
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    12
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    13
typedef (overloaded) ('a, 'b) bcontfun ("(_ \<Rightarrow>\<^sub>C /_)" [22, 21] 21) =
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    14
    "{f::'a::topological_space \<Rightarrow> 'b::metric_space. continuous_on UNIV f \<and> bounded (range f)}"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    15
  morphisms apply_bcontfun Bcontfun
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    16
  by (auto intro: continuous_intros simp: bounded_def)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    17
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    18
declare [[coercion "apply_bcontfun :: ('a::topological_space \<Rightarrow>\<^sub>C'b::metric_space) \<Rightarrow> 'a \<Rightarrow> 'b"]]
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    19
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    20
setup_lifting type_definition_bcontfun
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    21
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    22
lemma continuous_on_apply_bcontfun[intro, simp]: "continuous_on T (apply_bcontfun x)"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    23
  and bounded_apply_bcontfun[intro, simp]: "bounded (range (apply_bcontfun x))"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    24
  using apply_bcontfun[of x]
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    25
  by (auto simp: intro: continuous_on_subset)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    26
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    27
lemma bcontfun_eqI: "(\<And>x. apply_bcontfun f x = apply_bcontfun g x) \<Longrightarrow> f = g"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    28
  by transfer auto
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    29
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    30
lemma bcontfunE:
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    31
  assumes "continuous_on UNIV f" "bounded (range f)"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    32
  obtains g where "f = apply_bcontfun g"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    33
  by (blast intro: apply_bcontfun_cases assms)
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    34
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    35
lift_definition const_bcontfun::"'b::metric_space \<Rightarrow> ('a::topological_space \<Rightarrow>\<^sub>C 'b)" is "\<lambda>c _. c"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    36
  by (auto intro!: continuous_intros simp: image_def)
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    37
62101
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61969
diff changeset
    38
(* TODO: Generalize to uniform spaces? *)
59453
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
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    42
lift_definition dist_bcontfun :: "'a \<Rightarrow>\<^sub>C 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>C 'b \<Rightarrow> real"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    43
  is "\<lambda>f g. (SUP x. dist (f x) (g x))" .
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    44
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    45
definition uniformity_bcontfun :: "('a \<Rightarrow>\<^sub>C 'b \<times> 'a \<Rightarrow>\<^sub>C 'b) filter"
62101
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61969
diff changeset
    46
  where "uniformity_bcontfun = (INF e:{0 <..}. principal {(x, y). dist x y < e})"
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61969
diff changeset
    47
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    48
definition open_bcontfun :: "('a \<Rightarrow>\<^sub>C 'b) set \<Rightarrow> bool"
62101
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61969
diff changeset
    49
  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
    50
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    51
lemma bounded_dist_le_SUP_dist:
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    52
  "bounded (range f) \<Longrightarrow> bounded (range g) \<Longrightarrow> dist (f x) (g x) \<le> (SUP x. dist (f x) (g x))"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    53
  by (auto intro!: cSUP_upper bounded_imp_bdd_above bounded_dist_comp)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    54
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    55
lemma dist_bounded:
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    56
  fixes f g :: "'a \<Rightarrow>\<^sub>C 'b"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    57
  shows "dist (f x) (g x) \<le> dist f g"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    58
  by transfer (auto intro!: bounded_dist_le_SUP_dist)
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    59
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    60
lemma dist_bound:
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    61
  fixes f g :: "'a \<Rightarrow>\<^sub>C 'b"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    62
  assumes "\<And>x. dist (f x) (g x) \<le> b"
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    63
  shows "dist f g \<le> b"
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    64
  using assms
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    65
  by transfer (auto intro!: cSUP_least)
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    66
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    67
lemma dist_fun_lt_imp_dist_val_lt:
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    68
  fixes f g :: "'a \<Rightarrow>\<^sub>C 'b"
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    69
  assumes "dist f g < e"
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    70
  shows "dist (f x) (g x) < e"
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    71
  using dist_bounded assms by (rule le_less_trans)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    72
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    73
instance
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    74
proof
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    75
  fix f g h :: "'a \<Rightarrow>\<^sub>C 'b"
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    76
  show "dist f g = 0 \<longleftrightarrow> f = g"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    77
  proof
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    78
    have "\<And>x. dist (f x) (g x) \<le> dist f g"
60421
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
    79
      by (rule dist_bounded)
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    80
    also assume "dist f g = 0"
60421
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
    81
    finally show "f = g"
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    82
      by (auto simp: apply_bcontfun_inject[symmetric])
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 62101
diff changeset
    83
  qed (auto simp: dist_bcontfun_def intro!: cSup_eq)
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    84
  show "dist f g \<le> dist f h + dist g h"
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    85
  proof (rule dist_bound)
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    86
    fix x
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    87
    have "dist (f x) (g x) \<le> dist (f x) (h x) + dist (g x) (h x)"
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    88
      by (rule dist_triangle2)
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    89
    also have "dist (f x) (h x) \<le> dist f h"
60421
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
    90
      by (rule dist_bounded)
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    91
    also have "dist (g x) (h x) \<le> dist g h"
60421
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
    92
      by (rule dist_bounded)
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
    93
    finally show "dist (f x) (g x) \<le> dist f h + dist g h"
60421
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
    94
      by simp
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    95
  qed
62101
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61969
diff changeset
    96
qed (rule open_bcontfun_def uniformity_bcontfun_def)+
60421
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
    97
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    98
end
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
    99
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   100
lift_definition PiC::"'a::topological_space set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> ('a \<Rightarrow>\<^sub>C 'b::metric_space) set"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   101
  is "\<lambda>I X. Pi I X \<inter> {f. continuous_on UNIV f \<and> bounded (range f)}"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   102
  by auto
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   103
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   104
lemma mem_PiC_iff: "x \<in> PiC I X \<longleftrightarrow> apply_bcontfun x \<in> Pi I X"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   105
  by transfer simp
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   106
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   107
lemmas mem_PiCD = mem_PiC_iff[THEN iffD1]
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   108
  and mem_PiCI = mem_PiC_iff[THEN iffD2]
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   109
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   110
lemma tendsto_bcontfun_uniform_limit:
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   111
  fixes f::"'i \<Rightarrow> 'a::topological_space \<Rightarrow>\<^sub>C 'b::metric_space"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   112
  assumes "(f \<longlongrightarrow> l) F"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   113
  shows "uniform_limit UNIV f l F"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   114
proof (rule uniform_limitI)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   115
  fix e::real assume "e > 0"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   116
  from tendstoD[OF assms this] have "\<forall>\<^sub>F x in F. dist (f x) l < e" .
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   117
  then show "\<forall>\<^sub>F n in F. \<forall>x\<in>UNIV. dist ((f n) x) (l x) < e"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   118
    by eventually_elim (auto simp: dist_fun_lt_imp_dist_val_lt)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   119
qed
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   120
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   121
lemma uniform_limit_tendsto_bcontfun:
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   122
  fixes f::"'i \<Rightarrow> 'a::topological_space \<Rightarrow>\<^sub>C 'b::metric_space"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   123
    and l::"'a::topological_space \<Rightarrow>\<^sub>C 'b::metric_space"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   124
  assumes "uniform_limit UNIV f l F"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   125
  shows "(f \<longlongrightarrow> l) F"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   126
proof (rule tendstoI)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   127
  fix e::real assume "e > 0"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   128
  then have "e / 2 > 0" by simp
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   129
  from uniform_limitD[OF assms this]
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   130
  have "\<forall>\<^sub>F i in F. \<forall>x. dist (f i x) (l x) < e / 2" by simp
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   131
  then have "\<forall>\<^sub>F x in F. dist (f x) l \<le> e / 2"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   132
    by eventually_elim (blast intro: dist_bound less_imp_le)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   133
  then show "\<forall>\<^sub>F x in F. dist (f x) l < e"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   134
    by eventually_elim (use \<open>0 < e\<close> in auto)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   135
qed
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   136
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   137
lemma uniform_limit_bcontfunE:
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   138
  fixes f::"'i \<Rightarrow> 'a::topological_space \<Rightarrow>\<^sub>C 'b::metric_space"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   139
    and l::"'a::topological_space \<Rightarrow> 'b::metric_space"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   140
  assumes "uniform_limit UNIV f l F" "F \<noteq> bot"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   141
  obtains l'::"'a::topological_space \<Rightarrow>\<^sub>C 'b::metric_space"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   142
  where "l = l'" "(f \<longlongrightarrow> l') F"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   143
  by (metis (mono_tags, lifting) always_eventually apply_bcontfun apply_bcontfun_cases assms
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   144
      mem_Collect_eq uniform_limit_bounded uniform_limit_tendsto_bcontfun uniform_limit_theorem)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   145
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   146
lemma closed_PiC:
60421
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
   147
  fixes I :: "'a::metric_space set"
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
   148
    and X :: "'a \<Rightarrow> 'b::complete_space set"
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   149
  assumes "\<And>i. i \<in> I \<Longrightarrow> closed (X i)"
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   150
  shows "closed (PiC I X)"
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   151
  unfolding closed_sequential_limits
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   152
proof safe
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   153
  fix f l
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   154
  assume seq: "\<forall>n. f n \<in> PiC I X" and lim: "f \<longlonglongrightarrow> l"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   155
  show "l \<in> PiC I X"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   156
  proof (safe intro!: mem_PiCI)
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   157
    fix x assume "x \<in> I"
60421
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
   158
    then have "closed (X x)"
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
   159
      using assms by simp
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   160
    moreover have "eventually (\<lambda>i. f i x \<in> X x) sequentially"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   161
      using seq \<open>x \<in> I\<close>
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   162
      by (auto intro!: eventuallyI dest!: mem_PiCD simp: Pi_iff)
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   163
    moreover note sequentially_bot
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   164
    moreover have "(\<lambda>n. (f n) x) \<longlonglongrightarrow> l x"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   165
      using tendsto_bcontfun_uniform_limit[OF lim]
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   166
      by (rule tendsto_uniform_limitI) simp
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   167
    ultimately show "l x \<in> X x"
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   168
      by (rule Lim_in_closed_set)
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   169
  qed
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   170
qed
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   171
60421
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
   172
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60017
diff changeset
   173
subsection \<open>Complete Space\<close>
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   174
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   175
instance bcontfun :: (metric_space, complete_space) complete_space
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   176
proof
60421
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
   177
  fix f :: "nat \<Rightarrow> ('a, 'b) bcontfun"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61260
diff changeset
   178
  assume "Cauchy f"  \<comment> \<open>Cauchy equals uniform convergence\<close>
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   179
  then obtain g where "uniform_limit UNIV f g sequentially"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   180
    using uniformly_convergent_eq_cauchy[of "\<lambda>_. True" f]
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   181
    unfolding Cauchy_def uniform_limit_sequentially_iff
60421
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
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   184
  from uniform_limit_bcontfunE[OF this sequentially_bot]
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   185
  obtain l' where "g = apply_bcontfun l'" "(f \<longlonglongrightarrow> l')" by metis
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   186
  then show "convergent f"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   187
    by (intro convergentI)
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   188
qed
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   189
60421
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
   190
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
   191
subsection \<open>Supremum norm for a normed vector space\<close>
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   192
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   193
instantiation bcontfun :: (topological_space, real_normed_vector) real_vector
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   194
begin
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   195
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   196
lift_definition uminus_bcontfun::"('a \<Rightarrow>\<^sub>C 'b) \<Rightarrow> 'a \<Rightarrow>\<^sub>C 'b" is "\<lambda>f x. - f x"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   197
  by (auto intro!: continuous_intros)
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   198
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   199
lift_definition plus_bcontfun::"('a \<Rightarrow>\<^sub>C 'b) \<Rightarrow> ('a \<Rightarrow>\<^sub>C 'b) \<Rightarrow> 'a \<Rightarrow>\<^sub>C 'b"  is "\<lambda>f g x. f x + g x"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   200
  by (auto simp: intro!: continuous_intros bounded_plus_comp)
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   201
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   202
lift_definition minus_bcontfun::"('a \<Rightarrow>\<^sub>C 'b) \<Rightarrow> ('a \<Rightarrow>\<^sub>C 'b) \<Rightarrow> 'a \<Rightarrow>\<^sub>C 'b"  is "\<lambda>f g x. f x - g x"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   203
  by (auto simp: intro!: continuous_intros bounded_minus_comp)
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   204
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   205
lift_definition zero_bcontfun::"'a \<Rightarrow>\<^sub>C 'b" is "\<lambda>_. 0"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   206
  by (auto intro!: continuous_intros simp: image_def)
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   207
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   208
lemma const_bcontfun_0_eq_0[simp]: "const_bcontfun 0 = 0"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   209
  by transfer simp
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   210
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   211
lift_definition scaleR_bcontfun::"real \<Rightarrow> ('a \<Rightarrow>\<^sub>C 'b) \<Rightarrow> 'a \<Rightarrow>\<^sub>C 'b"  is "\<lambda>r g x. r *\<^sub>R g x"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   212
  by (auto simp: intro!: continuous_intros bounded_scaleR_comp)
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   213
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   214
lemmas [simp] =
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   215
  const_bcontfun.rep_eq
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   216
  uminus_bcontfun.rep_eq
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   217
  plus_bcontfun.rep_eq
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   218
  minus_bcontfun.rep_eq
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   219
  zero_bcontfun.rep_eq
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   220
  scaleR_bcontfun.rep_eq
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   221
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   222
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   223
instance
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   224
  by standard (auto intro!: bcontfun_eqI simp: algebra_simps)
60421
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
   225
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   226
end
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   227
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   228
lemma bounded_norm_le_SUP_norm:
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   229
  "bounded (range f) \<Longrightarrow> norm (f x) \<le> (SUP x. norm (f x))"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   230
  by (auto intro!: cSUP_upper bounded_imp_bdd_above simp: bounded_norm_comp)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   231
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   232
instantiation bcontfun :: (topological_space, real_normed_vector) real_normed_vector
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   233
begin
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   234
60421
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
   235
definition norm_bcontfun :: "('a, 'b) bcontfun \<Rightarrow> real"
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
   236
  where "norm_bcontfun f = dist f 0"
59453
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 "sgn (f::('a,'b) bcontfun) = f /\<^sub>R norm f"
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
instance
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   241
proof
60421
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
   242
  fix a :: real
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
   243
  fix f g :: "('a, 'b) bcontfun"
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   244
  show "dist f g = norm (f - g)"
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   245
    unfolding norm_bcontfun_def
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   246
    by transfer (simp add: dist_norm)
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   247
  show "norm (f + g) \<le> norm f + norm g"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   248
    unfolding norm_bcontfun_def
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   249
    by transfer
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   250
      (auto intro!: cSUP_least norm_triangle_le add_mono bounded_norm_le_SUP_norm simp: dist_norm)
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   251
  show "norm (a *\<^sub>R f) = \<bar>a\<bar> * norm f"
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   252
    unfolding norm_bcontfun_def
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   253
    apply transfer
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   254
    by (rule trans[OF _ continuous_at_Sup_mono[symmetric]])
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   255
      (auto intro!: monoI mult_left_mono continuous_intros bounded_imp_bdd_above
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   256
        simp: bounded_norm_comp)
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   257
qed (auto simp: norm_bcontfun_def sgn_bcontfun_def)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   258
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   259
end
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   260
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   261
lemma norm_bounded:
60421
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
   262
  fixes f :: "('a::topological_space, 'b::real_normed_vector) bcontfun"
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   263
  shows "norm (apply_bcontfun f x) \<le> norm f"
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   264
  using dist_bounded[of f x 0]
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   265
  by (simp add: dist_norm)
59453
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 norm_bound:
60421
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
   268
  fixes f :: "('a::topological_space, 'b::real_normed_vector) bcontfun"
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   269
  assumes "\<And>x. norm (apply_bcontfun f x) \<le> b"
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   270
  shows "norm f \<le> b"
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   271
  using dist_bound[of f 0 b] assms
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   272
  by (simp add: dist_norm)
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   273
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   274
subsection \<open>(bounded) continuous extenstion\<close>
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   275
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   276
lemma integral_clamp:
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   277
  "integral {t0::real..clamp t0 t1 x} f =
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   278
    (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
   279
  by (auto simp: clamp_def)
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   280
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   281
lemma continuous_on_interval_bcontfunE:
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   282
  fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::metric_space"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   283
  assumes "continuous_on {a .. b} f"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   284
  obtains g::"'a \<Rightarrow>\<^sub>C 'b" where
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   285
    "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> g x = f x"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   286
    "\<And>x. g x = f (clamp a b x)"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   287
proof -
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   288
  define g where "g \<equiv> ext_cont f a b"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   289
  have "continuous_on UNIV g"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   290
    using assms
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   291
    by (auto intro!: continuous_on_ext_cont simp: g_def cbox_interval)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   292
  moreover
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   293
  have "bounded (range g)"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   294
    by (auto simp: g_def ext_cont_def cbox_interval
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   295
        intro!: compact_interval clamp_bounded compact_imp_bounded[OF compact_continuous_image] assms)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   296
  ultimately
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   297
  obtain h where h: "g = apply_bcontfun h" by (rule bcontfunE)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   298
  then have "h x = f x" if "a \<le> x" "x \<le> b" for x
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   299
    by (auto simp: h[symmetric] g_def cbox_interval that)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   300
  moreover
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   301
  have "h x = f (clamp a b x)" for x
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   302
    by (auto simp: h[symmetric] g_def ext_cont_def)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   303
  ultimately show ?thesis ..
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   304
qed
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   305
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   306
lifting_update bcontfun.lifting
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 65036
diff changeset
   307
lifting_forget bcontfun.lifting
59453
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   308
4736ff5a41d8 moved bcontfun from AFP/Ordinary_Differential_Equations
hoelzl
parents:
diff changeset
   309
end