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