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