src/HOL/Multivariate_Analysis/Weierstrass.thy
changeset 60987 ea00d17eba3b
child 61222 05d28dc76e5c
equal deleted inserted replaced
60986:077f663b6c24 60987:ea00d17eba3b
       
     1 section{*Bernstein-Weierstrass and Stone-Weierstrass Theorems*}
       
     2 
       
     3 theory Weierstrass
       
     4 imports Uniform_Limit Path_Connected
       
     5 
       
     6 begin
       
     7 
       
     8 (*FIXME: simplification changes to be enforced globally*)
       
     9 declare of_nat_Suc [simp del]
       
    10 
       
    11 (*Power.thy:*)
       
    12 declare power_divide [where b = "numeral w" for w, simp del]
       
    13 
       
    14 subsection {*Bernstein polynomials*}
       
    15 
       
    16 definition Bernstein :: "[nat,nat,real] \<Rightarrow> real" where
       
    17   "Bernstein n k x \<equiv> of_nat (n choose k) * x ^ k * (1 - x) ^ (n - k)"
       
    18 
       
    19 lemma Bernstein_nonneg: "\<lbrakk>0 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> 0 \<le> Bernstein n k x"
       
    20   by (simp add: Bernstein_def)
       
    21 
       
    22 lemma Bernstein_pos: "\<lbrakk>0 < x; x < 1; k \<le> n\<rbrakk> \<Longrightarrow> 0 < Bernstein n k x"
       
    23   by (simp add: Bernstein_def)
       
    24 
       
    25 lemma sum_Bernstein [simp]: "(\<Sum> k = 0..n. Bernstein n k x) = 1"
       
    26   using binomial_ring [of x "1-x" n]
       
    27   by (simp add: Bernstein_def)
       
    28 
       
    29 lemma binomial_deriv1:
       
    30     "(\<Sum>k=0..n. (of_nat k * of_nat (n choose k)) * a^(k-1) * b^(n-k)) =
       
    31      of_nat n * (a+b::real) ^ (n-1)"
       
    32   apply (rule DERIV_unique [where f = "\<lambda>a. (a+b)^n" and x=a])
       
    33   apply (subst binomial_ring)
       
    34   apply (rule derivative_eq_intros setsum.cong | simp)+
       
    35   done
       
    36 
       
    37 lemma binomial_deriv2:
       
    38     "(\<Sum>k=0..n. (of_nat k * of_nat (k-1) * of_nat (n choose k)) * a^(k-2) * b^(n-k)) =
       
    39      of_nat n * of_nat (n-1) * (a+b::real) ^ (n-2)"
       
    40   apply (rule DERIV_unique [where f = "\<lambda>a. of_nat n * (a+b::real) ^ (n-1)" and x=a])
       
    41   apply (subst binomial_deriv1 [symmetric])
       
    42   apply (rule derivative_eq_intros setsum.cong | simp add: Num.numeral_2_eq_2)+
       
    43   done
       
    44 
       
    45 lemma sum_k_Bernstein [simp]: "(\<Sum>k = 0..n. real k * Bernstein n k x) = of_nat n * x"
       
    46   apply (subst binomial_deriv1 [of n x "1-x", simplified, symmetric])
       
    47   apply (simp add: setsum_left_distrib)
       
    48   apply (auto simp: Bernstein_def real_of_nat_def algebra_simps realpow_num_eq_if intro!: setsum.cong)
       
    49   done
       
    50 
       
    51 lemma sum_kk_Bernstein [simp]: "(\<Sum> k = 0..n. real k * (real k - 1) * Bernstein n k x) = real n * (real n - 1) * x\<^sup>2"
       
    52 proof -
       
    53   have "(\<Sum> k = 0..n. real k * (real k - 1) * Bernstein n k x) = real_of_nat n * real_of_nat (n - Suc 0) * x\<^sup>2"
       
    54     apply (subst binomial_deriv2 [of n x "1-x", simplified, symmetric])
       
    55     apply (simp add: setsum_left_distrib)
       
    56     apply (rule setsum.cong [OF refl])
       
    57     apply (simp add: Bernstein_def power2_eq_square algebra_simps real_of_nat_def)
       
    58     apply (rename_tac k)
       
    59     apply (subgoal_tac "k = 0 \<or> k = 1 \<or> (\<exists>k'. k = Suc (Suc k'))")
       
    60     apply (force simp add: field_simps of_nat_Suc power2_eq_square)
       
    61     by presburger
       
    62   also have "... = n * (n - 1) * x\<^sup>2"
       
    63     by auto
       
    64   finally show ?thesis
       
    65     by auto
       
    66 qed
       
    67 
       
    68 subsection {*Explicit Bernstein version of the 1D Weierstrass approximation theorem*}
       
    69 
       
    70 lemma Bernstein_Weierstrass:
       
    71   fixes f :: "real \<Rightarrow> real"
       
    72   assumes contf: "continuous_on {0..1} f" and e: "0 < e"
       
    73     shows "\<exists>N. \<forall>n x. N \<le> n \<and> x \<in> {0..1}
       
    74                     \<longrightarrow> abs(f x - (\<Sum>k = 0..n. f(k/n) * Bernstein n k x)) < e"
       
    75 proof -
       
    76   have "bounded (f ` {0..1})"
       
    77     using compact_continuous_image compact_imp_bounded contf by blast
       
    78   then obtain M where M: "\<And>x. 0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> \<bar>f x\<bar> \<le> M"
       
    79     by (force simp add: bounded_iff)
       
    80   then have Mge0: "0 \<le> M" by force
       
    81   have ucontf: "uniformly_continuous_on {0..1} f"
       
    82     using compact_uniformly_continuous contf by blast
       
    83   then obtain d where d: "d>0" "\<And>x x'. \<lbrakk> x \<in> {0..1}; x' \<in> {0..1}; \<bar>x' - x\<bar> < d\<rbrakk> \<Longrightarrow> \<bar>f x' - f x\<bar> < e/2"
       
    84      apply (rule uniformly_continuous_onE [where e = "e/2"])
       
    85      using e by (auto simp: dist_norm)
       
    86   { fix n::nat and x::real
       
    87     assume n: "Suc (nat\<lceil>4*M/(e*d\<^sup>2)\<rceil>) \<le> n" and x: "0 \<le> x" "x \<le> 1"
       
    88     have "0 < n" using n by simp
       
    89     have ed0: "- (e * d\<^sup>2) < 0"
       
    90       using e `0<d` by simp
       
    91     also have "... \<le> M * 4"
       
    92       using `0\<le>M` by simp
       
    93     finally have [simp]: "real (nat \<lceil>4 * M / (e * d\<^sup>2)\<rceil>) = real \<lceil>4 * M / (e * d\<^sup>2)\<rceil>"
       
    94       using `0\<le>M` e `0<d`
       
    95       by (simp add: Real.real_of_nat_Suc field_simps)
       
    96     have "4*M/(e*d\<^sup>2) + 1 \<le> real (Suc (nat\<lceil>4*M/(e*d\<^sup>2)\<rceil>))"
       
    97       by (simp add: Real.real_of_nat_Suc)
       
    98     also have "... \<le> real n"
       
    99       using n by (simp add: Real.real_of_nat_Suc field_simps)
       
   100     finally have nbig: "4*M/(e*d\<^sup>2) + 1 \<le> real n" .
       
   101     have sum_bern: "(\<Sum>k = 0..n. (x - k/n)\<^sup>2 * Bernstein n k x) = x * (1 - x) / n"
       
   102     proof -
       
   103       have *: "\<And>a b x::real. (a - b)\<^sup>2 * x = a * (a - 1) * x + (1 - 2 * b) * a * x + b * b * x"
       
   104         by (simp add: algebra_simps power2_eq_square)
       
   105       have "(\<Sum> k = 0..n. (k - n * x)\<^sup>2 * Bernstein n k x) = n * x * (1 - x)"
       
   106         apply (simp add: * setsum.distrib)
       
   107         apply (simp add: setsum_right_distrib [symmetric] mult.assoc)
       
   108         apply (simp add: algebra_simps power2_eq_square)
       
   109         done
       
   110       then have "(\<Sum> k = 0..n. (k - n * x)\<^sup>2 * Bernstein n k x)/n^2 = x * (1 - x) / n"
       
   111         by (simp add: power2_eq_square)
       
   112       then show ?thesis
       
   113         using n by (simp add: setsum_divide_distrib divide_simps mult.commute power2_commute)
       
   114     qed
       
   115     { fix k
       
   116       assume k: "k \<le> n"
       
   117       then have kn: "0 \<le> k / n" "k / n \<le> 1"
       
   118         by (auto simp: divide_simps)
       
   119       consider (lessd) "abs(x - k / n) < d" | (ged) "d \<le> abs(x - k / n)"
       
   120         by linarith
       
   121       then have "\<bar>(f x - f (k/n))\<bar> \<le> e/2 + 2 * M / d\<^sup>2 * (x - k/n)\<^sup>2"
       
   122       proof cases
       
   123         case lessd
       
   124         then have "\<bar>(f x - f (k/n))\<bar> < e/2"
       
   125           using d x kn by (simp add: abs_minus_commute)
       
   126         also have "... \<le> (e/2 + 2 * M / d\<^sup>2 * (x - k/n)\<^sup>2)"
       
   127           using Mge0 d by simp
       
   128         finally show ?thesis by simp
       
   129       next
       
   130         case ged
       
   131         then have dle: "d\<^sup>2 \<le> (x - k/n)\<^sup>2"
       
   132           by (metis d(1) less_eq_real_def power2_abs power_mono)
       
   133         have "\<bar>(f x - f (k/n))\<bar> \<le> \<bar>f x\<bar> + \<bar>f (k/n)\<bar>"
       
   134           by (rule abs_triangle_ineq4)
       
   135         also have "... \<le> M+M"
       
   136           by (meson M add_mono_thms_linordered_semiring(1) kn x)
       
   137         also have "... \<le> 2 * M * ((x - k/n)\<^sup>2 / d\<^sup>2)"
       
   138           apply simp
       
   139           apply (rule Rings.ordered_semiring_class.mult_left_mono [of 1 "((x - k/n)\<^sup>2 / d\<^sup>2)", simplified])
       
   140           using dle `d>0` `M\<ge>0` by auto
       
   141         also have "... \<le> e/2 + 2 * M / d\<^sup>2 * (x - k/n)\<^sup>2"
       
   142           using e  by simp
       
   143         finally show ?thesis .
       
   144         qed
       
   145     } note * = this
       
   146     have "\<bar>f x - (\<Sum> k = 0..n. f(k / n) * Bernstein n k x)\<bar> \<le> \<bar>\<Sum> k = 0..n. (f x - f(k / n)) * Bernstein n k x\<bar>"
       
   147       by (simp add: setsum_subtractf setsum_right_distrib [symmetric] algebra_simps)
       
   148     also have "... \<le> (\<Sum> k = 0..n. (e/2 + (2 * M / d\<^sup>2) * (x - k / n)\<^sup>2) * Bernstein n k x)"
       
   149       apply (rule order_trans [OF setsum_abs setsum_mono])
       
   150       using *
       
   151       apply (simp add: abs_mult Bernstein_nonneg x mult_right_mono)
       
   152       done
       
   153     also have "... \<le> e/2 + (2 * M) / (d\<^sup>2 * n)"
       
   154       apply (simp only: setsum.distrib Rings.semiring_class.distrib_right setsum_right_distrib [symmetric] mult.assoc sum_bern)
       
   155       using `d>0` x
       
   156       apply (simp add: divide_simps Mge0 mult_le_one mult_left_le)
       
   157       done
       
   158     also have "... < e"
       
   159       apply (simp add: field_simps)
       
   160       using `d>0` nbig e `n>0`
       
   161       apply (simp add: divide_simps algebra_simps)
       
   162       using ed0 by linarith
       
   163     finally have "\<bar>f x - (\<Sum>k = 0..n. f (real k / real n) * Bernstein n k x)\<bar> < e" .
       
   164   }
       
   165   then show ?thesis
       
   166     by auto
       
   167 qed
       
   168 
       
   169 
       
   170 subsection {*General Stone-Weierstrass theorem*}
       
   171 
       
   172 text\<open>Source:
       
   173 Bruno Brosowski and Frank Deutsch.
       
   174 An Elementary Proof of the Stone-Weierstrass Theorem.
       
   175 Proceedings of the American Mathematical Society
       
   176 Volume 81, Number 1, January 1981.\<close>
       
   177 
       
   178 locale function_ring_on =
       
   179   fixes R :: "('a::t2_space \<Rightarrow> real) set" and s :: "'a set"
       
   180   assumes compact: "compact s"
       
   181   assumes continuous: "f \<in> R \<Longrightarrow> continuous_on s f"
       
   182   assumes add: "f \<in> R \<Longrightarrow> g \<in> R \<Longrightarrow> (\<lambda>x. f x + g x) \<in> R"
       
   183   assumes mult: "f \<in> R \<Longrightarrow> g \<in> R \<Longrightarrow> (\<lambda>x. f x * g x) \<in> R"
       
   184   assumes const: "(\<lambda>_. c) \<in> R"
       
   185   assumes separable: "x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> x \<noteq> y \<Longrightarrow> \<exists>f\<in>R. f x \<noteq> f y"
       
   186 
       
   187 begin
       
   188   lemma minus: "f \<in> R \<Longrightarrow> (\<lambda>x. - f x) \<in> R"
       
   189     by (frule mult [OF const [of "-1"]]) simp
       
   190 
       
   191   lemma diff: "f \<in> R \<Longrightarrow> g \<in> R \<Longrightarrow> (\<lambda>x. f x - g x) \<in> R"
       
   192     unfolding diff_conv_add_uminus by (metis add minus)
       
   193 
       
   194   lemma power: "f \<in> R \<Longrightarrow> (\<lambda>x. f x ^ n) \<in> R"
       
   195     by (induct n) (auto simp: const mult)
       
   196 
       
   197   lemma setsum: "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> f i \<in> R\<rbrakk> \<Longrightarrow> (\<lambda>x. \<Sum>i \<in> I. f i x) \<in> R"
       
   198     by (induct I rule: finite_induct; simp add: const add)
       
   199 
       
   200   lemma setprod: "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> f i \<in> R\<rbrakk> \<Longrightarrow> (\<lambda>x. \<Prod>i \<in> I. f i x) \<in> R"
       
   201     by (induct I rule: finite_induct; simp add: const mult)
       
   202 
       
   203   definition normf :: "('a::t2_space \<Rightarrow> real) \<Rightarrow> real"
       
   204     where "normf f \<equiv> SUP x:s. \<bar>f x\<bar>"
       
   205 
       
   206   lemma normf_upper: "\<lbrakk>continuous_on s f; x \<in> s\<rbrakk> \<Longrightarrow> \<bar>f x\<bar> \<le> normf f"
       
   207     apply (simp add: normf_def)
       
   208     apply (rule cSUP_upper, assumption)
       
   209     by (simp add: bounded_imp_bdd_above compact compact_continuous_image compact_imp_bounded continuous_on_rabs)
       
   210 
       
   211 lemma normf_least: "s \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> s \<Longrightarrow> \<bar>f x\<bar> \<le> M) \<Longrightarrow> normf f \<le> M"
       
   212     by (simp add: normf_def cSUP_least)
       
   213 
       
   214 end
       
   215 
       
   216 lemma (in function_ring_on) one:
       
   217   assumes U: "open U" and t0: "t0 \<in> s" "t0 \<in> U" and t1: "t1 \<in> s-U"
       
   218     shows "\<exists>V. open V \<and> t0 \<in> V \<and> s \<inter> V \<subseteq> U \<and>
       
   219                (\<forall>e>0. \<exists>f \<in> R. f ` s \<subseteq> {0..1} \<and> (\<forall>t \<in> s \<inter> V. f t < e) \<and> (\<forall>t \<in> s - U. f t > 1 - e))"
       
   220 proof -
       
   221   have "\<exists>pt \<in> R. pt t0 = 0 \<and> pt t > 0 \<and> pt ` s \<subseteq> {0..1}" if t: "t \<in> s - U" for t
       
   222   proof -
       
   223     have "t \<noteq> t0" using t t0 by auto
       
   224     then obtain g where g: "g \<in> R" "g t \<noteq> g t0"
       
   225       using separable t0  by (metis Diff_subset subset_eq t)
       
   226     def h \<equiv> "\<lambda>x. g x - g t0"
       
   227     have "h \<in> R"
       
   228       unfolding h_def by (fast intro: g const diff)
       
   229     then have hsq: "(\<lambda>w. (h w)\<^sup>2) \<in> R"
       
   230       by (simp add: power2_eq_square mult)
       
   231     have "h t \<noteq> h t0"
       
   232       by (simp add: h_def g)
       
   233     then have "h t \<noteq> 0"
       
   234       by (simp add: h_def)
       
   235     then have ht2: "0 < (h t)^2"
       
   236       by simp
       
   237     also have "... \<le> normf (\<lambda>w. (h w)\<^sup>2)"
       
   238       using t normf_upper [where x=t] continuous [OF hsq] by force
       
   239     finally have nfp: "0 < normf (\<lambda>w. (h w)\<^sup>2)" .
       
   240     def p \<equiv> "\<lambda>x. (1 / normf (\<lambda>w. (h w)\<^sup>2)) * (h x)^2"
       
   241     have "p \<in> R"
       
   242       unfolding p_def by (fast intro: hsq const mult)
       
   243     moreover have "p t0 = 0"
       
   244       by (simp add: p_def h_def)
       
   245     moreover have "p t > 0"
       
   246       using nfp ht2 by (simp add: p_def)
       
   247     moreover have "\<And>x. x \<in> s \<Longrightarrow> p x \<in> {0..1}"
       
   248       using nfp normf_upper [OF continuous [OF hsq] ] by (auto simp: p_def)
       
   249     ultimately show "\<exists>pt \<in> R. pt t0 = 0 \<and> pt t > 0 \<and> pt ` s \<subseteq> {0..1}"
       
   250       by auto
       
   251   qed
       
   252   then obtain pf where pf: "\<And>t. t \<in> s-U \<Longrightarrow> pf t \<in> R \<and> pf t t0 = 0 \<and> pf t t > 0"
       
   253                    and pf01: "\<And>t. t \<in> s-U \<Longrightarrow> pf t ` s \<subseteq> {0..1}"
       
   254     by metis
       
   255   have com_sU: "compact (s-U)"
       
   256     using compact closed_inter_compact U by (simp add: Diff_eq compact_inter_closed open_closed)
       
   257   have "\<And>t. t \<in> s-U \<Longrightarrow> \<exists>A. open A \<and> A \<inter> s = {x\<in>s. 0 < pf t x}"
       
   258     apply (rule open_Collect_positive)
       
   259     by (metis pf continuous)
       
   260   then obtain Uf where Uf: "\<And>t. t \<in> s-U \<Longrightarrow> open (Uf t) \<and> (Uf t) \<inter> s = {x\<in>s. 0 < pf t x}"
       
   261     by metis
       
   262   then have open_Uf: "\<And>t. t \<in> s-U \<Longrightarrow> open (Uf t)"
       
   263     by blast
       
   264   have tUft: "\<And>t. t \<in> s-U \<Longrightarrow> t \<in> Uf t"
       
   265     using pf Uf by blast
       
   266   then have *: "s-U \<subseteq> (\<Union>x \<in> s-U. Uf x)"
       
   267     by blast
       
   268   obtain subU where subU: "subU \<subseteq> s - U" "finite subU" "s - U \<subseteq> (\<Union>x \<in> subU. Uf x)"
       
   269     by (blast intro: that open_Uf compactE_image [OF com_sU _ *])
       
   270   then have [simp]: "subU \<noteq> {}"
       
   271     using t1 by auto
       
   272   then have cardp: "card subU > 0" using subU
       
   273     by (simp add: card_gt_0_iff)
       
   274   def p \<equiv> "\<lambda>x. (1 / card subU) * (\<Sum>t \<in> subU. pf t x)"
       
   275   have pR: "p \<in> R"
       
   276     unfolding p_def using subU pf by (fast intro: pf const mult setsum)
       
   277   have pt0 [simp]: "p t0 = 0"
       
   278     using subU pf by (auto simp: p_def intro: setsum.neutral)
       
   279   have pt_pos: "p t > 0" if t: "t \<in> s-U" for t
       
   280   proof -
       
   281     obtain i where i: "i \<in> subU" "t \<in> Uf i" using subU t by blast
       
   282     show ?thesis
       
   283       using subU i t
       
   284       apply (clarsimp simp: p_def divide_simps)
       
   285       apply (rule setsum_pos2 [OF `finite subU`])
       
   286       using Uf t pf01 apply auto
       
   287       apply (force elim!: subsetCE)
       
   288       done
       
   289   qed
       
   290   have p01: "p x \<in> {0..1}" if t: "x \<in> s" for x
       
   291   proof -
       
   292     have "0 \<le> p x"
       
   293       using subU cardp t
       
   294       apply (simp add: p_def divide_simps setsum_nonneg)
       
   295       apply (rule setsum_nonneg)
       
   296       using pf01 by force
       
   297     moreover have "p x \<le> 1"
       
   298       using subU cardp t
       
   299       apply (simp add: p_def divide_simps setsum_nonneg real_of_nat_def)
       
   300       apply (rule setsum_bounded_above [where 'a=real and K=1, simplified])
       
   301       using pf01 by force
       
   302     ultimately show ?thesis
       
   303       by auto
       
   304   qed
       
   305   have "compact (p ` (s-U))"
       
   306     by (meson Diff_subset com_sU compact_continuous_image continuous continuous_on_subset pR)
       
   307   then have "open (- (p ` (s-U)))"
       
   308     by (simp add: compact_imp_closed open_Compl)
       
   309   moreover have "0 \<in> - (p ` (s-U))"
       
   310     by (metis (no_types) ComplI image_iff not_less_iff_gr_or_eq pt_pos)
       
   311   ultimately obtain delta0 where delta0: "delta0 > 0" "ball 0 delta0 \<subseteq> - (p ` (s-U))"
       
   312     by (auto simp: elim!: openE)
       
   313   then have pt_delta: "\<And>x. x \<in> s-U \<Longrightarrow> p x \<ge> delta0"
       
   314     by (force simp: ball_def dist_norm dest: p01)
       
   315   def \<delta> \<equiv> "delta0/2"
       
   316   have "delta0 \<le> 1" using delta0 p01 [of t1] t1
       
   317       by (force simp: ball_def dist_norm dest: p01)
       
   318   with delta0 have \<delta>01: "0 < \<delta>" "\<delta> < 1"
       
   319     by (auto simp: \<delta>_def)
       
   320   have pt_\<delta>: "\<And>x. x \<in> s-U \<Longrightarrow> p x \<ge> \<delta>"
       
   321     using pt_delta delta0 by (force simp: \<delta>_def)
       
   322   have "\<exists>A. open A \<and> A \<inter> s = {x\<in>s. p x < \<delta>/2}"
       
   323     by (rule open_Collect_less_Int [OF continuous [OF pR] continuous_on_const])
       
   324   then obtain V where V: "open V" "V \<inter> s = {x\<in>s. p x < \<delta>/2}"
       
   325     by blast
       
   326   def k \<equiv> "nat\<lfloor>1/\<delta>\<rfloor> + 1"
       
   327   have "k>0"  by (simp add: k_def)
       
   328   have "k-1 \<le> 1/\<delta>"
       
   329     using \<delta>01 by (simp add: k_def)
       
   330   with \<delta>01 have "k \<le> (1+\<delta>)/\<delta>"
       
   331     by (auto simp: algebra_simps add_divide_distrib)
       
   332   also have "... < 2/\<delta>"
       
   333     using \<delta>01 by (auto simp: divide_simps)
       
   334   finally have k2\<delta>: "k < 2/\<delta>" .
       
   335   have "1/\<delta> < k"
       
   336     using \<delta>01 unfolding k_def by linarith
       
   337   with \<delta>01 k2\<delta> have k\<delta>: "1 < k*\<delta>" "k*\<delta> < 2"
       
   338     by (auto simp: divide_simps)
       
   339   def q \<equiv> "\<lambda>n t. (1 - p t ^ n) ^ (k^n)"
       
   340   have qR: "q n \<in> R" for n
       
   341     by (simp add: q_def const diff power pR)
       
   342   have q01: "\<And>n t. t \<in> s \<Longrightarrow> q n t \<in> {0..1}"
       
   343     using p01 by (simp add: q_def power_le_one algebra_simps)
       
   344   have qt0 [simp]: "\<And>n. n>0 \<Longrightarrow> q n t0 = 1"
       
   345     using t0 pf by (simp add: q_def power_0_left)
       
   346   { fix t and n::nat
       
   347     assume t: "t \<in> s \<inter> V"
       
   348     with `k>0` V have "k * p t < k * \<delta> / 2"
       
   349        by force
       
   350     then have "1 - (k * \<delta> / 2)^n \<le> 1 - (k * p t)^n"
       
   351       using  `k>0` p01 t by (simp add: power_mono)
       
   352     also have "... \<le> q n t"
       
   353       using Bernoulli_inequality [of "- ((p t)^n)" "k^n"]
       
   354       apply (simp add: q_def)
       
   355       by (metis IntE atLeastAtMost_iff p01 power_le_one power_mult_distrib t)
       
   356     finally have "1 - (k * \<delta> / 2) ^ n \<le> q n t" .
       
   357   } note limitV = this
       
   358   { fix t and n::nat
       
   359     assume t: "t \<in> s - U"
       
   360     with `k>0` U have "k * \<delta> \<le> k * p t"
       
   361       by (simp add: pt_\<delta>)
       
   362     with k\<delta> have kpt: "1 < k * p t"
       
   363       by (blast intro: less_le_trans)
       
   364     have ptn_pos: "0 < p t ^ n"
       
   365       using pt_pos [OF t] by simp
       
   366     have ptn_le: "p t ^ n \<le> 1"
       
   367       by (meson DiffE atLeastAtMost_iff p01 power_le_one t)
       
   368     have "q n t = (1/(k^n * (p t)^n)) * (1 - p t ^ n) ^ (k^n) * k^n * (p t)^n"
       
   369       using pt_pos [OF t] `k>0` by (simp add: q_def)
       
   370     also have "... \<le> (1/(k * (p t))^n) * (1 - p t ^ n) ^ (k^n) * (1 + k^n * (p t)^n)"
       
   371       using pt_pos [OF t] `k>0`
       
   372       apply simp
       
   373       apply (simp only: times_divide_eq_right [symmetric])
       
   374       apply (rule mult_left_mono [of "1::real", simplified])
       
   375       apply (simp_all add: power_mult_distrib)
       
   376       apply (rule zero_le_power)
       
   377       using ptn_le by linarith
       
   378     also have "... \<le> (1/(k * (p t))^n) * (1 - p t ^ n) ^ (k^n) * (1 + (p t)^n) ^ (k^n)"
       
   379       apply (rule mult_left_mono [OF Bernoulli_inequality [of "p t ^ n" "k^n"]])
       
   380       using `k>0` ptn_pos ptn_le
       
   381       apply (auto simp: power_mult_distrib)
       
   382       done
       
   383     also have "... = (1/(k * (p t))^n) * (1 - p t ^ (2*n)) ^ (k^n)"
       
   384       using pt_pos [OF t] `k>0`
       
   385       by (simp add: algebra_simps power_mult power2_eq_square power_mult_distrib [symmetric])
       
   386     also have "... \<le> (1/(k * (p t))^n) * 1"
       
   387       apply (rule mult_left_mono [OF power_le_one])
       
   388       apply (metis diff_le_iff(1) less_eq_real_def mult.commute power_le_one power_mult ptn_pos ptn_le)
       
   389       using pt_pos [OF t] `k>0`
       
   390       apply auto
       
   391       done
       
   392     also have "... \<le> (1 / (k*\<delta>))^n"
       
   393       using `k>0` \<delta>01  power_mono pt_\<delta> t
       
   394       by (fastforce simp: field_simps)
       
   395     finally have "q n t \<le> (1 / (real k * \<delta>)) ^ n " .
       
   396   } note limitNonU = this
       
   397   def NN \<equiv> "\<lambda>e. 1 + nat \<lceil>max (ln e / ln (real k * \<delta> / 2)) (- ln e / ln (real k * \<delta>))\<rceil>"
       
   398   have NN: "of_nat (NN e) > ln e / ln (real k * \<delta> / 2)"  "of_nat (NN e) > - ln e / ln (real k * \<delta>)"
       
   399               if "0<e" for e
       
   400       unfolding NN_def  by linarith+
       
   401   have NN1: "\<And>e. e>0 \<Longrightarrow> (k * \<delta> / 2)^NN e < e"
       
   402     apply (subst Transcendental.ln_less_cancel_iff [symmetric])
       
   403       prefer 3 apply (subst ln_realpow)
       
   404     using `k>0` `\<delta>>0` NN  k\<delta>
       
   405     apply (force simp add: field_simps)+
       
   406     done
       
   407   have NN0: "\<And>e. e>0 \<Longrightarrow> (1/(k*\<delta>))^NN e < e"
       
   408     apply (subst Transcendental.ln_less_cancel_iff [symmetric])
       
   409       prefer 3 apply (subst ln_realpow)
       
   410     using `k>0` `\<delta>>0` NN k\<delta>
       
   411     apply (force simp add: field_simps ln_div)+
       
   412     done
       
   413   { fix t and e::real
       
   414     assume "e>0"
       
   415     have "t \<in> s \<inter> V \<Longrightarrow> 1 - q (NN e) t < e" "t \<in> s - U \<Longrightarrow> q (NN e) t < e"
       
   416     proof -
       
   417       assume t: "t \<in> s \<inter> V"
       
   418       show "1 - q (NN e) t < e"
       
   419         by (metis add.commute diff_le_eq not_le limitV [OF t] less_le_trans [OF NN1 [OF `e>0`]])
       
   420     next
       
   421       assume t: "t \<in> s - U"
       
   422       show "q (NN e) t < e"
       
   423       using  limitNonU [OF t] less_le_trans [OF NN0 [OF `e>0`]] not_le by blast
       
   424     qed
       
   425   } then have "\<And>e. e > 0 \<Longrightarrow> \<exists>f\<in>R. f ` s \<subseteq> {0..1} \<and> (\<forall>t \<in> s \<inter> V. f t < e) \<and> (\<forall>t \<in> s - U. 1 - e < f t)"
       
   426     using q01
       
   427     by (rule_tac x="\<lambda>x. 1 - q (NN e) x" in bexI) (auto simp: algebra_simps intro: diff const qR)
       
   428   moreover have t0V: "t0 \<in> V"  "s \<inter> V \<subseteq> U"
       
   429     using pt_\<delta> t0 U V \<delta>01  by fastforce+
       
   430   ultimately show ?thesis using V t0V
       
   431     by blast
       
   432 qed
       
   433 
       
   434 text\<open>Non-trivial case, with @{term A} and @{term B} both non-empty\<close>
       
   435 lemma (in function_ring_on) two_special:
       
   436   assumes A: "closed A" "A \<subseteq> s" "a \<in> A"
       
   437       and B: "closed B" "B \<subseteq> s" "b \<in> B"
       
   438       and disj: "A \<inter> B = {}"
       
   439       and e: "0 < e" "e < 1"
       
   440     shows "\<exists>f \<in> R. f ` s \<subseteq> {0..1} \<and> (\<forall>x \<in> A. f x < e) \<and> (\<forall>x \<in> B. f x > 1 - e)"
       
   441 proof -
       
   442   { fix w
       
   443     assume "w \<in> A"
       
   444     then have "open ( - B)" "b \<in> s" "w \<notin> B" "w \<in> s"
       
   445       using assms by auto
       
   446     then have "\<exists>V. open V \<and> w \<in> V \<and> s \<inter> V \<subseteq> -B \<and>
       
   447                (\<forall>e>0. \<exists>f \<in> R. f ` s \<subseteq> {0..1} \<and> (\<forall>x \<in> s \<inter> V. f x < e) \<and> (\<forall>x \<in> s \<inter> B. f x > 1 - e))"
       
   448       using one [of "-B" w b] assms `w \<in> A` by simp
       
   449   }
       
   450   then obtain Vf where Vf:
       
   451          "\<And>w. w \<in> A \<Longrightarrow> open (Vf w) \<and> w \<in> Vf w \<and> s \<inter> Vf w \<subseteq> -B \<and>
       
   452                          (\<forall>e>0. \<exists>f \<in> R. f ` s \<subseteq> {0..1} \<and> (\<forall>x \<in> s \<inter> Vf w. f x < e) \<and> (\<forall>x \<in> s \<inter> B. f x > 1 - e))"
       
   453     by metis
       
   454   then have open_Vf: "\<And>w. w \<in> A \<Longrightarrow> open (Vf w)"
       
   455     by blast
       
   456   have tVft: "\<And>w. w \<in> A \<Longrightarrow> w \<in> Vf w"
       
   457     using Vf by blast
       
   458   then have setsum_max_0: "A \<subseteq> (\<Union>x \<in> A. Vf x)"
       
   459     by blast
       
   460   have com_A: "compact A" using A
       
   461     by (metis compact compact_inter_closed inf.absorb_iff2)
       
   462   obtain subA where subA: "subA \<subseteq> A" "finite subA" "A \<subseteq> (\<Union>x \<in> subA. Vf x)"
       
   463     by (blast intro: that open_Vf compactE_image [OF com_A _ setsum_max_0])
       
   464   then have [simp]: "subA \<noteq> {}"
       
   465     using `a \<in> A` by auto
       
   466   then have cardp: "card subA > 0" using subA
       
   467     by (simp add: card_gt_0_iff)
       
   468   have "\<And>w. w \<in> A \<Longrightarrow> \<exists>f \<in> R. f ` s \<subseteq> {0..1} \<and> (\<forall>x \<in> s \<inter> Vf w. f x < e / card subA) \<and> (\<forall>x \<in> s \<inter> B. f x > 1 - e / card subA)"
       
   469     using Vf e cardp by simp
       
   470   then obtain ff where ff:
       
   471          "\<And>w. w \<in> A \<Longrightarrow> ff w \<in> R \<and> ff w ` s \<subseteq> {0..1} \<and>
       
   472                          (\<forall>x \<in> s \<inter> Vf w. ff w x < e / card subA) \<and> (\<forall>x \<in> s \<inter> B. ff w x > 1 - e / card subA)"
       
   473     by metis
       
   474   def pff \<equiv> "\<lambda>x. (\<Prod>w \<in> subA. ff w x)"
       
   475   have pffR: "pff \<in> R"
       
   476     unfolding pff_def using subA ff by (auto simp: intro: setprod)
       
   477   moreover
       
   478   have pff01: "pff x \<in> {0..1}" if t: "x \<in> s" for x
       
   479   proof -
       
   480     have "0 \<le> pff x"
       
   481       using subA cardp t
       
   482       apply (simp add: pff_def divide_simps setsum_nonneg)
       
   483       apply (rule Groups_Big.linordered_semidom_class.setprod_nonneg)
       
   484       using ff by fastforce
       
   485     moreover have "pff x \<le> 1"
       
   486       using subA cardp t
       
   487       apply (simp add: pff_def divide_simps setsum_nonneg real_of_nat_def)
       
   488       apply (rule setprod_mono [where g = "\<lambda>x. 1", simplified])
       
   489       using ff by fastforce
       
   490     ultimately show ?thesis
       
   491       by auto
       
   492   qed
       
   493   moreover
       
   494   { fix v x
       
   495     assume v: "v \<in> subA" and x: "x \<in> Vf v" "x \<in> s"
       
   496     from subA v have "pff x = ff v x * (\<Prod>w \<in> subA - {v}. ff w x)"
       
   497       unfolding pff_def  by (metis setprod.remove)
       
   498     also have "... \<le> ff v x * 1"
       
   499       apply (rule Rings.ordered_semiring_class.mult_left_mono)
       
   500       apply (rule setprod_mono [where g = "\<lambda>x. 1", simplified])
       
   501       using ff [THEN conjunct2, THEN conjunct1] v subA x
       
   502       apply auto
       
   503       apply (meson atLeastAtMost_iff contra_subsetD imageI)
       
   504       apply (meson atLeastAtMost_iff contra_subsetD image_eqI)
       
   505       using atLeastAtMost_iff by blast
       
   506     also have "... < e / card subA"
       
   507       using ff [THEN conjunct2, THEN conjunct2, THEN conjunct1] v subA x
       
   508       by auto
       
   509     also have "... \<le> e"
       
   510       using cardp e by (simp add: divide_simps)
       
   511     finally have "pff x < e" .
       
   512   }
       
   513   then have "\<And>x. x \<in> A \<Longrightarrow> pff x < e"
       
   514     using A Vf subA by (metis UN_E contra_subsetD)
       
   515   moreover
       
   516   { fix x
       
   517     assume x: "x \<in> B"
       
   518     then have "x \<in> s"
       
   519       using B by auto
       
   520     have "1 - e \<le> (1 - e / card subA) ^ card subA"
       
   521       using Bernoulli_inequality [of "-e / card subA" "card subA"] e cardp
       
   522       by (auto simp: field_simps)
       
   523     also have "... = (\<Prod>w \<in> subA. 1 - e / card subA)"
       
   524       by (simp add: setprod_constant subA(2))
       
   525     also have "... < pff x"
       
   526       apply (simp add: pff_def)
       
   527       apply (rule setprod_mono_strict [where f = "\<lambda>x. 1 - e / card subA", simplified])
       
   528       apply (simp_all add: subA(2))
       
   529       apply (intro ballI conjI)
       
   530       using e apply (force simp: divide_simps)
       
   531       using ff [THEN conjunct2, THEN conjunct2, THEN conjunct2] subA B x
       
   532       apply blast
       
   533       done
       
   534     finally have "1 - e < pff x" .
       
   535   }
       
   536   ultimately
       
   537   show ?thesis by blast
       
   538 qed
       
   539 
       
   540 lemma (in function_ring_on) two:
       
   541   assumes A: "closed A" "A \<subseteq> s"
       
   542       and B: "closed B" "B \<subseteq> s"
       
   543       and disj: "A \<inter> B = {}"
       
   544       and e: "0 < e" "e < 1"
       
   545     shows "\<exists>f \<in> R. f ` s \<subseteq> {0..1} \<and> (\<forall>x \<in> A. f x < e) \<and> (\<forall>x \<in> B. f x > 1 - e)"
       
   546 proof (cases "A \<noteq> {} \<and> B \<noteq> {}")
       
   547   case True then show ?thesis
       
   548     apply (simp add: ex_in_conv [symmetric])
       
   549     using assms
       
   550     apply safe
       
   551     apply (force simp add: intro!: two_special)
       
   552     done
       
   553 next
       
   554   case False with e show ?thesis
       
   555     apply simp
       
   556     apply (erule disjE)
       
   557     apply (rule_tac [2] x="\<lambda>x. 0" in bexI)
       
   558     apply (rule_tac x="\<lambda>x. 1" in bexI)
       
   559     apply (auto simp: const)
       
   560     done
       
   561 qed
       
   562 
       
   563 text\<open>The special case where @{term f} is non-negative and @{term"e<1/3"}\<close>
       
   564 lemma (in function_ring_on) Stone_Weierstrass_special:
       
   565   assumes f: "continuous_on s f" and fpos: "\<And>x. x \<in> s \<Longrightarrow> f x \<ge> 0"
       
   566       and e: "0 < e" "e < 1/3"
       
   567   shows "\<exists>g \<in> R. \<forall>x\<in>s. \<bar>f x - g x\<bar> < 2*e"
       
   568 proof -
       
   569   def n \<equiv> "1 + nat \<lceil>normf f / e\<rceil>"
       
   570   def A \<equiv> "\<lambda>j::nat. {x \<in> s. f x \<le> (j - 1/3)*e}"
       
   571   def B \<equiv> "\<lambda>j::nat. {x \<in> s. f x \<ge> (j + 1/3)*e}"
       
   572   have ngt: "(n-1) * e \<ge> normf f" "n\<ge>1"
       
   573     using e
       
   574     apply (simp_all add: n_def field_simps real_of_nat_Suc)
       
   575     by (metis real_nat_ceiling_ge mult.commute not_less pos_less_divide_eq)
       
   576   then have ge_fx: "(n-1) * e \<ge> f x" if "x \<in> s" for x
       
   577     using f normf_upper that by fastforce
       
   578   { fix j
       
   579     have A: "closed (A j)" "A j \<subseteq> s"
       
   580       apply (simp_all add: A_def Collect_restrict)
       
   581       apply (rule continuous_on_closed_Collect_le [OF f continuous_on_const])
       
   582       apply (simp add: compact compact_imp_closed)
       
   583       done
       
   584     have B: "closed (B j)" "B j \<subseteq> s"
       
   585       apply (simp_all add: B_def Collect_restrict)
       
   586       apply (rule continuous_on_closed_Collect_le [OF continuous_on_const f])
       
   587       apply (simp add: compact compact_imp_closed)
       
   588       done
       
   589     have disj: "(A j) \<inter> (B j) = {}"
       
   590       using e by (auto simp: A_def B_def field_simps)
       
   591     have "\<exists>f \<in> R. f ` s \<subseteq> {0..1} \<and> (\<forall>x \<in> A j. f x < e/n) \<and> (\<forall>x \<in> B j. f x > 1 - e/n)"
       
   592       apply (rule two)
       
   593       using e A B disj ngt
       
   594       apply simp_all
       
   595       done
       
   596   }
       
   597   then obtain xf where xfR: "\<And>j. xf j \<in> R" and xf01: "\<And>j. xf j ` s \<subseteq> {0..1}"
       
   598                    and xfA: "\<And>x j. x \<in> A j \<Longrightarrow> xf j x < e/n"
       
   599                    and xfB: "\<And>x j. x \<in> B j \<Longrightarrow> xf j x > 1 - e/n"
       
   600     by metis
       
   601   def g \<equiv> "\<lambda>x. e * (\<Sum>i\<le>n. xf i x)"
       
   602   have gR: "g \<in> R"
       
   603     unfolding g_def by (fast intro: mult const setsum xfR)
       
   604   have gge0: "\<And>x. x \<in> s \<Longrightarrow> g x \<ge> 0"
       
   605     using e xf01 by (simp add: g_def zero_le_mult_iff image_subset_iff setsum_nonneg)
       
   606   have A0: "A 0 = {}"
       
   607     using fpos e by (fastforce simp: A_def)
       
   608   have An: "A n = s"
       
   609     using e ngt f normf_upper by (fastforce simp: A_def field_simps real_of_nat_diff)
       
   610   have Asub: "A j \<subseteq> A i" if "i\<ge>j" for i j
       
   611     using e that apply (clarsimp simp: A_def)
       
   612     apply (erule order_trans, simp)
       
   613     done
       
   614   { fix t
       
   615     assume t: "t \<in> s"
       
   616     def j \<equiv> "LEAST j. t \<in> A j"
       
   617     have jn: "j \<le> n"
       
   618       using t An by (simp add: Least_le j_def)
       
   619     have Aj: "t \<in> A j"
       
   620       using t An by (fastforce simp add: j_def intro: LeastI)
       
   621     then have Ai: "t \<in> A i" if "i\<ge>j" for i
       
   622       using Asub [OF that] by blast
       
   623     then have fj1: "f t \<le> (j - 1/3)*e"
       
   624       by (simp add: A_def)
       
   625     then have Anj: "t \<notin> A i" if "i<j" for i
       
   626       using  Aj  `i<j`
       
   627       apply (simp add: j_def)
       
   628       using not_less_Least by blast
       
   629     have j1: "1 \<le> j"
       
   630       using A0 Aj j_def not_less_eq_eq by (fastforce simp add: j_def)
       
   631     then have Anj: "t \<notin> A (j-1)"
       
   632       using Least_le by (fastforce simp add: j_def)
       
   633     then have fj2: "(j - 4/3)*e < f t"
       
   634       using j1 t  by (simp add: A_def real_of_nat_diff)
       
   635     have ***: "xf i t \<le> e/n" if "i\<ge>j" for i
       
   636       using xfA [OF Ai] that by (simp add: less_eq_real_def)
       
   637     { fix i
       
   638       assume "i+2 \<le> j"
       
   639       then obtain d where "i+2+d = j"
       
   640         using le_Suc_ex that by blast
       
   641       then have "t \<in> B i"
       
   642         using Anj e ge_fx [OF t] `1 \<le> n` fpos [OF t] t
       
   643         apply (simp add: A_def B_def)
       
   644         apply (clarsimp simp add: field_simps real_of_nat_diff not_le real_of_nat_Suc)
       
   645         apply (rule order_trans [of _ "e * 2 + (e * (real d * 3) + e * (real i * 3))"])
       
   646         apply auto
       
   647         done
       
   648       then have "xf i t > 1 - e/n"
       
   649         by (rule xfB)
       
   650     } note **** = this
       
   651     have xf_le1: "\<And>i. xf i t \<le> 1"
       
   652       using xf01 t by force
       
   653     have "g t = e * (\<Sum>i<j. xf i t) + e * (\<Sum>i=j..n. xf i t)"
       
   654       using j1 jn e
       
   655       apply (simp add: g_def distrib_left [symmetric])
       
   656       apply (subst setsum.union_disjoint [symmetric])
       
   657       apply (auto simp: ivl_disj_un)
       
   658       done
       
   659     also have "... \<le> e*j + e * ((Suc n - j)*e/n)"
       
   660       apply (rule add_mono)
       
   661       apply (simp_all only: mult_le_cancel_left_pos e real_of_nat_def)
       
   662       apply (rule setsum_bounded_above [OF xf_le1, where A = "lessThan j", simplified])
       
   663       using setsum_bounded_above [of "{j..n}" "\<lambda>i. xf i t", OF ***]
       
   664       apply simp
       
   665       done
       
   666     also have "... \<le> j*e + e*(n - j + 1)*e/n "
       
   667       using `1 \<le> n` e  by (simp add: field_simps)
       
   668     also have "... \<le> j*e + e*e"
       
   669       using `1 \<le> n` e j1 by (simp add: field_simps)
       
   670     also have "... < (j + 1/3)*e"
       
   671       using e by (auto simp: field_simps)
       
   672     finally have gj1: "g t < (j + 1 / 3) * e" .
       
   673     have gj2: "(j - 4/3)*e < g t"
       
   674     proof (cases "2 \<le> j")
       
   675       case False
       
   676       then have "j=1" using j1 by simp
       
   677       with t gge0 e show ?thesis by force
       
   678     next
       
   679       case True
       
   680       then have "(j - 4/3)*e < (j-1)*e - e^2"
       
   681         using e by (auto simp: real_of_nat_diff algebra_simps power2_eq_square)
       
   682       also have "... < (j-1)*e - ((j - 1)/n) * e^2"
       
   683         using e True jn by (simp add: power2_eq_square field_simps)
       
   684       also have "... = e * (j-1) * (1 - e/n)"
       
   685         by (simp add: power2_eq_square field_simps)
       
   686       also have "... \<le> e * (\<Sum>i\<le>j-2. xf i t)"
       
   687         using e
       
   688         apply simp
       
   689         apply (rule order_trans [OF _ setsum_bounded_below [OF less_imp_le [OF ****]]])
       
   690         using True
       
   691         apply (simp_all add: real_of_nat_def of_nat_Suc of_nat_diff)
       
   692         done
       
   693       also have "... \<le> g t"
       
   694         using jn e
       
   695         using e xf01 t
       
   696         apply (simp add: g_def zero_le_mult_iff image_subset_iff setsum_nonneg)
       
   697         apply (rule Groups_Big.setsum_mono2, auto)
       
   698         done
       
   699       finally show ?thesis .
       
   700     qed
       
   701     have "\<bar>f t - g t\<bar> < 2 * e"
       
   702       using fj1 fj2 gj1 gj2 by (simp add: abs_less_iff field_simps)
       
   703   }
       
   704   then show ?thesis
       
   705     by (rule_tac x=g in bexI) (auto intro: gR)
       
   706 qed
       
   707 
       
   708 text\<open>The ``unpretentious'' formulation\<close>
       
   709 lemma (in function_ring_on) Stone_Weierstrass_basic:
       
   710   assumes f: "continuous_on s f" and e: "e > 0"
       
   711   shows "\<exists>g \<in> R. \<forall>x\<in>s. \<bar>f x - g x\<bar> < e"
       
   712 proof -
       
   713   have "\<exists>g \<in> R. \<forall>x\<in>s. \<bar>(f x + normf f) - g x\<bar> < 2 * min (e/2) (1/4)"
       
   714     apply (rule Stone_Weierstrass_special)
       
   715     apply (rule Limits.continuous_on_add [OF f Topological_Spaces.continuous_on_const])
       
   716     using normf_upper [OF f] apply force
       
   717     apply (simp add: e, linarith)
       
   718     done
       
   719   then obtain g where "g \<in> R" "\<forall>x\<in>s. \<bar>g x - (f x + normf f)\<bar> < e"
       
   720     by force
       
   721   then show ?thesis
       
   722     apply (rule_tac x="\<lambda>x. g x - normf f" in bexI)
       
   723     apply (auto simp: algebra_simps intro: diff const)
       
   724     done
       
   725 qed
       
   726 
       
   727 
       
   728 theorem (in function_ring_on) Stone_Weierstrass:
       
   729   assumes f: "continuous_on s f"
       
   730   shows "\<exists>F\<in>UNIV \<rightarrow> R. LIM n sequentially. F n :> uniformly_on s f"
       
   731 proof -
       
   732   { fix e::real
       
   733     assume e: "0 < e"
       
   734     then obtain N::nat where N: "0 < N" "0 < inverse N" "inverse N < e"
       
   735       by (auto simp: real_arch_inv [of e])
       
   736     { fix n :: nat and x :: 'a and g :: "'a \<Rightarrow> real"
       
   737       assume n: "N \<le> n"  "\<forall>x\<in>s. \<bar>f x - g x\<bar> < 1 / (1 + real n)"
       
   738       assume x: "x \<in> s"
       
   739       have "\<not> real (Suc n) < inverse e"
       
   740         using `N \<le> n` N using less_imp_inverse_less by force
       
   741       then have "1 / (1 + real n) \<le> e"
       
   742         using e by (simp add: field_simps real_of_nat_Suc)
       
   743       then have "\<bar>f x - g x\<bar> < e"
       
   744         using n(2) x by auto
       
   745     } note * = this
       
   746     have "\<forall>\<^sub>F n in sequentially. \<forall>x\<in>s. \<bar>f x - (SOME g. g \<in> R \<and> (\<forall>x\<in>s. \<bar>f x - g x\<bar> < 1 / (1 + real n))) x\<bar> < e"
       
   747       apply (rule eventually_sequentiallyI [of N])
       
   748       apply (auto intro: someI2_bex [OF Stone_Weierstrass_basic [OF f]] *)
       
   749       done
       
   750   } then
       
   751   show ?thesis
       
   752     apply (rule_tac x="\<lambda>n::nat. SOME g. g \<in> R \<and> (\<forall>x\<in>s. \<bar>f x - g x\<bar> < 1 / (1 + n))" in bexI)
       
   753     prefer 2  apply (force intro: someI2_bex [OF Stone_Weierstrass_basic [OF f]])
       
   754     unfolding uniform_limit_iff
       
   755     apply (auto simp: dist_norm abs_minus_commute)
       
   756     done
       
   757 qed
       
   758 
       
   759 text{*A HOL Light formulation*}
       
   760 corollary Stone_Weierstrass_HOL:
       
   761   fixes R :: "('a::t2_space \<Rightarrow> real) set" and s :: "'a set"
       
   762   assumes "compact s"  "\<And>c. P(\<lambda>x. c::real)"
       
   763           "\<And>f. P f \<Longrightarrow> continuous_on s f"
       
   764           "\<And>f g. P(f) \<and> P(g) \<Longrightarrow> P(\<lambda>x. f x + g x)"  "\<And>f g. P(f) \<and> P(g) \<Longrightarrow> P(\<lambda>x. f x * g x)"
       
   765           "\<And>x y. x \<in> s \<and> y \<in> s \<and> ~(x = y) \<Longrightarrow> \<exists>f. P(f) \<and> ~(f x = f y)"
       
   766           "continuous_on s f"
       
   767        "0 < e"
       
   768     shows "\<exists>g. P(g) \<and> (\<forall>x \<in> s. abs(f x - g x) < e)"
       
   769 proof -
       
   770   interpret PR: function_ring_on "Collect P"
       
   771     apply unfold_locales
       
   772     using assms
       
   773     by auto
       
   774   show ?thesis
       
   775     using PR.Stone_Weierstrass_basic [OF `continuous_on s f` `0 < e`]
       
   776     by blast
       
   777 qed
       
   778 
       
   779 
       
   780 subsection {*Polynomial functions*}
       
   781 
       
   782 inductive real_polynomial_function :: "('a::real_normed_vector \<Rightarrow> real) \<Rightarrow> bool" where
       
   783     linear: "bounded_linear f \<Longrightarrow> real_polynomial_function f"
       
   784   | const: "real_polynomial_function (\<lambda>x. c)"
       
   785   | add:   "\<lbrakk>real_polynomial_function f; real_polynomial_function g\<rbrakk> \<Longrightarrow> real_polynomial_function (\<lambda>x. f x + g x)"
       
   786   | mult:  "\<lbrakk>real_polynomial_function f; real_polynomial_function g\<rbrakk> \<Longrightarrow> real_polynomial_function (\<lambda>x. f x * g x)"
       
   787 
       
   788 declare real_polynomial_function.intros [intro]
       
   789 
       
   790 definition polynomial_function :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> bool"
       
   791   where
       
   792    "polynomial_function p \<equiv> (\<forall>f. bounded_linear f \<longrightarrow> real_polynomial_function (f o p))"
       
   793 
       
   794 lemma real_polynomial_function_eq: "real_polynomial_function p = polynomial_function p"
       
   795 unfolding polynomial_function_def
       
   796 proof
       
   797   assume "real_polynomial_function p"
       
   798   then show " \<forall>f. bounded_linear f \<longrightarrow> real_polynomial_function (f \<circ> p)"
       
   799   proof (induction p rule: real_polynomial_function.induct)
       
   800     case (linear h) then show ?case
       
   801       by (auto simp: bounded_linear_compose real_polynomial_function.linear)
       
   802   next
       
   803     case (const h) then show ?case
       
   804       by (simp add: real_polynomial_function.const)
       
   805   next
       
   806     case (add h) then show ?case
       
   807       by (force simp add: bounded_linear_def linear_add real_polynomial_function.add)
       
   808   next
       
   809     case (mult h) then show ?case
       
   810       by (force simp add: real_bounded_linear const real_polynomial_function.mult)
       
   811   qed
       
   812 next
       
   813   assume [rule_format, OF bounded_linear_ident]: "\<forall>f. bounded_linear f \<longrightarrow> real_polynomial_function (f \<circ> p)"
       
   814   then show "real_polynomial_function p"
       
   815     by (simp add: o_def)
       
   816 qed
       
   817 
       
   818 lemma polynomial_function_const [iff]: "polynomial_function (\<lambda>x. c)"
       
   819   by (simp add: polynomial_function_def o_def const)
       
   820 
       
   821 lemma polynomial_function_bounded_linear:
       
   822   "bounded_linear f \<Longrightarrow> polynomial_function f"
       
   823   by (simp add: polynomial_function_def o_def bounded_linear_compose real_polynomial_function.linear)
       
   824 
       
   825 lemma polynomial_function_id [iff]: "polynomial_function(\<lambda>x. x)"
       
   826   by (simp add: polynomial_function_bounded_linear)
       
   827 
       
   828 lemma polynomial_function_add [intro]:
       
   829     "\<lbrakk>polynomial_function f; polynomial_function g\<rbrakk> \<Longrightarrow> polynomial_function (\<lambda>x. f x + g x)"
       
   830   by (auto simp: polynomial_function_def bounded_linear_def linear_add real_polynomial_function.add o_def)
       
   831 
       
   832 lemma polynomial_function_mult [intro]:
       
   833   assumes f: "polynomial_function f" and g: "polynomial_function g"
       
   834     shows "polynomial_function (\<lambda>x. f x *\<^sub>R g x)"
       
   835   using g
       
   836   apply (auto simp: polynomial_function_def bounded_linear_def Real_Vector_Spaces.linear.scaleR  const real_polynomial_function.mult o_def)
       
   837   apply (rule mult)
       
   838   using f
       
   839   apply (auto simp: real_polynomial_function_eq)
       
   840   done
       
   841 
       
   842 lemma polynomial_function_cmul [intro]:
       
   843   assumes f: "polynomial_function f"
       
   844     shows "polynomial_function (\<lambda>x. c *\<^sub>R f x)"
       
   845   by (rule polynomial_function_mult [OF polynomial_function_const f])
       
   846 
       
   847 lemma polynomial_function_minus [intro]:
       
   848   assumes f: "polynomial_function f"
       
   849     shows "polynomial_function (\<lambda>x. - f x)"
       
   850   using polynomial_function_cmul [OF f, of "-1"] by simp
       
   851 
       
   852 lemma polynomial_function_diff [intro]:
       
   853     "\<lbrakk>polynomial_function f; polynomial_function g\<rbrakk> \<Longrightarrow> polynomial_function (\<lambda>x. f x - g x)"
       
   854   unfolding add_uminus_conv_diff [symmetric]
       
   855   by (metis polynomial_function_add polynomial_function_minus)
       
   856 
       
   857 lemma polynomial_function_setsum [intro]:
       
   858     "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> polynomial_function (\<lambda>x. f x i)\<rbrakk> \<Longrightarrow> polynomial_function (\<lambda>x. setsum (f x) I)"
       
   859 by (induct I rule: finite_induct) auto
       
   860 
       
   861 lemma real_polynomial_function_minus [intro]:
       
   862     "real_polynomial_function f \<Longrightarrow> real_polynomial_function (\<lambda>x. - f x)"
       
   863   using polynomial_function_minus [of f]
       
   864   by (simp add: real_polynomial_function_eq)
       
   865 
       
   866 lemma real_polynomial_function_diff [intro]:
       
   867     "\<lbrakk>real_polynomial_function f; real_polynomial_function g\<rbrakk> \<Longrightarrow> real_polynomial_function (\<lambda>x. f x - g x)"
       
   868   using polynomial_function_diff [of f]
       
   869   by (simp add: real_polynomial_function_eq)
       
   870 
       
   871 lemma real_polynomial_function_setsum [intro]:
       
   872     "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> real_polynomial_function (\<lambda>x. f x i)\<rbrakk> \<Longrightarrow> real_polynomial_function (\<lambda>x. setsum (f x) I)"
       
   873   using polynomial_function_setsum [of I f]
       
   874   by (simp add: real_polynomial_function_eq)
       
   875 
       
   876 lemma real_polynomial_function_power [intro]:
       
   877     "real_polynomial_function f \<Longrightarrow> real_polynomial_function (\<lambda>x. f x ^ n)"
       
   878   by (induct n) (simp_all add: const mult)
       
   879 
       
   880 lemma real_polynomial_function_compose [intro]:
       
   881   assumes f: "polynomial_function f" and g: "real_polynomial_function g"
       
   882     shows "real_polynomial_function (g o f)"
       
   883   using g
       
   884   apply (induction g rule: real_polynomial_function.induct)
       
   885   using f
       
   886   apply (simp_all add: polynomial_function_def o_def const add mult)
       
   887   done
       
   888 
       
   889 lemma polynomial_function_compose [intro]:
       
   890   assumes f: "polynomial_function f" and g: "polynomial_function g"
       
   891     shows "polynomial_function (g o f)"
       
   892   using g real_polynomial_function_compose [OF f]
       
   893   by (auto simp: polynomial_function_def o_def)
       
   894 
       
   895 lemma setsum_max_0:
       
   896   fixes x::real (*in fact "'a::comm_ring_1"*)
       
   897   shows "(\<Sum>i = 0..max m n. x^i * (if i \<le> m then a i else 0)) = (\<Sum>i = 0..m. x^i * a i)"
       
   898 proof -
       
   899   have "(\<Sum>i = 0..max m n. x^i * (if i \<le> m then a i else 0)) = (\<Sum>i = 0..max m n. (if i \<le> m then x^i * a i else 0))"
       
   900     by (auto simp: algebra_simps intro: setsum.cong)
       
   901   also have "... = (\<Sum>i = 0..m. (if i \<le> m then x^i * a i else 0))"
       
   902     by (rule setsum.mono_neutral_right) auto
       
   903   also have "... = (\<Sum>i = 0..m. x^i * a i)"
       
   904     by (auto simp: algebra_simps intro: setsum.cong)
       
   905   finally show ?thesis .
       
   906 qed
       
   907 
       
   908 lemma real_polynomial_function_imp_setsum:
       
   909   assumes "real_polynomial_function f"
       
   910     shows "\<exists>a n::nat. f = (\<lambda>x. \<Sum>i=0..n. a i * x ^ i)"
       
   911 using assms
       
   912 proof (induct f)
       
   913   case (linear f)
       
   914   then show ?case
       
   915     apply (clarsimp simp add: real_bounded_linear)
       
   916     apply (rule_tac x="\<lambda>i. if i=0 then 0 else c" in exI)
       
   917     apply (rule_tac x=1 in exI)
       
   918     apply (simp add: mult_ac)
       
   919     done
       
   920 next
       
   921   case (const c)
       
   922   show ?case
       
   923     apply (rule_tac x="\<lambda>i. c" in exI)
       
   924     apply (rule_tac x=0 in exI)
       
   925     apply (auto simp: mult_ac real_of_nat_Suc)
       
   926     done
       
   927   case (add f1 f2)
       
   928   then obtain a1 n1 a2 n2 where
       
   929     "f1 = (\<lambda>x. \<Sum>i = 0..n1. a1 i * x ^ i)" "f2 = (\<lambda>x. \<Sum>i = 0..n2. a2 i * x ^ i)"
       
   930     by auto
       
   931   then show ?case
       
   932     apply (rule_tac x="\<lambda>i. (if i \<le> n1 then a1 i else 0) + (if i \<le> n2 then a2 i else 0)" in exI)
       
   933     apply (rule_tac x="max n1 n2" in exI)
       
   934     using setsum_max_0 [where m=n1 and n=n2] setsum_max_0 [where m=n2 and n=n1]
       
   935     apply (simp add: setsum.distrib algebra_simps max.commute)
       
   936     done
       
   937   case (mult f1 f2)
       
   938   then obtain a1 n1 a2 n2 where
       
   939     "f1 = (\<lambda>x. \<Sum>i = 0..n1. a1 i * x ^ i)" "f2 = (\<lambda>x. \<Sum>i = 0..n2. a2 i * x ^ i)"
       
   940     by auto
       
   941   then obtain b1 b2 where
       
   942     "f1 = (\<lambda>x. \<Sum>i = 0..n1. b1 i * x ^ i)" "f2 = (\<lambda>x. \<Sum>i = 0..n2. b2 i * x ^ i)"
       
   943     "b1 = (\<lambda>i. if i\<le>n1 then a1 i else 0)" "b2 = (\<lambda>i. if i\<le>n2 then a2 i else 0)"
       
   944     by auto
       
   945   then show ?case
       
   946     apply (rule_tac x="\<lambda>i. \<Sum>k\<le>i. b1 k * b2 (i - k)" in exI)
       
   947     apply (rule_tac x="n1+n2" in exI)
       
   948     using polynomial_product [of n1 b1 n2 b2]
       
   949     apply (simp add: Set_Interval.atLeast0AtMost)
       
   950     done
       
   951 qed
       
   952 
       
   953 lemma real_polynomial_function_iff_setsum:
       
   954      "real_polynomial_function f \<longleftrightarrow> (\<exists>a n::nat. f = (\<lambda>x. \<Sum>i=0..n. a i * x ^ i))"
       
   955   apply (rule iffI)
       
   956   apply (erule real_polynomial_function_imp_setsum)
       
   957   apply (auto simp: linear mult const real_polynomial_function_power real_polynomial_function_setsum)
       
   958   done
       
   959 
       
   960 lemma polynomial_function_iff_Basis_inner:
       
   961   fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space"
       
   962   shows "polynomial_function f \<longleftrightarrow> (\<forall>b\<in>Basis. real_polynomial_function (\<lambda>x. inner (f x) b))"
       
   963         (is "?lhs = ?rhs")
       
   964 unfolding polynomial_function_def
       
   965 proof (intro iffI allI impI)
       
   966   assume "\<forall>h. bounded_linear h \<longrightarrow> real_polynomial_function (h \<circ> f)"
       
   967   then show ?rhs
       
   968     by (force simp add: bounded_linear_inner_left o_def)
       
   969 next
       
   970   fix h :: "'b \<Rightarrow> real"
       
   971   assume rp: "\<forall>b\<in>Basis. real_polynomial_function (\<lambda>x. f x \<bullet> b)" and h: "bounded_linear h"
       
   972   have "real_polynomial_function (h \<circ> (\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b))"
       
   973     apply (rule real_polynomial_function_compose [OF _  linear [OF h]])
       
   974     using rp
       
   975     apply (auto simp: real_polynomial_function_eq polynomial_function_mult)
       
   976     done
       
   977   then show "real_polynomial_function (h \<circ> f)"
       
   978     by (simp add: euclidean_representation_setsum_fun)
       
   979 qed
       
   980 
       
   981 subsection {*Stone-Weierstrass theorem for polynomial functions*}
       
   982 
       
   983 text\<open>First, we need to show that they are continous, differentiable and separable.\<close>
       
   984 
       
   985 lemma continuous_real_polymonial_function:
       
   986   assumes "real_polynomial_function f"
       
   987     shows "continuous (at x) f"
       
   988 using assms
       
   989 by (induct f) (auto simp: linear_continuous_at)
       
   990 
       
   991 lemma continuous_polymonial_function:
       
   992   fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space"
       
   993   assumes "polynomial_function f"
       
   994     shows "continuous (at x) f"
       
   995   apply (rule euclidean_isCont)
       
   996   using assms apply (simp add: polynomial_function_iff_Basis_inner)
       
   997   apply (force dest: continuous_real_polymonial_function intro: isCont_scaleR)
       
   998   done
       
   999 
       
  1000 lemma continuous_on_polymonial_function:
       
  1001   fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space"
       
  1002   assumes "polynomial_function f"
       
  1003     shows "continuous_on s f"
       
  1004   using continuous_polymonial_function [OF assms] continuous_at_imp_continuous_on
       
  1005   by blast
       
  1006 
       
  1007 lemma has_real_derivative_polynomial_function:
       
  1008   assumes "real_polynomial_function p"
       
  1009     shows "\<exists>p'. real_polynomial_function p' \<and>
       
  1010                  (\<forall>x. (p has_real_derivative (p' x)) (at x))"
       
  1011 using assms
       
  1012 proof (induct p)
       
  1013   case (linear p)
       
  1014   then show ?case
       
  1015     by (force simp: real_bounded_linear const intro!: derivative_eq_intros)
       
  1016 next
       
  1017   case (const c)
       
  1018   show ?case
       
  1019     by (rule_tac x="\<lambda>x. 0" in exI) auto
       
  1020   case (add f1 f2)
       
  1021   then obtain p1 p2 where
       
  1022     "real_polynomial_function p1" "\<And>x. (f1 has_real_derivative p1 x) (at x)"
       
  1023     "real_polynomial_function p2" "\<And>x. (f2 has_real_derivative p2 x) (at x)"
       
  1024     by auto
       
  1025   then show ?case
       
  1026     apply (rule_tac x="\<lambda>x. p1 x + p2 x" in exI)
       
  1027     apply (auto intro!: derivative_eq_intros)
       
  1028     done
       
  1029   case (mult f1 f2)
       
  1030   then obtain p1 p2 where
       
  1031     "real_polynomial_function p1" "\<And>x. (f1 has_real_derivative p1 x) (at x)"
       
  1032     "real_polynomial_function p2" "\<And>x. (f2 has_real_derivative p2 x) (at x)"
       
  1033     by auto
       
  1034   then show ?case
       
  1035     using mult
       
  1036     apply (rule_tac x="\<lambda>x. f1 x * p2 x + f2 x * p1 x" in exI)
       
  1037     apply (auto intro!: derivative_eq_intros)
       
  1038     done
       
  1039 qed
       
  1040 
       
  1041 lemma has_vector_derivative_polynomial_function:
       
  1042   fixes p :: "real \<Rightarrow> 'a::euclidean_space"
       
  1043   assumes "polynomial_function p"
       
  1044     shows "\<exists>p'. polynomial_function p' \<and>
       
  1045                  (\<forall>x. (p has_vector_derivative (p' x)) (at x))"
       
  1046 proof -
       
  1047   { fix b :: 'a
       
  1048     assume "b \<in> Basis"
       
  1049     then
       
  1050     obtain p' where p': "real_polynomial_function p'" and pd: "\<And>x. ((\<lambda>x. p x \<bullet> b) has_real_derivative p' x) (at x)"
       
  1051       using assms [unfolded polynomial_function_iff_Basis_inner, rule_format]  `b \<in> Basis`
       
  1052       has_real_derivative_polynomial_function
       
  1053       by blast
       
  1054     have "\<exists>q. polynomial_function q \<and> (\<forall>x. ((\<lambda>u. (p u \<bullet> b) *\<^sub>R b) has_vector_derivative q x) (at x))"
       
  1055       apply (rule_tac x="\<lambda>x. p' x *\<^sub>R b" in exI)
       
  1056       using `b \<in> Basis` p'
       
  1057       apply (simp add: polynomial_function_iff_Basis_inner inner_Basis)
       
  1058       apply (auto intro: derivative_eq_intros pd)
       
  1059       done
       
  1060   }
       
  1061   then obtain qf where qf:
       
  1062       "\<And>b. b \<in> Basis \<Longrightarrow> polynomial_function (qf b)"
       
  1063       "\<And>b x. b \<in> Basis \<Longrightarrow> ((\<lambda>u. (p u \<bullet> b) *\<^sub>R b) has_vector_derivative qf b x) (at x)"
       
  1064     by metis
       
  1065   show ?thesis
       
  1066     apply (subst euclidean_representation_setsum_fun [of p, symmetric])
       
  1067     apply (rule_tac x="\<lambda>x. \<Sum>b\<in>Basis. qf b x" in exI)
       
  1068     apply (auto intro: has_vector_derivative_setsum qf)
       
  1069     done
       
  1070 qed
       
  1071 
       
  1072 lemma real_polynomial_function_separable:
       
  1073   fixes x :: "'a::euclidean_space"
       
  1074   assumes "x \<noteq> y" shows "\<exists>f. real_polynomial_function f \<and> f x \<noteq> f y"
       
  1075 proof -
       
  1076   have "real_polynomial_function (\<lambda>u. \<Sum>b\<in>Basis. (inner (x-u) b)^2)"
       
  1077     apply (rule real_polynomial_function_setsum)
       
  1078     apply (auto simp: algebra_simps real_polynomial_function_power real_polynomial_function_diff
       
  1079                  const linear bounded_linear_inner_left)
       
  1080     done
       
  1081   then show ?thesis
       
  1082     apply (intro exI conjI, assumption)
       
  1083     using assms
       
  1084     apply (force simp add: euclidean_eq_iff [of x y] setsum_nonneg_eq_0_iff algebra_simps)
       
  1085     done
       
  1086 qed
       
  1087 
       
  1088 lemma Stone_Weierstrass_real_polynomial_function:
       
  1089   fixes f :: "'a::euclidean_space \<Rightarrow> real"
       
  1090   assumes "compact s" "continuous_on s f" "0 < e"
       
  1091     shows "\<exists>g. real_polynomial_function g \<and> (\<forall>x \<in> s. abs(f x - g x) < e)"
       
  1092 proof -
       
  1093   interpret PR: function_ring_on "Collect real_polynomial_function"
       
  1094     apply unfold_locales
       
  1095     using assms continuous_on_polymonial_function real_polynomial_function_eq
       
  1096     apply (auto intro: real_polynomial_function_separable)
       
  1097     done
       
  1098   show ?thesis
       
  1099     using PR.Stone_Weierstrass_basic [OF `continuous_on s f` `0 < e`]
       
  1100     by blast
       
  1101 qed
       
  1102 
       
  1103 lemma Stone_Weierstrass_polynomial_function:
       
  1104   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
  1105   assumes s: "compact s"
       
  1106       and f: "continuous_on s f"
       
  1107       and e: "0 < e"
       
  1108     shows "\<exists>g. polynomial_function g \<and> (\<forall>x \<in> s. norm(f x - g x) < e)"
       
  1109 proof -
       
  1110   { fix b :: 'b
       
  1111     assume "b \<in> Basis"
       
  1112     have "\<exists>p. real_polynomial_function p \<and> (\<forall>x \<in> s. abs(f x \<bullet> b - p x) < e / DIM('b))"
       
  1113       apply (rule exE [OF Stone_Weierstrass_real_polynomial_function [OF s _, of "\<lambda>x. f x \<bullet> b" "e / card Basis"]])
       
  1114       using e f
       
  1115       apply (auto simp: Euclidean_Space.DIM_positive intro: continuous_intros)
       
  1116       done
       
  1117   }
       
  1118   then obtain pf where pf:
       
  1119       "\<And>b. b \<in> Basis \<Longrightarrow> real_polynomial_function (pf b) \<and> (\<forall>x \<in> s. abs(f x \<bullet> b - pf b x) < e / DIM('b))"
       
  1120       apply (rule bchoice [rule_format, THEN exE])
       
  1121       apply assumption
       
  1122       apply (force simp add: intro: that)
       
  1123       done
       
  1124   have "polynomial_function (\<lambda>x. \<Sum>b\<in>Basis. pf b x *\<^sub>R b)"
       
  1125     using pf
       
  1126     by (simp add: polynomial_function_setsum polynomial_function_mult real_polynomial_function_eq)
       
  1127   moreover
       
  1128   { fix x
       
  1129     assume "x \<in> s"
       
  1130     have "norm (\<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b - pf b x *\<^sub>R b) \<le> (\<Sum>b\<in>Basis. norm ((f x \<bullet> b) *\<^sub>R b - pf b x *\<^sub>R b))"
       
  1131       by (rule norm_setsum)
       
  1132     also have "... < of_nat DIM('b) * (e / DIM('b))"
       
  1133       apply (rule setsum_bounded_above_strict)
       
  1134       apply (simp add: Real_Vector_Spaces.scaleR_diff_left [symmetric] pf `x \<in> s`)
       
  1135       apply (rule DIM_positive)
       
  1136       done
       
  1137     also have "... = e"
       
  1138       using DIM_positive by (simp add: field_simps)
       
  1139     finally have "norm (\<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b - pf b x *\<^sub>R b) < e" .
       
  1140   }
       
  1141   ultimately
       
  1142   show ?thesis
       
  1143     apply (subst euclidean_representation_setsum_fun [of f, symmetric])
       
  1144     apply (rule_tac x="\<lambda>x. \<Sum>b\<in>Basis. pf b x *\<^sub>R b" in exI)
       
  1145     apply (auto simp: setsum_subtractf [symmetric])
       
  1146     done
       
  1147 qed
       
  1148 
       
  1149 
       
  1150 subsection\<open>Polynomial functions as paths\<close>
       
  1151 
       
  1152 text{*One application is to pick a smooth approximation to a path,
       
  1153 or just pick a smooth path anyway in an open connected set*}
       
  1154 
       
  1155 lemma path_polynomial_function:
       
  1156     fixes g  :: "real \<Rightarrow> 'b::euclidean_space"
       
  1157     shows "polynomial_function g \<Longrightarrow> path g"
       
  1158   by (simp add: path_def continuous_on_polymonial_function)
       
  1159 
       
  1160 lemma path_approx_polynomial_function:
       
  1161     fixes g :: "real \<Rightarrow> 'b::euclidean_space"
       
  1162     assumes "path g" "0 < e"
       
  1163     shows "\<exists>p. polynomial_function p \<and>
       
  1164                 pathstart p = pathstart g \<and>
       
  1165                 pathfinish p = pathfinish g \<and>
       
  1166                 (\<forall>t \<in> {0..1}. norm(p t - g t) < e)"
       
  1167 proof -
       
  1168   obtain q where poq: "polynomial_function q" and noq: "\<And>x. x \<in> {0..1} \<Longrightarrow> norm (g x - q x) < e/4"
       
  1169     using Stone_Weierstrass_polynomial_function [of "{0..1}" g "e/4"] assms
       
  1170     by (auto simp: path_def)
       
  1171   have pf: "polynomial_function (\<lambda>t. q t + (g 0 - q 0) + t *\<^sub>R (g 1 - q 1 - (g 0 - q 0)))"
       
  1172     by (force simp add: poq)
       
  1173   have *: "\<And>t. t \<in> {0..1} \<Longrightarrow> norm (((q t - g t) + (g 0 - q 0)) + (t *\<^sub>R (g 1 - q 1) + t *\<^sub>R (q 0 - g 0))) < (e/4 + e/4) + (e/4+e/4)"
       
  1174     apply (intro Real_Vector_Spaces.norm_add_less)
       
  1175     using noq
       
  1176     apply (auto simp: norm_minus_commute intro: le_less_trans [OF mult_left_le_one_le noq] simp del: less_divide_eq_numeral1)
       
  1177     done
       
  1178   show ?thesis
       
  1179     apply (intro exI conjI)
       
  1180     apply (rule pf)
       
  1181     using *
       
  1182     apply (auto simp add: pathstart_def pathfinish_def algebra_simps)
       
  1183     done
       
  1184 qed
       
  1185 
       
  1186 lemma connected_open_polynomial_connected:
       
  1187   fixes s :: "'a::euclidean_space set"
       
  1188   assumes s: "open s" "connected s"
       
  1189       and "x \<in> s" "y \<in> s"
       
  1190     shows "\<exists>g. polynomial_function g \<and> path_image g \<subseteq> s \<and>
       
  1191                pathstart g = x \<and> pathfinish g = y"
       
  1192 proof -
       
  1193   have "path_connected s" using assms
       
  1194     by (simp add: connected_open_path_connected)
       
  1195   with `x \<in> s` `y \<in> s` obtain p where p: "path p" "path_image p \<subseteq> s" "pathstart p = x" "pathfinish p = y"
       
  1196     by (force simp: path_connected_def)
       
  1197   have "\<exists>e. 0 < e \<and> (\<forall>x \<in> path_image p. ball x e \<subseteq> s)"
       
  1198   proof (cases "s = UNIV")
       
  1199     case True then show ?thesis
       
  1200       by (simp add: gt_ex)
       
  1201   next
       
  1202     case False
       
  1203     then have "- s \<noteq> {}" by blast
       
  1204     then show ?thesis
       
  1205       apply (rule_tac x="setdist (path_image p) (-s)" in exI)
       
  1206       using s p
       
  1207       apply (simp add: setdist_gt_0_compact_closed compact_path_image open_closed)
       
  1208       using setdist_le_dist [of _ "path_image p" _ "-s"]
       
  1209       by fastforce
       
  1210   qed
       
  1211   then obtain e where "0 < e"and eb: "\<And>x. x \<in> path_image p \<Longrightarrow> ball x e \<subseteq> s"
       
  1212     by auto
       
  1213   show ?thesis
       
  1214     using path_approx_polynomial_function [OF `path p` `0 < e`]
       
  1215     apply clarify
       
  1216     apply (intro exI conjI, assumption)
       
  1217     using p
       
  1218     apply (fastforce simp add: dist_norm path_image_def norm_minus_commute intro: eb [THEN subsetD])+
       
  1219     done
       
  1220 qed
       
  1221 
       
  1222 hide_fact linear add mult const
       
  1223 
       
  1224 end