src/HOL/Real_Vector_Spaces.thy
author hoelzl
Tue, 05 Nov 2013 09:45:02 +0100
changeset 55715 c4159fe6fa46
parent 55682 b1d955791529
child 55862 03ff4d1e6784
permissions -rw-r--r--
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
     1 (*  Title:      HOL/Real_Vector_Spaces.thy
     2     Author:     Brian Huffman
     3     Author:     Johannes Hölzl
     4 *)
     5 
     6 header {* Vector Spaces and Algebras over the Reals *}
     7 
     8 theory Real_Vector_Spaces
     9 imports Real Topological_Spaces
    10 begin
    11 
    12 subsection {* Locale for additive functions *}
    13 
    14 locale additive =
    15   fixes f :: "'a::ab_group_add \<Rightarrow> 'b::ab_group_add"
    16   assumes add: "f (x + y) = f x + f y"
    17 begin
    18 
    19 lemma zero: "f 0 = 0"
    20 proof -
    21   have "f 0 = f (0 + 0)" by simp
    22   also have "\<dots> = f 0 + f 0" by (rule add)
    23   finally show "f 0 = 0" by simp
    24 qed
    25 
    26 lemma minus: "f (- x) = - f x"
    27 proof -
    28   have "f (- x) + f x = f (- x + x)" by (rule add [symmetric])
    29   also have "\<dots> = - f x + f x" by (simp add: zero)
    30   finally show "f (- x) = - f x" by (rule add_right_imp_eq)
    31 qed
    32 
    33 lemma diff: "f (x - y) = f x - f y"
    34   using add [of x "- y"] by (simp add: minus)
    35 
    36 lemma setsum: "f (setsum g A) = (\<Sum>x\<in>A. f (g x))"
    37 apply (cases "finite A")
    38 apply (induct set: finite)
    39 apply (simp add: zero)
    40 apply (simp add: add)
    41 apply (simp add: zero)
    42 done
    43 
    44 end
    45 
    46 subsection {* Vector spaces *}
    47 
    48 locale vector_space =
    49   fixes scale :: "'a::field \<Rightarrow> 'b::ab_group_add \<Rightarrow> 'b"
    50   assumes scale_right_distrib [algebra_simps]:
    51     "scale a (x + y) = scale a x + scale a y"
    52   and scale_left_distrib [algebra_simps]:
    53     "scale (a + b) x = scale a x + scale b x"
    54   and scale_scale [simp]: "scale a (scale b x) = scale (a * b) x"
    55   and scale_one [simp]: "scale 1 x = x"
    56 begin
    57 
    58 lemma scale_left_commute:
    59   "scale a (scale b x) = scale b (scale a x)"
    60 by (simp add: mult_commute)
    61 
    62 lemma scale_zero_left [simp]: "scale 0 x = 0"
    63   and scale_minus_left [simp]: "scale (- a) x = - (scale a x)"
    64   and scale_left_diff_distrib [algebra_simps]:
    65         "scale (a - b) x = scale a x - scale b x"
    66   and scale_setsum_left: "scale (setsum f A) x = (\<Sum>a\<in>A. scale (f a) x)"
    67 proof -
    68   interpret s: additive "\<lambda>a. scale a x"
    69     proof qed (rule scale_left_distrib)
    70   show "scale 0 x = 0" by (rule s.zero)
    71   show "scale (- a) x = - (scale a x)" by (rule s.minus)
    72   show "scale (a - b) x = scale a x - scale b x" by (rule s.diff)
    73   show "scale (setsum f A) x = (\<Sum>a\<in>A. scale (f a) x)" by (rule s.setsum)
    74 qed
    75 
    76 lemma scale_zero_right [simp]: "scale a 0 = 0"
    77   and scale_minus_right [simp]: "scale a (- x) = - (scale a x)"
    78   and scale_right_diff_distrib [algebra_simps]:
    79         "scale a (x - y) = scale a x - scale a y"
    80   and scale_setsum_right: "scale a (setsum f A) = (\<Sum>x\<in>A. scale a (f x))"
    81 proof -
    82   interpret s: additive "\<lambda>x. scale a x"
    83     proof qed (rule scale_right_distrib)
    84   show "scale a 0 = 0" by (rule s.zero)
    85   show "scale a (- x) = - (scale a x)" by (rule s.minus)
    86   show "scale a (x - y) = scale a x - scale a y" by (rule s.diff)
    87   show "scale a (setsum f A) = (\<Sum>x\<in>A. scale a (f x))" by (rule s.setsum)
    88 qed
    89 
    90 lemma scale_eq_0_iff [simp]:
    91   "scale a x = 0 \<longleftrightarrow> a = 0 \<or> x = 0"
    92 proof cases
    93   assume "a = 0" thus ?thesis by simp
    94 next
    95   assume anz [simp]: "a \<noteq> 0"
    96   { assume "scale a x = 0"
    97     hence "scale (inverse a) (scale a x) = 0" by simp
    98     hence "x = 0" by simp }
    99   thus ?thesis by force
   100 qed
   101 
   102 lemma scale_left_imp_eq:
   103   "\<lbrakk>a \<noteq> 0; scale a x = scale a y\<rbrakk> \<Longrightarrow> x = y"
   104 proof -
   105   assume nonzero: "a \<noteq> 0"
   106   assume "scale a x = scale a y"
   107   hence "scale a (x - y) = 0"
   108      by (simp add: scale_right_diff_distrib)
   109   hence "x - y = 0" by (simp add: nonzero)
   110   thus "x = y" by (simp only: right_minus_eq)
   111 qed
   112 
   113 lemma scale_right_imp_eq:
   114   "\<lbrakk>x \<noteq> 0; scale a x = scale b x\<rbrakk> \<Longrightarrow> a = b"
   115 proof -
   116   assume nonzero: "x \<noteq> 0"
   117   assume "scale a x = scale b x"
   118   hence "scale (a - b) x = 0"
   119      by (simp add: scale_left_diff_distrib)
   120   hence "a - b = 0" by (simp add: nonzero)
   121   thus "a = b" by (simp only: right_minus_eq)
   122 qed
   123 
   124 lemma scale_cancel_left [simp]:
   125   "scale a x = scale a y \<longleftrightarrow> x = y \<or> a = 0"
   126 by (auto intro: scale_left_imp_eq)
   127 
   128 lemma scale_cancel_right [simp]:
   129   "scale a x = scale b x \<longleftrightarrow> a = b \<or> x = 0"
   130 by (auto intro: scale_right_imp_eq)
   131 
   132 end
   133 
   134 subsection {* Real vector spaces *}
   135 
   136 class scaleR =
   137   fixes scaleR :: "real \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "*\<^sub>R" 75)
   138 begin
   139 
   140 abbreviation
   141   divideR :: "'a \<Rightarrow> real \<Rightarrow> 'a" (infixl "'/\<^sub>R" 70)
   142 where
   143   "x /\<^sub>R r == scaleR (inverse r) x"
   144 
   145 end
   146 
   147 class real_vector = scaleR + ab_group_add +
   148   assumes scaleR_add_right: "scaleR a (x + y) = scaleR a x + scaleR a y"
   149   and scaleR_add_left: "scaleR (a + b) x = scaleR a x + scaleR b x"
   150   and scaleR_scaleR: "scaleR a (scaleR b x) = scaleR (a * b) x"
   151   and scaleR_one: "scaleR 1 x = x"
   152 
   153 interpretation real_vector:
   154   vector_space "scaleR :: real \<Rightarrow> 'a \<Rightarrow> 'a::real_vector"
   155 apply unfold_locales
   156 apply (rule scaleR_add_right)
   157 apply (rule scaleR_add_left)
   158 apply (rule scaleR_scaleR)
   159 apply (rule scaleR_one)
   160 done
   161 
   162 text {* Recover original theorem names *}
   163 
   164 lemmas scaleR_left_commute = real_vector.scale_left_commute
   165 lemmas scaleR_zero_left = real_vector.scale_zero_left
   166 lemmas scaleR_minus_left = real_vector.scale_minus_left
   167 lemmas scaleR_diff_left = real_vector.scale_left_diff_distrib
   168 lemmas scaleR_setsum_left = real_vector.scale_setsum_left
   169 lemmas scaleR_zero_right = real_vector.scale_zero_right
   170 lemmas scaleR_minus_right = real_vector.scale_minus_right
   171 lemmas scaleR_diff_right = real_vector.scale_right_diff_distrib
   172 lemmas scaleR_setsum_right = real_vector.scale_setsum_right
   173 lemmas scaleR_eq_0_iff = real_vector.scale_eq_0_iff
   174 lemmas scaleR_left_imp_eq = real_vector.scale_left_imp_eq
   175 lemmas scaleR_right_imp_eq = real_vector.scale_right_imp_eq
   176 lemmas scaleR_cancel_left = real_vector.scale_cancel_left
   177 lemmas scaleR_cancel_right = real_vector.scale_cancel_right
   178 
   179 text {* Legacy names *}
   180 
   181 lemmas scaleR_left_distrib = scaleR_add_left
   182 lemmas scaleR_right_distrib = scaleR_add_right
   183 lemmas scaleR_left_diff_distrib = scaleR_diff_left
   184 lemmas scaleR_right_diff_distrib = scaleR_diff_right
   185 
   186 lemma scaleR_minus1_left [simp]:
   187   fixes x :: "'a::real_vector"
   188   shows "scaleR (-1) x = - x"
   189   using scaleR_minus_left [of 1 x] by simp
   190 
   191 class real_algebra = real_vector + ring +
   192   assumes mult_scaleR_left [simp]: "scaleR a x * y = scaleR a (x * y)"
   193   and mult_scaleR_right [simp]: "x * scaleR a y = scaleR a (x * y)"
   194 
   195 class real_algebra_1 = real_algebra + ring_1
   196 
   197 class real_div_algebra = real_algebra_1 + division_ring
   198 
   199 class real_field = real_div_algebra + field
   200 
   201 instantiation real :: real_field
   202 begin
   203 
   204 definition
   205   real_scaleR_def [simp]: "scaleR a x = a * x"
   206 
   207 instance proof
   208 qed (simp_all add: algebra_simps)
   209 
   210 end
   211 
   212 interpretation scaleR_left: additive "(\<lambda>a. scaleR a x::'a::real_vector)"
   213 proof qed (rule scaleR_left_distrib)
   214 
   215 interpretation scaleR_right: additive "(\<lambda>x. scaleR a x::'a::real_vector)"
   216 proof qed (rule scaleR_right_distrib)
   217 
   218 lemma nonzero_inverse_scaleR_distrib:
   219   fixes x :: "'a::real_div_algebra" shows
   220   "\<lbrakk>a \<noteq> 0; x \<noteq> 0\<rbrakk> \<Longrightarrow> inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
   221 by (rule inverse_unique, simp)
   222 
   223 lemma inverse_scaleR_distrib:
   224   fixes x :: "'a::{real_div_algebra, division_ring_inverse_zero}"
   225   shows "inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
   226 apply (case_tac "a = 0", simp)
   227 apply (case_tac "x = 0", simp)
   228 apply (erule (1) nonzero_inverse_scaleR_distrib)
   229 done
   230 
   231 
   232 subsection {* Embedding of the Reals into any @{text real_algebra_1}:
   233 @{term of_real} *}
   234 
   235 definition
   236   of_real :: "real \<Rightarrow> 'a::real_algebra_1" where
   237   "of_real r = scaleR r 1"
   238 
   239 lemma scaleR_conv_of_real: "scaleR r x = of_real r * x"
   240 by (simp add: of_real_def)
   241 
   242 lemma of_real_0 [simp]: "of_real 0 = 0"
   243 by (simp add: of_real_def)
   244 
   245 lemma of_real_1 [simp]: "of_real 1 = 1"
   246 by (simp add: of_real_def)
   247 
   248 lemma of_real_add [simp]: "of_real (x + y) = of_real x + of_real y"
   249 by (simp add: of_real_def scaleR_left_distrib)
   250 
   251 lemma of_real_minus [simp]: "of_real (- x) = - of_real x"
   252 by (simp add: of_real_def)
   253 
   254 lemma of_real_diff [simp]: "of_real (x - y) = of_real x - of_real y"
   255 by (simp add: of_real_def scaleR_left_diff_distrib)
   256 
   257 lemma of_real_mult [simp]: "of_real (x * y) = of_real x * of_real y"
   258 by (simp add: of_real_def mult_commute)
   259 
   260 lemma nonzero_of_real_inverse:
   261   "x \<noteq> 0 \<Longrightarrow> of_real (inverse x) =
   262    inverse (of_real x :: 'a::real_div_algebra)"
   263 by (simp add: of_real_def nonzero_inverse_scaleR_distrib)
   264 
   265 lemma of_real_inverse [simp]:
   266   "of_real (inverse x) =
   267    inverse (of_real x :: 'a::{real_div_algebra, division_ring_inverse_zero})"
   268 by (simp add: of_real_def inverse_scaleR_distrib)
   269 
   270 lemma nonzero_of_real_divide:
   271   "y \<noteq> 0 \<Longrightarrow> of_real (x / y) =
   272    (of_real x / of_real y :: 'a::real_field)"
   273 by (simp add: divide_inverse nonzero_of_real_inverse)
   274 
   275 lemma of_real_divide [simp]:
   276   "of_real (x / y) =
   277    (of_real x / of_real y :: 'a::{real_field, field_inverse_zero})"
   278 by (simp add: divide_inverse)
   279 
   280 lemma of_real_power [simp]:
   281   "of_real (x ^ n) = (of_real x :: 'a::{real_algebra_1}) ^ n"
   282 by (induct n) simp_all
   283 
   284 lemma of_real_eq_iff [simp]: "(of_real x = of_real y) = (x = y)"
   285 by (simp add: of_real_def)
   286 
   287 lemma inj_of_real:
   288   "inj of_real"
   289   by (auto intro: injI)
   290 
   291 lemmas of_real_eq_0_iff [simp] = of_real_eq_iff [of _ 0, simplified]
   292 
   293 lemma of_real_eq_id [simp]: "of_real = (id :: real \<Rightarrow> real)"
   294 proof
   295   fix r
   296   show "of_real r = id r"
   297     by (simp add: of_real_def)
   298 qed
   299 
   300 text{*Collapse nested embeddings*}
   301 lemma of_real_of_nat_eq [simp]: "of_real (of_nat n) = of_nat n"
   302 by (induct n) auto
   303 
   304 lemma of_real_of_int_eq [simp]: "of_real (of_int z) = of_int z"
   305 by (cases z rule: int_diff_cases, simp)
   306 
   307 lemma of_real_numeral: "of_real (numeral w) = numeral w"
   308 using of_real_of_int_eq [of "numeral w"] by simp
   309 
   310 lemma of_real_neg_numeral: "of_real (neg_numeral w) = neg_numeral w"
   311 using of_real_of_int_eq [of "neg_numeral w"] by simp
   312 
   313 text{*Every real algebra has characteristic zero*}
   314 
   315 instance real_algebra_1 < ring_char_0
   316 proof
   317   from inj_of_real inj_of_nat have "inj (of_real \<circ> of_nat)" by (rule inj_comp)
   318   then show "inj (of_nat :: nat \<Rightarrow> 'a)" by (simp add: comp_def)
   319 qed
   320 
   321 instance real_field < field_char_0 ..
   322 
   323 
   324 subsection {* The Set of Real Numbers *}
   325 
   326 definition Reals :: "'a::real_algebra_1 set" where
   327   "Reals = range of_real"
   328 
   329 notation (xsymbols)
   330   Reals  ("\<real>")
   331 
   332 lemma Reals_of_real [simp]: "of_real r \<in> Reals"
   333 by (simp add: Reals_def)
   334 
   335 lemma Reals_of_int [simp]: "of_int z \<in> Reals"
   336 by (subst of_real_of_int_eq [symmetric], rule Reals_of_real)
   337 
   338 lemma Reals_of_nat [simp]: "of_nat n \<in> Reals"
   339 by (subst of_real_of_nat_eq [symmetric], rule Reals_of_real)
   340 
   341 lemma Reals_numeral [simp]: "numeral w \<in> Reals"
   342 by (subst of_real_numeral [symmetric], rule Reals_of_real)
   343 
   344 lemma Reals_neg_numeral [simp]: "neg_numeral w \<in> Reals"
   345 by (subst of_real_neg_numeral [symmetric], rule Reals_of_real)
   346 
   347 lemma Reals_0 [simp]: "0 \<in> Reals"
   348 apply (unfold Reals_def)
   349 apply (rule range_eqI)
   350 apply (rule of_real_0 [symmetric])
   351 done
   352 
   353 lemma Reals_1 [simp]: "1 \<in> Reals"
   354 apply (unfold Reals_def)
   355 apply (rule range_eqI)
   356 apply (rule of_real_1 [symmetric])
   357 done
   358 
   359 lemma Reals_add [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a + b \<in> Reals"
   360 apply (auto simp add: Reals_def)
   361 apply (rule range_eqI)
   362 apply (rule of_real_add [symmetric])
   363 done
   364 
   365 lemma Reals_minus [simp]: "a \<in> Reals \<Longrightarrow> - a \<in> Reals"
   366 apply (auto simp add: Reals_def)
   367 apply (rule range_eqI)
   368 apply (rule of_real_minus [symmetric])
   369 done
   370 
   371 lemma Reals_diff [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a - b \<in> Reals"
   372 apply (auto simp add: Reals_def)
   373 apply (rule range_eqI)
   374 apply (rule of_real_diff [symmetric])
   375 done
   376 
   377 lemma Reals_mult [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a * b \<in> Reals"
   378 apply (auto simp add: Reals_def)
   379 apply (rule range_eqI)
   380 apply (rule of_real_mult [symmetric])
   381 done
   382 
   383 lemma nonzero_Reals_inverse:
   384   fixes a :: "'a::real_div_algebra"
   385   shows "\<lbrakk>a \<in> Reals; a \<noteq> 0\<rbrakk> \<Longrightarrow> inverse a \<in> Reals"
   386 apply (auto simp add: Reals_def)
   387 apply (rule range_eqI)
   388 apply (erule nonzero_of_real_inverse [symmetric])
   389 done
   390 
   391 lemma Reals_inverse [simp]:
   392   fixes a :: "'a::{real_div_algebra, division_ring_inverse_zero}"
   393   shows "a \<in> Reals \<Longrightarrow> inverse a \<in> Reals"
   394 apply (auto simp add: Reals_def)
   395 apply (rule range_eqI)
   396 apply (rule of_real_inverse [symmetric])
   397 done
   398 
   399 lemma nonzero_Reals_divide:
   400   fixes a b :: "'a::real_field"
   401   shows "\<lbrakk>a \<in> Reals; b \<in> Reals; b \<noteq> 0\<rbrakk> \<Longrightarrow> a / b \<in> Reals"
   402 apply (auto simp add: Reals_def)
   403 apply (rule range_eqI)
   404 apply (erule nonzero_of_real_divide [symmetric])
   405 done
   406 
   407 lemma Reals_divide [simp]:
   408   fixes a b :: "'a::{real_field, field_inverse_zero}"
   409   shows "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a / b \<in> Reals"
   410 apply (auto simp add: Reals_def)
   411 apply (rule range_eqI)
   412 apply (rule of_real_divide [symmetric])
   413 done
   414 
   415 lemma Reals_power [simp]:
   416   fixes a :: "'a::{real_algebra_1}"
   417   shows "a \<in> Reals \<Longrightarrow> a ^ n \<in> Reals"
   418 apply (auto simp add: Reals_def)
   419 apply (rule range_eqI)
   420 apply (rule of_real_power [symmetric])
   421 done
   422 
   423 lemma Reals_cases [cases set: Reals]:
   424   assumes "q \<in> \<real>"
   425   obtains (of_real) r where "q = of_real r"
   426   unfolding Reals_def
   427 proof -
   428   from `q \<in> \<real>` have "q \<in> range of_real" unfolding Reals_def .
   429   then obtain r where "q = of_real r" ..
   430   then show thesis ..
   431 qed
   432 
   433 lemma Reals_induct [case_names of_real, induct set: Reals]:
   434   "q \<in> \<real> \<Longrightarrow> (\<And>r. P (of_real r)) \<Longrightarrow> P q"
   435   by (rule Reals_cases) auto
   436 
   437 
   438 subsection {* Real normed vector spaces *}
   439 
   440 class dist =
   441   fixes dist :: "'a \<Rightarrow> 'a \<Rightarrow> real"
   442 
   443 class norm =
   444   fixes norm :: "'a \<Rightarrow> real"
   445 
   446 class sgn_div_norm = scaleR + norm + sgn +
   447   assumes sgn_div_norm: "sgn x = x /\<^sub>R norm x"
   448 
   449 class dist_norm = dist + norm + minus +
   450   assumes dist_norm: "dist x y = norm (x - y)"
   451 
   452 class open_dist = "open" + dist +
   453   assumes open_dist: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)"
   454 
   455 class real_normed_vector = real_vector + sgn_div_norm + dist_norm + open_dist +
   456   assumes norm_eq_zero [simp]: "norm x = 0 \<longleftrightarrow> x = 0"
   457   and norm_triangle_ineq: "norm (x + y) \<le> norm x + norm y"
   458   and norm_scaleR [simp]: "norm (scaleR a x) = \<bar>a\<bar> * norm x"
   459 begin
   460 
   461 lemma norm_ge_zero [simp]: "0 \<le> norm x"
   462 proof -
   463   have "0 = norm (x + -1 *\<^sub>R x)" 
   464     using scaleR_add_left[of 1 "-1" x] norm_scaleR[of 0 x] by (simp add: scaleR_one)
   465   also have "\<dots> \<le> norm x + norm (-1 *\<^sub>R x)" by (rule norm_triangle_ineq)
   466   finally show ?thesis by simp
   467 qed
   468 
   469 end
   470 
   471 class real_normed_algebra = real_algebra + real_normed_vector +
   472   assumes norm_mult_ineq: "norm (x * y) \<le> norm x * norm y"
   473 
   474 class real_normed_algebra_1 = real_algebra_1 + real_normed_algebra +
   475   assumes norm_one [simp]: "norm 1 = 1"
   476 
   477 class real_normed_div_algebra = real_div_algebra + real_normed_vector +
   478   assumes norm_mult: "norm (x * y) = norm x * norm y"
   479 
   480 class real_normed_field = real_field + real_normed_div_algebra
   481 
   482 instance real_normed_div_algebra < real_normed_algebra_1
   483 proof
   484   fix x y :: 'a
   485   show "norm (x * y) \<le> norm x * norm y"
   486     by (simp add: norm_mult)
   487 next
   488   have "norm (1 * 1::'a) = norm (1::'a) * norm (1::'a)"
   489     by (rule norm_mult)
   490   thus "norm (1::'a) = 1" by simp
   491 qed
   492 
   493 lemma norm_zero [simp]: "norm (0::'a::real_normed_vector) = 0"
   494 by simp
   495 
   496 lemma zero_less_norm_iff [simp]:
   497   fixes x :: "'a::real_normed_vector"
   498   shows "(0 < norm x) = (x \<noteq> 0)"
   499 by (simp add: order_less_le)
   500 
   501 lemma norm_not_less_zero [simp]:
   502   fixes x :: "'a::real_normed_vector"
   503   shows "\<not> norm x < 0"
   504 by (simp add: linorder_not_less)
   505 
   506 lemma norm_le_zero_iff [simp]:
   507   fixes x :: "'a::real_normed_vector"
   508   shows "(norm x \<le> 0) = (x = 0)"
   509 by (simp add: order_le_less)
   510 
   511 lemma norm_minus_cancel [simp]:
   512   fixes x :: "'a::real_normed_vector"
   513   shows "norm (- x) = norm x"
   514 proof -
   515   have "norm (- x) = norm (scaleR (- 1) x)"
   516     by (simp only: scaleR_minus_left scaleR_one)
   517   also have "\<dots> = \<bar>- 1\<bar> * norm x"
   518     by (rule norm_scaleR)
   519   finally show ?thesis by simp
   520 qed
   521 
   522 lemma norm_minus_commute:
   523   fixes a b :: "'a::real_normed_vector"
   524   shows "norm (a - b) = norm (b - a)"
   525 proof -
   526   have "norm (- (b - a)) = norm (b - a)"
   527     by (rule norm_minus_cancel)
   528   thus ?thesis by simp
   529 qed
   530 
   531 lemma norm_triangle_ineq2:
   532   fixes a b :: "'a::real_normed_vector"
   533   shows "norm a - norm b \<le> norm (a - b)"
   534 proof -
   535   have "norm (a - b + b) \<le> norm (a - b) + norm b"
   536     by (rule norm_triangle_ineq)
   537   thus ?thesis by simp
   538 qed
   539 
   540 lemma norm_triangle_ineq3:
   541   fixes a b :: "'a::real_normed_vector"
   542   shows "\<bar>norm a - norm b\<bar> \<le> norm (a - b)"
   543 apply (subst abs_le_iff)
   544 apply auto
   545 apply (rule norm_triangle_ineq2)
   546 apply (subst norm_minus_commute)
   547 apply (rule norm_triangle_ineq2)
   548 done
   549 
   550 lemma norm_triangle_ineq4:
   551   fixes a b :: "'a::real_normed_vector"
   552   shows "norm (a - b) \<le> norm a + norm b"
   553 proof -
   554   have "norm (a + - b) \<le> norm a + norm (- b)"
   555     by (rule norm_triangle_ineq)
   556   then show ?thesis by simp
   557 qed
   558 
   559 lemma norm_diff_ineq:
   560   fixes a b :: "'a::real_normed_vector"
   561   shows "norm a - norm b \<le> norm (a + b)"
   562 proof -
   563   have "norm a - norm (- b) \<le> norm (a - - b)"
   564     by (rule norm_triangle_ineq2)
   565   thus ?thesis by simp
   566 qed
   567 
   568 lemma norm_diff_triangle_ineq:
   569   fixes a b c d :: "'a::real_normed_vector"
   570   shows "norm ((a + b) - (c + d)) \<le> norm (a - c) + norm (b - d)"
   571 proof -
   572   have "norm ((a + b) - (c + d)) = norm ((a - c) + (b - d))"
   573     by (simp add: algebra_simps)
   574   also have "\<dots> \<le> norm (a - c) + norm (b - d)"
   575     by (rule norm_triangle_ineq)
   576   finally show ?thesis .
   577 qed
   578 
   579 lemma abs_norm_cancel [simp]:
   580   fixes a :: "'a::real_normed_vector"
   581   shows "\<bar>norm a\<bar> = norm a"
   582 by (rule abs_of_nonneg [OF norm_ge_zero])
   583 
   584 lemma norm_add_less:
   585   fixes x y :: "'a::real_normed_vector"
   586   shows "\<lbrakk>norm x < r; norm y < s\<rbrakk> \<Longrightarrow> norm (x + y) < r + s"
   587 by (rule order_le_less_trans [OF norm_triangle_ineq add_strict_mono])
   588 
   589 lemma norm_mult_less:
   590   fixes x y :: "'a::real_normed_algebra"
   591   shows "\<lbrakk>norm x < r; norm y < s\<rbrakk> \<Longrightarrow> norm (x * y) < r * s"
   592 apply (rule order_le_less_trans [OF norm_mult_ineq])
   593 apply (simp add: mult_strict_mono')
   594 done
   595 
   596 lemma norm_of_real [simp]:
   597   "norm (of_real r :: 'a::real_normed_algebra_1) = \<bar>r\<bar>"
   598 unfolding of_real_def by simp
   599 
   600 lemma norm_numeral [simp]:
   601   "norm (numeral w::'a::real_normed_algebra_1) = numeral w"
   602 by (subst of_real_numeral [symmetric], subst norm_of_real, simp)
   603 
   604 lemma norm_neg_numeral [simp]:
   605   "norm (neg_numeral w::'a::real_normed_algebra_1) = numeral w"
   606 by (subst of_real_neg_numeral [symmetric], subst norm_of_real, simp)
   607 
   608 lemma norm_of_int [simp]:
   609   "norm (of_int z::'a::real_normed_algebra_1) = \<bar>of_int z\<bar>"
   610 by (subst of_real_of_int_eq [symmetric], rule norm_of_real)
   611 
   612 lemma norm_of_nat [simp]:
   613   "norm (of_nat n::'a::real_normed_algebra_1) = of_nat n"
   614 apply (subst of_real_of_nat_eq [symmetric])
   615 apply (subst norm_of_real, simp)
   616 done
   617 
   618 lemma nonzero_norm_inverse:
   619   fixes a :: "'a::real_normed_div_algebra"
   620   shows "a \<noteq> 0 \<Longrightarrow> norm (inverse a) = inverse (norm a)"
   621 apply (rule inverse_unique [symmetric])
   622 apply (simp add: norm_mult [symmetric])
   623 done
   624 
   625 lemma norm_inverse:
   626   fixes a :: "'a::{real_normed_div_algebra, division_ring_inverse_zero}"
   627   shows "norm (inverse a) = inverse (norm a)"
   628 apply (case_tac "a = 0", simp)
   629 apply (erule nonzero_norm_inverse)
   630 done
   631 
   632 lemma nonzero_norm_divide:
   633   fixes a b :: "'a::real_normed_field"
   634   shows "b \<noteq> 0 \<Longrightarrow> norm (a / b) = norm a / norm b"
   635 by (simp add: divide_inverse norm_mult nonzero_norm_inverse)
   636 
   637 lemma norm_divide:
   638   fixes a b :: "'a::{real_normed_field, field_inverse_zero}"
   639   shows "norm (a / b) = norm a / norm b"
   640 by (simp add: divide_inverse norm_mult norm_inverse)
   641 
   642 lemma norm_power_ineq:
   643   fixes x :: "'a::{real_normed_algebra_1}"
   644   shows "norm (x ^ n) \<le> norm x ^ n"
   645 proof (induct n)
   646   case 0 show "norm (x ^ 0) \<le> norm x ^ 0" by simp
   647 next
   648   case (Suc n)
   649   have "norm (x * x ^ n) \<le> norm x * norm (x ^ n)"
   650     by (rule norm_mult_ineq)
   651   also from Suc have "\<dots> \<le> norm x * norm x ^ n"
   652     using norm_ge_zero by (rule mult_left_mono)
   653   finally show "norm (x ^ Suc n) \<le> norm x ^ Suc n"
   654     by simp
   655 qed
   656 
   657 lemma norm_power:
   658   fixes x :: "'a::{real_normed_div_algebra}"
   659   shows "norm (x ^ n) = norm x ^ n"
   660 by (induct n) (simp_all add: norm_mult)
   661 
   662 
   663 subsection {* Metric spaces *}
   664 
   665 class metric_space = open_dist +
   666   assumes dist_eq_0_iff [simp]: "dist x y = 0 \<longleftrightarrow> x = y"
   667   assumes dist_triangle2: "dist x y \<le> dist x z + dist y z"
   668 begin
   669 
   670 lemma dist_self [simp]: "dist x x = 0"
   671 by simp
   672 
   673 lemma zero_le_dist [simp]: "0 \<le> dist x y"
   674 using dist_triangle2 [of x x y] by simp
   675 
   676 lemma zero_less_dist_iff: "0 < dist x y \<longleftrightarrow> x \<noteq> y"
   677 by (simp add: less_le)
   678 
   679 lemma dist_not_less_zero [simp]: "\<not> dist x y < 0"
   680 by (simp add: not_less)
   681 
   682 lemma dist_le_zero_iff [simp]: "dist x y \<le> 0 \<longleftrightarrow> x = y"
   683 by (simp add: le_less)
   684 
   685 lemma dist_commute: "dist x y = dist y x"
   686 proof (rule order_antisym)
   687   show "dist x y \<le> dist y x"
   688     using dist_triangle2 [of x y x] by simp
   689   show "dist y x \<le> dist x y"
   690     using dist_triangle2 [of y x y] by simp
   691 qed
   692 
   693 lemma dist_triangle: "dist x z \<le> dist x y + dist y z"
   694 using dist_triangle2 [of x z y] by (simp add: dist_commute)
   695 
   696 lemma dist_triangle3: "dist x y \<le> dist a x + dist a y"
   697 using dist_triangle2 [of x y a] by (simp add: dist_commute)
   698 
   699 lemma dist_triangle_alt:
   700   shows "dist y z <= dist x y + dist x z"
   701 by (rule dist_triangle3)
   702 
   703 lemma dist_pos_lt:
   704   shows "x \<noteq> y ==> 0 < dist x y"
   705 by (simp add: zero_less_dist_iff)
   706 
   707 lemma dist_nz:
   708   shows "x \<noteq> y \<longleftrightarrow> 0 < dist x y"
   709 by (simp add: zero_less_dist_iff)
   710 
   711 lemma dist_triangle_le:
   712   shows "dist x z + dist y z <= e \<Longrightarrow> dist x y <= e"
   713 by (rule order_trans [OF dist_triangle2])
   714 
   715 lemma dist_triangle_lt:
   716   shows "dist x z + dist y z < e ==> dist x y < e"
   717 by (rule le_less_trans [OF dist_triangle2])
   718 
   719 lemma dist_triangle_half_l:
   720   shows "dist x1 y < e / 2 \<Longrightarrow> dist x2 y < e / 2 \<Longrightarrow> dist x1 x2 < e"
   721 by (rule dist_triangle_lt [where z=y], simp)
   722 
   723 lemma dist_triangle_half_r:
   724   shows "dist y x1 < e / 2 \<Longrightarrow> dist y x2 < e / 2 \<Longrightarrow> dist x1 x2 < e"
   725 by (rule dist_triangle_half_l, simp_all add: dist_commute)
   726 
   727 subclass topological_space
   728 proof
   729   have "\<exists>e::real. 0 < e"
   730     by (fast intro: zero_less_one)
   731   then show "open UNIV"
   732     unfolding open_dist by simp
   733 next
   734   fix S T assume "open S" "open T"
   735   then show "open (S \<inter> T)"
   736     unfolding open_dist
   737     apply clarify
   738     apply (drule (1) bspec)+
   739     apply (clarify, rename_tac r s)
   740     apply (rule_tac x="min r s" in exI, simp)
   741     done
   742 next
   743   fix K assume "\<forall>S\<in>K. open S" thus "open (\<Union>K)"
   744     unfolding open_dist by fast
   745 qed
   746 
   747 lemma open_ball: "open {y. dist x y < d}"
   748 proof (unfold open_dist, intro ballI)
   749   fix y assume *: "y \<in> {y. dist x y < d}"
   750   then show "\<exists>e>0. \<forall>z. dist z y < e \<longrightarrow> z \<in> {y. dist x y < d}"
   751     by (auto intro!: exI[of _ "d - dist x y"] simp: field_simps dist_triangle_lt)
   752 qed
   753 
   754 subclass first_countable_topology
   755 proof
   756   fix x 
   757   show "\<exists>A::nat \<Rightarrow> 'a set. (\<forall>i. x \<in> A i \<and> open (A i)) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>i. A i \<subseteq> S))"
   758   proof (safe intro!: exI[of _ "\<lambda>n. {y. dist x y < inverse (Suc n)}"])
   759     fix S assume "open S" "x \<in> S"
   760     then obtain e where e: "0 < e" and "{y. dist x y < e} \<subseteq> S"
   761       by (auto simp: open_dist subset_eq dist_commute)
   762     moreover
   763     from e obtain i where "inverse (Suc i) < e"
   764       by (auto dest!: reals_Archimedean)
   765     then have "{y. dist x y < inverse (Suc i)} \<subseteq> {y. dist x y < e}"
   766       by auto
   767     ultimately show "\<exists>i. {y. dist x y < inverse (Suc i)} \<subseteq> S"
   768       by blast
   769   qed (auto intro: open_ball)
   770 qed
   771 
   772 end
   773 
   774 instance metric_space \<subseteq> t2_space
   775 proof
   776   fix x y :: "'a::metric_space"
   777   assume xy: "x \<noteq> y"
   778   let ?U = "{y'. dist x y' < dist x y / 2}"
   779   let ?V = "{x'. dist y x' < dist x y / 2}"
   780   have th0: "\<And>d x y z. (d x z :: real) \<le> d x y + d y z \<Longrightarrow> d y z = d z y
   781                \<Longrightarrow> \<not>(d x y * 2 < d x z \<and> d z y * 2 < d x z)" by arith
   782   have "open ?U \<and> open ?V \<and> x \<in> ?U \<and> y \<in> ?V \<and> ?U \<inter> ?V = {}"
   783     using dist_pos_lt[OF xy] th0[of dist, OF dist_triangle dist_commute]
   784     using open_ball[of _ "dist x y / 2"] by auto
   785   then show "\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
   786     by blast
   787 qed
   788 
   789 text {* Every normed vector space is a metric space. *}
   790 
   791 instance real_normed_vector < metric_space
   792 proof
   793   fix x y :: 'a show "dist x y = 0 \<longleftrightarrow> x = y"
   794     unfolding dist_norm by simp
   795 next
   796   fix x y z :: 'a show "dist x y \<le> dist x z + dist y z"
   797     unfolding dist_norm
   798     using norm_triangle_ineq4 [of "x - z" "y - z"] by simp
   799 qed
   800 
   801 subsection {* Class instances for real numbers *}
   802 
   803 instantiation real :: real_normed_field
   804 begin
   805 
   806 definition dist_real_def:
   807   "dist x y = \<bar>x - y\<bar>"
   808 
   809 definition open_real_def [code del]:
   810   "open (S :: real set) \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)"
   811 
   812 definition real_norm_def [simp]:
   813   "norm r = \<bar>r\<bar>"
   814 
   815 instance
   816 apply (intro_classes, unfold real_norm_def real_scaleR_def)
   817 apply (rule dist_real_def)
   818 apply (rule open_real_def)
   819 apply (simp add: sgn_real_def)
   820 apply (rule abs_eq_0)
   821 apply (rule abs_triangle_ineq)
   822 apply (rule abs_mult)
   823 apply (rule abs_mult)
   824 done
   825 
   826 end
   827 
   828 code_abort "open :: real set \<Rightarrow> bool"
   829 
   830 instance real :: linorder_topology
   831 proof
   832   show "(open :: real set \<Rightarrow> bool) = generate_topology (range lessThan \<union> range greaterThan)"
   833   proof (rule ext, safe)
   834     fix S :: "real set" assume "open S"
   835     then obtain f where "\<forall>x\<in>S. 0 < f x \<and> (\<forall>y. dist y x < f x \<longrightarrow> y \<in> S)"
   836       unfolding open_real_def bchoice_iff ..
   837     then have *: "S = (\<Union>x\<in>S. {x - f x <..} \<inter> {..< x + f x})"
   838       by (fastforce simp: dist_real_def)
   839     show "generate_topology (range lessThan \<union> range greaterThan) S"
   840       apply (subst *)
   841       apply (intro generate_topology_Union generate_topology.Int)
   842       apply (auto intro: generate_topology.Basis)
   843       done
   844   next
   845     fix S :: "real set" assume "generate_topology (range lessThan \<union> range greaterThan) S"
   846     moreover have "\<And>a::real. open {..<a}"
   847       unfolding open_real_def dist_real_def
   848     proof clarify
   849       fix x a :: real assume "x < a"
   850       hence "0 < a - x \<and> (\<forall>y. \<bar>y - x\<bar> < a - x \<longrightarrow> y \<in> {..<a})" by auto
   851       thus "\<exists>e>0. \<forall>y. \<bar>y - x\<bar> < e \<longrightarrow> y \<in> {..<a}" ..
   852     qed
   853     moreover have "\<And>a::real. open {a <..}"
   854       unfolding open_real_def dist_real_def
   855     proof clarify
   856       fix x a :: real assume "a < x"
   857       hence "0 < x - a \<and> (\<forall>y. \<bar>y - x\<bar> < x - a \<longrightarrow> y \<in> {a<..})" by auto
   858       thus "\<exists>e>0. \<forall>y. \<bar>y - x\<bar> < e \<longrightarrow> y \<in> {a<..}" ..
   859     qed
   860     ultimately show "open S"
   861       by induct auto
   862   qed
   863 qed
   864 
   865 instance real :: linear_continuum_topology ..
   866 
   867 lemmas open_real_greaterThan = open_greaterThan[where 'a=real]
   868 lemmas open_real_lessThan = open_lessThan[where 'a=real]
   869 lemmas open_real_greaterThanLessThan = open_greaterThanLessThan[where 'a=real]
   870 lemmas closed_real_atMost = closed_atMost[where 'a=real]
   871 lemmas closed_real_atLeast = closed_atLeast[where 'a=real]
   872 lemmas closed_real_atLeastAtMost = closed_atLeastAtMost[where 'a=real]
   873 
   874 subsection {* Extra type constraints *}
   875 
   876 text {* Only allow @{term "open"} in class @{text topological_space}. *}
   877 
   878 setup {* Sign.add_const_constraint
   879   (@{const_name "open"}, SOME @{typ "'a::topological_space set \<Rightarrow> bool"}) *}
   880 
   881 text {* Only allow @{term dist} in class @{text metric_space}. *}
   882 
   883 setup {* Sign.add_const_constraint
   884   (@{const_name dist}, SOME @{typ "'a::metric_space \<Rightarrow> 'a \<Rightarrow> real"}) *}
   885 
   886 text {* Only allow @{term norm} in class @{text real_normed_vector}. *}
   887 
   888 setup {* Sign.add_const_constraint
   889   (@{const_name norm}, SOME @{typ "'a::real_normed_vector \<Rightarrow> real"}) *}
   890 
   891 subsection {* Sign function *}
   892 
   893 lemma norm_sgn:
   894   "norm (sgn(x::'a::real_normed_vector)) = (if x = 0 then 0 else 1)"
   895 by (simp add: sgn_div_norm)
   896 
   897 lemma sgn_zero [simp]: "sgn(0::'a::real_normed_vector) = 0"
   898 by (simp add: sgn_div_norm)
   899 
   900 lemma sgn_zero_iff: "(sgn(x::'a::real_normed_vector) = 0) = (x = 0)"
   901 by (simp add: sgn_div_norm)
   902 
   903 lemma sgn_minus: "sgn (- x) = - sgn(x::'a::real_normed_vector)"
   904 by (simp add: sgn_div_norm)
   905 
   906 lemma sgn_scaleR:
   907   "sgn (scaleR r x) = scaleR (sgn r) (sgn(x::'a::real_normed_vector))"
   908 by (simp add: sgn_div_norm mult_ac)
   909 
   910 lemma sgn_one [simp]: "sgn (1::'a::real_normed_algebra_1) = 1"
   911 by (simp add: sgn_div_norm)
   912 
   913 lemma sgn_of_real:
   914   "sgn (of_real r::'a::real_normed_algebra_1) = of_real (sgn r)"
   915 unfolding of_real_def by (simp only: sgn_scaleR sgn_one)
   916 
   917 lemma sgn_mult:
   918   fixes x y :: "'a::real_normed_div_algebra"
   919   shows "sgn (x * y) = sgn x * sgn y"
   920 by (simp add: sgn_div_norm norm_mult mult_commute)
   921 
   922 lemma real_sgn_eq: "sgn (x::real) = x / \<bar>x\<bar>"
   923 by (simp add: sgn_div_norm divide_inverse)
   924 
   925 lemma real_sgn_pos: "0 < (x::real) \<Longrightarrow> sgn x = 1"
   926 unfolding real_sgn_eq by simp
   927 
   928 lemma real_sgn_neg: "(x::real) < 0 \<Longrightarrow> sgn x = -1"
   929 unfolding real_sgn_eq by simp
   930 
   931 lemma norm_conv_dist: "norm x = dist x 0"
   932   unfolding dist_norm by simp
   933 
   934 subsection {* Bounded Linear and Bilinear Operators *}
   935 
   936 locale linear = additive f for f :: "'a::real_vector \<Rightarrow> 'b::real_vector" +
   937   assumes scaleR: "f (scaleR r x) = scaleR r (f x)"
   938 
   939 lemma linearI:
   940   assumes "\<And>x y. f (x + y) = f x + f y"
   941   assumes "\<And>c x. f (c *\<^sub>R x) = c *\<^sub>R f x"
   942   shows "linear f"
   943   by default (rule assms)+
   944 
   945 locale bounded_linear = linear f for f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" +
   946   assumes bounded: "\<exists>K. \<forall>x. norm (f x) \<le> norm x * K"
   947 begin
   948 
   949 lemma pos_bounded:
   950   "\<exists>K>0. \<forall>x. norm (f x) \<le> norm x * K"
   951 proof -
   952   obtain K where K: "\<And>x. norm (f x) \<le> norm x * K"
   953     using bounded by fast
   954   show ?thesis
   955   proof (intro exI impI conjI allI)
   956     show "0 < max 1 K"
   957       by (rule order_less_le_trans [OF zero_less_one le_maxI1])
   958   next
   959     fix x
   960     have "norm (f x) \<le> norm x * K" using K .
   961     also have "\<dots> \<le> norm x * max 1 K"
   962       by (rule mult_left_mono [OF le_maxI2 norm_ge_zero])
   963     finally show "norm (f x) \<le> norm x * max 1 K" .
   964   qed
   965 qed
   966 
   967 lemma nonneg_bounded:
   968   "\<exists>K\<ge>0. \<forall>x. norm (f x) \<le> norm x * K"
   969 proof -
   970   from pos_bounded
   971   show ?thesis by (auto intro: order_less_imp_le)
   972 qed
   973 
   974 end
   975 
   976 lemma bounded_linear_intro:
   977   assumes "\<And>x y. f (x + y) = f x + f y"
   978   assumes "\<And>r x. f (scaleR r x) = scaleR r (f x)"
   979   assumes "\<And>x. norm (f x) \<le> norm x * K"
   980   shows "bounded_linear f"
   981   by default (fast intro: assms)+
   982 
   983 locale bounded_bilinear =
   984   fixes prod :: "['a::real_normed_vector, 'b::real_normed_vector]
   985                  \<Rightarrow> 'c::real_normed_vector"
   986     (infixl "**" 70)
   987   assumes add_left: "prod (a + a') b = prod a b + prod a' b"
   988   assumes add_right: "prod a (b + b') = prod a b + prod a b'"
   989   assumes scaleR_left: "prod (scaleR r a) b = scaleR r (prod a b)"
   990   assumes scaleR_right: "prod a (scaleR r b) = scaleR r (prod a b)"
   991   assumes bounded: "\<exists>K. \<forall>a b. norm (prod a b) \<le> norm a * norm b * K"
   992 begin
   993 
   994 lemma pos_bounded:
   995   "\<exists>K>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
   996 apply (cut_tac bounded, erule exE)
   997 apply (rule_tac x="max 1 K" in exI, safe)
   998 apply (rule order_less_le_trans [OF zero_less_one le_maxI1])
   999 apply (drule spec, drule spec, erule order_trans)
  1000 apply (rule mult_left_mono [OF le_maxI2])
  1001 apply (intro mult_nonneg_nonneg norm_ge_zero)
  1002 done
  1003 
  1004 lemma nonneg_bounded:
  1005   "\<exists>K\<ge>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
  1006 proof -
  1007   from pos_bounded
  1008   show ?thesis by (auto intro: order_less_imp_le)
  1009 qed
  1010 
  1011 lemma additive_right: "additive (\<lambda>b. prod a b)"
  1012 by (rule additive.intro, rule add_right)
  1013 
  1014 lemma additive_left: "additive (\<lambda>a. prod a b)"
  1015 by (rule additive.intro, rule add_left)
  1016 
  1017 lemma zero_left: "prod 0 b = 0"
  1018 by (rule additive.zero [OF additive_left])
  1019 
  1020 lemma zero_right: "prod a 0 = 0"
  1021 by (rule additive.zero [OF additive_right])
  1022 
  1023 lemma minus_left: "prod (- a) b = - prod a b"
  1024 by (rule additive.minus [OF additive_left])
  1025 
  1026 lemma minus_right: "prod a (- b) = - prod a b"
  1027 by (rule additive.minus [OF additive_right])
  1028 
  1029 lemma diff_left:
  1030   "prod (a - a') b = prod a b - prod a' b"
  1031 by (rule additive.diff [OF additive_left])
  1032 
  1033 lemma diff_right:
  1034   "prod a (b - b') = prod a b - prod a b'"
  1035 by (rule additive.diff [OF additive_right])
  1036 
  1037 lemma bounded_linear_left:
  1038   "bounded_linear (\<lambda>a. a ** b)"
  1039 apply (cut_tac bounded, safe)
  1040 apply (rule_tac K="norm b * K" in bounded_linear_intro)
  1041 apply (rule add_left)
  1042 apply (rule scaleR_left)
  1043 apply (simp add: mult_ac)
  1044 done
  1045 
  1046 lemma bounded_linear_right:
  1047   "bounded_linear (\<lambda>b. a ** b)"
  1048 apply (cut_tac bounded, safe)
  1049 apply (rule_tac K="norm a * K" in bounded_linear_intro)
  1050 apply (rule add_right)
  1051 apply (rule scaleR_right)
  1052 apply (simp add: mult_ac)
  1053 done
  1054 
  1055 lemma prod_diff_prod:
  1056   "(x ** y - a ** b) = (x - a) ** (y - b) + (x - a) ** b + a ** (y - b)"
  1057 by (simp add: diff_left diff_right)
  1058 
  1059 end
  1060 
  1061 lemma bounded_linear_ident[simp]: "bounded_linear (\<lambda>x. x)"
  1062   by default (auto intro!: exI[of _ 1])
  1063 
  1064 lemma bounded_linear_zero[simp]: "bounded_linear (\<lambda>x. 0)"
  1065   by default (auto intro!: exI[of _ 1])
  1066 
  1067 lemma bounded_linear_add:
  1068   assumes "bounded_linear f"
  1069   assumes "bounded_linear g"
  1070   shows "bounded_linear (\<lambda>x. f x + g x)"
  1071 proof -
  1072   interpret f: bounded_linear f by fact
  1073   interpret g: bounded_linear g by fact
  1074   show ?thesis
  1075   proof
  1076     from f.bounded obtain Kf where Kf: "\<And>x. norm (f x) \<le> norm x * Kf" by blast
  1077     from g.bounded obtain Kg where Kg: "\<And>x. norm (g x) \<le> norm x * Kg" by blast
  1078     show "\<exists>K. \<forall>x. norm (f x + g x) \<le> norm x * K"
  1079       using add_mono[OF Kf Kg]
  1080       by (intro exI[of _ "Kf + Kg"]) (auto simp: field_simps intro: norm_triangle_ineq order_trans)
  1081   qed (simp_all add: f.add g.add f.scaleR g.scaleR scaleR_right_distrib)
  1082 qed
  1083 
  1084 lemma bounded_linear_minus:
  1085   assumes "bounded_linear f"
  1086   shows "bounded_linear (\<lambda>x. - f x)"
  1087 proof -
  1088   interpret f: bounded_linear f by fact
  1089   show ?thesis apply (unfold_locales)
  1090     apply (simp add: f.add)
  1091     apply (simp add: f.scaleR)
  1092     apply (simp add: f.bounded)
  1093     done
  1094 qed
  1095 
  1096 lemma bounded_linear_compose:
  1097   assumes "bounded_linear f"
  1098   assumes "bounded_linear g"
  1099   shows "bounded_linear (\<lambda>x. f (g x))"
  1100 proof -
  1101   interpret f: bounded_linear f by fact
  1102   interpret g: bounded_linear g by fact
  1103   show ?thesis proof (unfold_locales)
  1104     fix x y show "f (g (x + y)) = f (g x) + f (g y)"
  1105       by (simp only: f.add g.add)
  1106   next
  1107     fix r x show "f (g (scaleR r x)) = scaleR r (f (g x))"
  1108       by (simp only: f.scaleR g.scaleR)
  1109   next
  1110     from f.pos_bounded
  1111     obtain Kf where f: "\<And>x. norm (f x) \<le> norm x * Kf" and Kf: "0 < Kf" by fast
  1112     from g.pos_bounded
  1113     obtain Kg where g: "\<And>x. norm (g x) \<le> norm x * Kg" by fast
  1114     show "\<exists>K. \<forall>x. norm (f (g x)) \<le> norm x * K"
  1115     proof (intro exI allI)
  1116       fix x
  1117       have "norm (f (g x)) \<le> norm (g x) * Kf"
  1118         using f .
  1119       also have "\<dots> \<le> (norm x * Kg) * Kf"
  1120         using g Kf [THEN order_less_imp_le] by (rule mult_right_mono)
  1121       also have "(norm x * Kg) * Kf = norm x * (Kg * Kf)"
  1122         by (rule mult_assoc)
  1123       finally show "norm (f (g x)) \<le> norm x * (Kg * Kf)" .
  1124     qed
  1125   qed
  1126 qed
  1127 
  1128 lemma bounded_bilinear_mult:
  1129   "bounded_bilinear (op * :: 'a \<Rightarrow> 'a \<Rightarrow> 'a::real_normed_algebra)"
  1130 apply (rule bounded_bilinear.intro)
  1131 apply (rule distrib_right)
  1132 apply (rule distrib_left)
  1133 apply (rule mult_scaleR_left)
  1134 apply (rule mult_scaleR_right)
  1135 apply (rule_tac x="1" in exI)
  1136 apply (simp add: norm_mult_ineq)
  1137 done
  1138 
  1139 lemma bounded_linear_mult_left:
  1140   "bounded_linear (\<lambda>x::'a::real_normed_algebra. x * y)"
  1141   using bounded_bilinear_mult
  1142   by (rule bounded_bilinear.bounded_linear_left)
  1143 
  1144 lemma bounded_linear_mult_right:
  1145   "bounded_linear (\<lambda>y::'a::real_normed_algebra. x * y)"
  1146   using bounded_bilinear_mult
  1147   by (rule bounded_bilinear.bounded_linear_right)
  1148 
  1149 lemmas bounded_linear_mult_const =
  1150   bounded_linear_mult_left [THEN bounded_linear_compose]
  1151 
  1152 lemmas bounded_linear_const_mult =
  1153   bounded_linear_mult_right [THEN bounded_linear_compose]
  1154 
  1155 lemma bounded_linear_divide:
  1156   "bounded_linear (\<lambda>x::'a::real_normed_field. x / y)"
  1157   unfolding divide_inverse by (rule bounded_linear_mult_left)
  1158 
  1159 lemma bounded_bilinear_scaleR: "bounded_bilinear scaleR"
  1160 apply (rule bounded_bilinear.intro)
  1161 apply (rule scaleR_left_distrib)
  1162 apply (rule scaleR_right_distrib)
  1163 apply simp
  1164 apply (rule scaleR_left_commute)
  1165 apply (rule_tac x="1" in exI, simp)
  1166 done
  1167 
  1168 lemma bounded_linear_scaleR_left: "bounded_linear (\<lambda>r. scaleR r x)"
  1169   using bounded_bilinear_scaleR
  1170   by (rule bounded_bilinear.bounded_linear_left)
  1171 
  1172 lemma bounded_linear_scaleR_right: "bounded_linear (\<lambda>x. scaleR r x)"
  1173   using bounded_bilinear_scaleR
  1174   by (rule bounded_bilinear.bounded_linear_right)
  1175 
  1176 lemma bounded_linear_of_real: "bounded_linear (\<lambda>r. of_real r)"
  1177   unfolding of_real_def by (rule bounded_linear_scaleR_left)
  1178 
  1179 lemma real_bounded_linear:
  1180   fixes f :: "real \<Rightarrow> real"
  1181   shows "bounded_linear f \<longleftrightarrow> (\<exists>c::real. f = (\<lambda>x. x * c))"
  1182 proof -
  1183   { fix x assume "bounded_linear f"
  1184     then interpret bounded_linear f .
  1185     from scaleR[of x 1] have "f x = x * f 1"
  1186       by simp }
  1187   then show ?thesis
  1188     by (auto intro: exI[of _ "f 1"] bounded_linear_mult_left)
  1189 qed
  1190 
  1191 instance real_normed_algebra_1 \<subseteq> perfect_space
  1192 proof
  1193   fix x::'a
  1194   show "\<not> open {x}"
  1195     unfolding open_dist dist_norm
  1196     by (clarsimp, rule_tac x="x + of_real (e/2)" in exI, simp)
  1197 qed
  1198 
  1199 subsection {* Filters and Limits on Metric Space *}
  1200 
  1201 lemma eventually_nhds_metric:
  1202   fixes a :: "'a :: metric_space"
  1203   shows "eventually P (nhds a) \<longleftrightarrow> (\<exists>d>0. \<forall>x. dist x a < d \<longrightarrow> P x)"
  1204 unfolding eventually_nhds open_dist
  1205 apply safe
  1206 apply fast
  1207 apply (rule_tac x="{x. dist x a < d}" in exI, simp)
  1208 apply clarsimp
  1209 apply (rule_tac x="d - dist x a" in exI, clarsimp)
  1210 apply (simp only: less_diff_eq)
  1211 apply (erule le_less_trans [OF dist_triangle])
  1212 done
  1213 
  1214 lemma eventually_at:
  1215   fixes a :: "'a :: metric_space"
  1216   shows "eventually P (at a within S) \<longleftrightarrow> (\<exists>d>0. \<forall>x\<in>S. x \<noteq> a \<and> dist x a < d \<longrightarrow> P x)"
  1217   unfolding eventually_at_filter eventually_nhds_metric by (auto simp: dist_nz)
  1218 
  1219 lemma eventually_at_le:
  1220   fixes a :: "'a::metric_space"
  1221   shows "eventually P (at a within S) \<longleftrightarrow> (\<exists>d>0. \<forall>x\<in>S. x \<noteq> a \<and> dist x a \<le> d \<longrightarrow> P x)"
  1222   unfolding eventually_at_filter eventually_nhds_metric
  1223   apply auto
  1224   apply (rule_tac x="d / 2" in exI)
  1225   apply auto
  1226   done
  1227 
  1228 lemma tendstoI:
  1229   fixes l :: "'a :: metric_space"
  1230   assumes "\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>x. dist (f x) l < e) F"
  1231   shows "(f ---> l) F"
  1232   apply (rule topological_tendstoI)
  1233   apply (simp add: open_dist)
  1234   apply (drule (1) bspec, clarify)
  1235   apply (drule assms)
  1236   apply (erule eventually_elim1, simp)
  1237   done
  1238 
  1239 lemma tendstoD:
  1240   fixes l :: "'a :: metric_space"
  1241   shows "(f ---> l) F \<Longrightarrow> 0 < e \<Longrightarrow> eventually (\<lambda>x. dist (f x) l < e) F"
  1242   apply (drule_tac S="{x. dist x l < e}" in topological_tendstoD)
  1243   apply (clarsimp simp add: open_dist)
  1244   apply (rule_tac x="e - dist x l" in exI, clarsimp)
  1245   apply (simp only: less_diff_eq)
  1246   apply (erule le_less_trans [OF dist_triangle])
  1247   apply simp
  1248   apply simp
  1249   done
  1250 
  1251 lemma tendsto_iff:
  1252   fixes l :: "'a :: metric_space"
  1253   shows "(f ---> l) F \<longleftrightarrow> (\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) F)"
  1254   using tendstoI tendstoD by fast
  1255 
  1256 lemma metric_tendsto_imp_tendsto:
  1257   fixes a :: "'a :: metric_space" and b :: "'b :: metric_space"
  1258   assumes f: "(f ---> a) F"
  1259   assumes le: "eventually (\<lambda>x. dist (g x) b \<le> dist (f x) a) F"
  1260   shows "(g ---> b) F"
  1261 proof (rule tendstoI)
  1262   fix e :: real assume "0 < e"
  1263   with f have "eventually (\<lambda>x. dist (f x) a < e) F" by (rule tendstoD)
  1264   with le show "eventually (\<lambda>x. dist (g x) b < e) F"
  1265     using le_less_trans by (rule eventually_elim2)
  1266 qed
  1267 
  1268 lemma filterlim_real_sequentially: "LIM x sequentially. real x :> at_top"
  1269   unfolding filterlim_at_top
  1270   apply (intro allI)
  1271   apply (rule_tac c="natceiling (Z + 1)" in eventually_sequentiallyI)
  1272   apply (auto simp: natceiling_le_eq)
  1273   done
  1274 
  1275 subsubsection {* Limits of Sequences *}
  1276 
  1277 lemma LIMSEQ_def: "X ----> (L::'a::metric_space) \<longleftrightarrow> (\<forall>r>0. \<exists>no. \<forall>n\<ge>no. dist (X n) L < r)"
  1278   unfolding tendsto_iff eventually_sequentially ..
  1279 
  1280 lemma LIMSEQ_iff_nz: "X ----> (L::'a::metric_space) = (\<forall>r>0. \<exists>no>0. \<forall>n\<ge>no. dist (X n) L < r)"
  1281   unfolding LIMSEQ_def by (metis Suc_leD zero_less_Suc)
  1282 
  1283 lemma metric_LIMSEQ_I:
  1284   "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. dist (X n) L < r) \<Longrightarrow> X ----> (L::'a::metric_space)"
  1285 by (simp add: LIMSEQ_def)
  1286 
  1287 lemma metric_LIMSEQ_D:
  1288   "\<lbrakk>X ----> (L::'a::metric_space); 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. dist (X n) L < r"
  1289 by (simp add: LIMSEQ_def)
  1290 
  1291 
  1292 subsubsection {* Limits of Functions *}
  1293 
  1294 lemma LIM_def: "f -- (a::'a::metric_space) --> (L::'b::metric_space) =
  1295      (\<forall>r > 0. \<exists>s > 0. \<forall>x. x \<noteq> a & dist x a < s
  1296         --> dist (f x) L < r)"
  1297   unfolding tendsto_iff eventually_at by simp
  1298 
  1299 lemma metric_LIM_I:
  1300   "(\<And>r. 0 < r \<Longrightarrow> \<exists>s>0. \<forall>x. x \<noteq> a \<and> dist x a < s \<longrightarrow> dist (f x) L < r)
  1301     \<Longrightarrow> f -- (a::'a::metric_space) --> (L::'b::metric_space)"
  1302 by (simp add: LIM_def)
  1303 
  1304 lemma metric_LIM_D:
  1305   "\<lbrakk>f -- (a::'a::metric_space) --> (L::'b::metric_space); 0 < r\<rbrakk>
  1306     \<Longrightarrow> \<exists>s>0. \<forall>x. x \<noteq> a \<and> dist x a < s \<longrightarrow> dist (f x) L < r"
  1307 by (simp add: LIM_def)
  1308 
  1309 lemma metric_LIM_imp_LIM:
  1310   assumes f: "f -- a --> (l::'a::metric_space)"
  1311   assumes le: "\<And>x. x \<noteq> a \<Longrightarrow> dist (g x) m \<le> dist (f x) l"
  1312   shows "g -- a --> (m::'b::metric_space)"
  1313   by (rule metric_tendsto_imp_tendsto [OF f]) (auto simp add: eventually_at_topological le)
  1314 
  1315 lemma metric_LIM_equal2:
  1316   assumes 1: "0 < R"
  1317   assumes 2: "\<And>x. \<lbrakk>x \<noteq> a; dist x a < R\<rbrakk> \<Longrightarrow> f x = g x"
  1318   shows "g -- a --> l \<Longrightarrow> f -- (a::'a::metric_space) --> l"
  1319 apply (rule topological_tendstoI)
  1320 apply (drule (2) topological_tendstoD)
  1321 apply (simp add: eventually_at, safe)
  1322 apply (rule_tac x="min d R" in exI, safe)
  1323 apply (simp add: 1)
  1324 apply (simp add: 2)
  1325 done
  1326 
  1327 lemma metric_LIM_compose2:
  1328   assumes f: "f -- (a::'a::metric_space) --> b"
  1329   assumes g: "g -- b --> c"
  1330   assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> f x \<noteq> b"
  1331   shows "(\<lambda>x. g (f x)) -- a --> c"
  1332   using inj
  1333   by (intro tendsto_compose_eventually[OF g f]) (auto simp: eventually_at)
  1334 
  1335 lemma metric_isCont_LIM_compose2:
  1336   fixes f :: "'a :: metric_space \<Rightarrow> _"
  1337   assumes f [unfolded isCont_def]: "isCont f a"
  1338   assumes g: "g -- f a --> l"
  1339   assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> f x \<noteq> f a"
  1340   shows "(\<lambda>x. g (f x)) -- a --> l"
  1341 by (rule metric_LIM_compose2 [OF f g inj])
  1342 
  1343 subsection {* Complete metric spaces *}
  1344 
  1345 subsection {* Cauchy sequences *}
  1346 
  1347 definition (in metric_space) Cauchy :: "(nat \<Rightarrow> 'a) \<Rightarrow> bool" where
  1348   "Cauchy X = (\<forall>e>0. \<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. dist (X m) (X n) < e)"
  1349 
  1350 subsection {* Cauchy Sequences *}
  1351 
  1352 lemma metric_CauchyI:
  1353   "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < e) \<Longrightarrow> Cauchy X"
  1354   by (simp add: Cauchy_def)
  1355 
  1356 lemma metric_CauchyD:
  1357   "Cauchy X \<Longrightarrow> 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < e"
  1358   by (simp add: Cauchy_def)
  1359 
  1360 lemma metric_Cauchy_iff2:
  1361   "Cauchy X = (\<forall>j. (\<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. dist (X m) (X n) < inverse(real (Suc j))))"
  1362 apply (simp add: Cauchy_def, auto)
  1363 apply (drule reals_Archimedean, safe)
  1364 apply (drule_tac x = n in spec, auto)
  1365 apply (rule_tac x = M in exI, auto)
  1366 apply (drule_tac x = m in spec, simp)
  1367 apply (drule_tac x = na in spec, auto)
  1368 done
  1369 
  1370 lemma Cauchy_iff2:
  1371   "Cauchy X = (\<forall>j. (\<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. \<bar>X m - X n\<bar> < inverse(real (Suc j))))"
  1372   unfolding metric_Cauchy_iff2 dist_real_def ..
  1373 
  1374 lemma Cauchy_subseq_Cauchy:
  1375   "\<lbrakk> Cauchy X; subseq f \<rbrakk> \<Longrightarrow> Cauchy (X o f)"
  1376 apply (auto simp add: Cauchy_def)
  1377 apply (drule_tac x=e in spec, clarify)
  1378 apply (rule_tac x=M in exI, clarify)
  1379 apply (blast intro: le_trans [OF _ seq_suble] dest!: spec)
  1380 done
  1381 
  1382 theorem LIMSEQ_imp_Cauchy:
  1383   assumes X: "X ----> a" shows "Cauchy X"
  1384 proof (rule metric_CauchyI)
  1385   fix e::real assume "0 < e"
  1386   hence "0 < e/2" by simp
  1387   with X have "\<exists>N. \<forall>n\<ge>N. dist (X n) a < e/2" by (rule metric_LIMSEQ_D)
  1388   then obtain N where N: "\<forall>n\<ge>N. dist (X n) a < e/2" ..
  1389   show "\<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m) (X n) < e"
  1390   proof (intro exI allI impI)
  1391     fix m assume "N \<le> m"
  1392     hence m: "dist (X m) a < e/2" using N by fast
  1393     fix n assume "N \<le> n"
  1394     hence n: "dist (X n) a < e/2" using N by fast
  1395     have "dist (X m) (X n) \<le> dist (X m) a + dist (X n) a"
  1396       by (rule dist_triangle2)
  1397     also from m n have "\<dots> < e" by simp
  1398     finally show "dist (X m) (X n) < e" .
  1399   qed
  1400 qed
  1401 
  1402 lemma convergent_Cauchy: "convergent X \<Longrightarrow> Cauchy X"
  1403 unfolding convergent_def
  1404 by (erule exE, erule LIMSEQ_imp_Cauchy)
  1405 
  1406 subsubsection {* Cauchy Sequences are Convergent *}
  1407 
  1408 class complete_space = metric_space +
  1409   assumes Cauchy_convergent: "Cauchy X \<Longrightarrow> convergent X"
  1410 
  1411 lemma Cauchy_convergent_iff:
  1412   fixes X :: "nat \<Rightarrow> 'a::complete_space"
  1413   shows "Cauchy X = convergent X"
  1414 by (fast intro: Cauchy_convergent convergent_Cauchy)
  1415 
  1416 subsection {* The set of real numbers is a complete metric space *}
  1417 
  1418 text {*
  1419 Proof that Cauchy sequences converge based on the one from
  1420 http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html
  1421 *}
  1422 
  1423 text {*
  1424   If sequence @{term "X"} is Cauchy, then its limit is the lub of
  1425   @{term "{r::real. \<exists>N. \<forall>n\<ge>N. r < X n}"}
  1426 *}
  1427 
  1428 lemma increasing_LIMSEQ:
  1429   fixes f :: "nat \<Rightarrow> real"
  1430   assumes inc: "\<And>n. f n \<le> f (Suc n)"
  1431       and bdd: "\<And>n. f n \<le> l"
  1432       and en: "\<And>e. 0 < e \<Longrightarrow> \<exists>n. l \<le> f n + e"
  1433   shows "f ----> l"
  1434 proof (rule increasing_tendsto)
  1435   fix x assume "x < l"
  1436   with dense[of 0 "l - x"] obtain e where "0 < e" "e < l - x"
  1437     by auto
  1438   from en[OF `0 < e`] obtain n where "l - e \<le> f n"
  1439     by (auto simp: field_simps)
  1440   with `e < l - x` `0 < e` have "x < f n" by simp
  1441   with incseq_SucI[of f, OF inc] show "eventually (\<lambda>n. x < f n) sequentially"
  1442     by (auto simp: eventually_sequentially incseq_def intro: less_le_trans)
  1443 qed (insert bdd, auto)
  1444 
  1445 lemma real_Cauchy_convergent:
  1446   fixes X :: "nat \<Rightarrow> real"
  1447   assumes X: "Cauchy X"
  1448   shows "convergent X"
  1449 proof -
  1450   def S \<equiv> "{x::real. \<exists>N. \<forall>n\<ge>N. x < X n}"
  1451   then have mem_S: "\<And>N x. \<forall>n\<ge>N. x < X n \<Longrightarrow> x \<in> S" by auto
  1452 
  1453   { fix N x assume N: "\<forall>n\<ge>N. X n < x"
  1454   fix y::real assume "y \<in> S"
  1455   hence "\<exists>M. \<forall>n\<ge>M. y < X n"
  1456     by (simp add: S_def)
  1457   then obtain M where "\<forall>n\<ge>M. y < X n" ..
  1458   hence "y < X (max M N)" by simp
  1459   also have "\<dots> < x" using N by simp
  1460   finally have "y \<le> x"
  1461     by (rule order_less_imp_le) }
  1462   note bound_isUb = this 
  1463 
  1464   obtain N where "\<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m) (X n) < 1"
  1465     using X[THEN metric_CauchyD, OF zero_less_one] by auto
  1466   hence N: "\<forall>n\<ge>N. dist (X n) (X N) < 1" by simp
  1467   have [simp]: "S \<noteq> {}"
  1468   proof (intro exI ex_in_conv[THEN iffD1])
  1469     from N have "\<forall>n\<ge>N. X N - 1 < X n"
  1470       by (simp add: abs_diff_less_iff dist_real_def)
  1471     thus "X N - 1 \<in> S" by (rule mem_S)
  1472   qed
  1473   have [simp]: "bdd_above S"
  1474   proof
  1475     from N have "\<forall>n\<ge>N. X n < X N + 1"
  1476       by (simp add: abs_diff_less_iff dist_real_def)
  1477     thus "\<And>s. s \<in> S \<Longrightarrow>  s \<le> X N + 1"
  1478       by (rule bound_isUb)
  1479   qed
  1480   have "X ----> Sup S"
  1481   proof (rule metric_LIMSEQ_I)
  1482   fix r::real assume "0 < r"
  1483   hence r: "0 < r/2" by simp
  1484   obtain N where "\<forall>n\<ge>N. \<forall>m\<ge>N. dist (X n) (X m) < r/2"
  1485     using metric_CauchyD [OF X r] by auto
  1486   hence "\<forall>n\<ge>N. dist (X n) (X N) < r/2" by simp
  1487   hence N: "\<forall>n\<ge>N. X N - r/2 < X n \<and> X n < X N + r/2"
  1488     by (simp only: dist_real_def abs_diff_less_iff)
  1489 
  1490   from N have "\<forall>n\<ge>N. X N - r/2 < X n" by fast
  1491   hence "X N - r/2 \<in> S" by (rule mem_S)
  1492   hence 1: "X N - r/2 \<le> Sup S" by (simp add: cSup_upper)
  1493 
  1494   from N have "\<forall>n\<ge>N. X n < X N + r/2" by fast
  1495   from bound_isUb[OF this]
  1496   have 2: "Sup S \<le> X N + r/2"
  1497     by (intro cSup_least) simp_all
  1498 
  1499   show "\<exists>N. \<forall>n\<ge>N. dist (X n) (Sup S) < r"
  1500   proof (intro exI allI impI)
  1501     fix n assume n: "N \<le> n"
  1502     from N n have "X n < X N + r/2" and "X N - r/2 < X n" by simp+
  1503     thus "dist (X n) (Sup S) < r" using 1 2
  1504       by (simp add: abs_diff_less_iff dist_real_def)
  1505   qed
  1506   qed
  1507   then show ?thesis unfolding convergent_def by auto
  1508 qed
  1509 
  1510 instance real :: complete_space
  1511   by intro_classes (rule real_Cauchy_convergent)
  1512 
  1513 class banach = real_normed_vector + complete_space
  1514 
  1515 instance real :: banach by default
  1516 
  1517 lemma tendsto_at_topI_sequentially:
  1518   fixes f :: "real \<Rightarrow> real"
  1519   assumes mono: "mono f"
  1520   assumes limseq: "(\<lambda>n. f (real n)) ----> y"
  1521   shows "(f ---> y) at_top"
  1522 proof (rule tendstoI)
  1523   fix e :: real assume "0 < e"
  1524   with limseq obtain N :: nat where N: "\<And>n. N \<le> n \<Longrightarrow> \<bar>f (real n) - y\<bar> < e"
  1525     by (auto simp: LIMSEQ_def dist_real_def)
  1526   { fix x :: real
  1527     obtain n where "x \<le> real_of_nat n"
  1528       using ex_le_of_nat[of x] ..
  1529     note monoD[OF mono this]
  1530     also have "f (real_of_nat n) \<le> y"
  1531       by (rule LIMSEQ_le_const[OF limseq])
  1532          (auto intro: exI[of _ n] monoD[OF mono] simp: real_eq_of_nat[symmetric])
  1533     finally have "f x \<le> y" . }
  1534   note le = this
  1535   have "eventually (\<lambda>x. real N \<le> x) at_top"
  1536     by (rule eventually_ge_at_top)
  1537   then show "eventually (\<lambda>x. dist (f x) y < e) at_top"
  1538   proof eventually_elim
  1539     fix x assume N': "real N \<le> x"
  1540     with N[of N] le have "y - f (real N) < e" by auto
  1541     moreover note monoD[OF mono N']
  1542     ultimately show "dist (f x) y < e"
  1543       using le[of x] by (auto simp: dist_real_def field_simps)
  1544   qed
  1545 qed
  1546 
  1547 end