src/HOL/Real/RealVector.thy
changeset 29197 6d4cb27ed19c
parent 29189 ee8572f3bb57
child 29198 418ed6411847
equal deleted inserted replaced
29189:ee8572f3bb57 29197:6d4cb27ed19c
     1 (*  Title:      HOL/Real/RealVector.thy
       
     2     Author:     Brian Huffman
       
     3 *)
       
     4 
       
     5 header {* Vector Spaces and Algebras over the Reals *}
       
     6 
       
     7 theory RealVector
       
     8 imports "~~/src/HOL/RealPow"
       
     9 begin
       
    10 
       
    11 subsection {* Locale for additive functions *}
       
    12 
       
    13 locale additive =
       
    14   fixes f :: "'a::ab_group_add \<Rightarrow> 'b::ab_group_add"
       
    15   assumes add: "f (x + y) = f x + f y"
       
    16 begin
       
    17 
       
    18 lemma zero: "f 0 = 0"
       
    19 proof -
       
    20   have "f 0 = f (0 + 0)" by simp
       
    21   also have "\<dots> = f 0 + f 0" by (rule add)
       
    22   finally show "f 0 = 0" by simp
       
    23 qed
       
    24 
       
    25 lemma minus: "f (- x) = - f x"
       
    26 proof -
       
    27   have "f (- x) + f x = f (- x + x)" by (rule add [symmetric])
       
    28   also have "\<dots> = - f x + f x" by (simp add: zero)
       
    29   finally show "f (- x) = - f x" by (rule add_right_imp_eq)
       
    30 qed
       
    31 
       
    32 lemma diff: "f (x - y) = f x - f y"
       
    33 by (simp add: diff_def add minus)
       
    34 
       
    35 lemma setsum: "f (setsum g A) = (\<Sum>x\<in>A. f (g x))"
       
    36 apply (cases "finite A")
       
    37 apply (induct set: finite)
       
    38 apply (simp add: zero)
       
    39 apply (simp add: add)
       
    40 apply (simp add: zero)
       
    41 done
       
    42 
       
    43 end
       
    44 
       
    45 subsection {* Vector spaces *}
       
    46 
       
    47 locale vector_space =
       
    48   fixes scale :: "'a::field \<Rightarrow> 'b::ab_group_add \<Rightarrow> 'b"
       
    49   assumes scale_right_distrib: "scale a (x + y) = scale a x + scale a y"
       
    50   and scale_left_distrib: "scale (a + b) x = scale a x + scale b x"
       
    51   and scale_scale [simp]: "scale a (scale b x) = scale (a * b) x"
       
    52   and scale_one [simp]: "scale 1 x = x"
       
    53 begin
       
    54 
       
    55 lemma scale_left_commute:
       
    56   "scale a (scale b x) = scale b (scale a x)"
       
    57 by (simp add: mult_commute)
       
    58 
       
    59 lemma scale_zero_left [simp]: "scale 0 x = 0"
       
    60   and scale_minus_left [simp]: "scale (- a) x = - (scale a x)"
       
    61   and scale_left_diff_distrib: "scale (a - b) x = scale a x - scale b x"
       
    62 proof -
       
    63   interpret s: additive ["\<lambda>a. scale a x"]
       
    64     proof qed (rule scale_left_distrib)
       
    65   show "scale 0 x = 0" by (rule s.zero)
       
    66   show "scale (- a) x = - (scale a x)" by (rule s.minus)
       
    67   show "scale (a - b) x = scale a x - scale b x" by (rule s.diff)
       
    68 qed
       
    69 
       
    70 lemma scale_zero_right [simp]: "scale a 0 = 0"
       
    71   and scale_minus_right [simp]: "scale a (- x) = - (scale a x)"
       
    72   and scale_right_diff_distrib: "scale a (x - y) = scale a x - scale a y"
       
    73 proof -
       
    74   interpret s: additive ["\<lambda>x. scale a x"]
       
    75     proof qed (rule scale_right_distrib)
       
    76   show "scale a 0 = 0" by (rule s.zero)
       
    77   show "scale a (- x) = - (scale a x)" by (rule s.minus)
       
    78   show "scale a (x - y) = scale a x - scale a y" by (rule s.diff)
       
    79 qed
       
    80 
       
    81 lemma scale_eq_0_iff [simp]:
       
    82   "scale a x = 0 \<longleftrightarrow> a = 0 \<or> x = 0"
       
    83 proof cases
       
    84   assume "a = 0" thus ?thesis by simp
       
    85 next
       
    86   assume anz [simp]: "a \<noteq> 0"
       
    87   { assume "scale a x = 0"
       
    88     hence "scale (inverse a) (scale a x) = 0" by simp
       
    89     hence "x = 0" by simp }
       
    90   thus ?thesis by force
       
    91 qed
       
    92 
       
    93 lemma scale_left_imp_eq:
       
    94   "\<lbrakk>a \<noteq> 0; scale a x = scale a y\<rbrakk> \<Longrightarrow> x = y"
       
    95 proof -
       
    96   assume nonzero: "a \<noteq> 0"
       
    97   assume "scale a x = scale a y"
       
    98   hence "scale a (x - y) = 0"
       
    99      by (simp add: scale_right_diff_distrib)
       
   100   hence "x - y = 0" by (simp add: nonzero)
       
   101   thus "x = y" by (simp only: right_minus_eq)
       
   102 qed
       
   103 
       
   104 lemma scale_right_imp_eq:
       
   105   "\<lbrakk>x \<noteq> 0; scale a x = scale b x\<rbrakk> \<Longrightarrow> a = b"
       
   106 proof -
       
   107   assume nonzero: "x \<noteq> 0"
       
   108   assume "scale a x = scale b x"
       
   109   hence "scale (a - b) x = 0"
       
   110      by (simp add: scale_left_diff_distrib)
       
   111   hence "a - b = 0" by (simp add: nonzero)
       
   112   thus "a = b" by (simp only: right_minus_eq)
       
   113 qed
       
   114 
       
   115 lemma scale_cancel_left:
       
   116   "scale a x = scale a y \<longleftrightarrow> x = y \<or> a = 0"
       
   117 by (auto intro: scale_left_imp_eq)
       
   118 
       
   119 lemma scale_cancel_right:
       
   120   "scale a x = scale b x \<longleftrightarrow> a = b \<or> x = 0"
       
   121 by (auto intro: scale_right_imp_eq)
       
   122 
       
   123 end
       
   124 
       
   125 subsection {* Real vector spaces *}
       
   126 
       
   127 class scaleR = type +
       
   128   fixes scaleR :: "real \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "*\<^sub>R" 75)
       
   129 begin
       
   130 
       
   131 abbreviation
       
   132   divideR :: "'a \<Rightarrow> real \<Rightarrow> 'a" (infixl "'/\<^sub>R" 70)
       
   133 where
       
   134   "x /\<^sub>R r == scaleR (inverse r) x"
       
   135 
       
   136 end
       
   137 
       
   138 instantiation real :: scaleR
       
   139 begin
       
   140 
       
   141 definition
       
   142   real_scaleR_def [simp]: "scaleR a x = a * x"
       
   143 
       
   144 instance ..
       
   145 
       
   146 end
       
   147 
       
   148 class real_vector = scaleR + ab_group_add +
       
   149   assumes scaleR_right_distrib: "scaleR a (x + y) = scaleR a x + scaleR a y"
       
   150   and scaleR_left_distrib: "scaleR (a + b) x = scaleR a x + scaleR b x"
       
   151   and scaleR_scaleR [simp]: "scaleR a (scaleR b x) = scaleR (a * b) x"
       
   152   and scaleR_one [simp]: "scaleR 1 x = x"
       
   153 
       
   154 interpretation real_vector:
       
   155   vector_space ["scaleR :: real \<Rightarrow> 'a \<Rightarrow> 'a::real_vector"]
       
   156 apply unfold_locales
       
   157 apply (rule scaleR_right_distrib)
       
   158 apply (rule scaleR_left_distrib)
       
   159 apply (rule scaleR_scaleR)
       
   160 apply (rule scaleR_one)
       
   161 done
       
   162 
       
   163 text {* Recover original theorem names *}
       
   164 
       
   165 lemmas scaleR_left_commute = real_vector.scale_left_commute
       
   166 lemmas scaleR_zero_left = real_vector.scale_zero_left
       
   167 lemmas scaleR_minus_left = real_vector.scale_minus_left
       
   168 lemmas scaleR_left_diff_distrib = real_vector.scale_left_diff_distrib
       
   169 lemmas scaleR_zero_right = real_vector.scale_zero_right
       
   170 lemmas scaleR_minus_right = real_vector.scale_minus_right
       
   171 lemmas scaleR_right_diff_distrib = real_vector.scale_right_diff_distrib
       
   172 lemmas scaleR_eq_0_iff = real_vector.scale_eq_0_iff
       
   173 lemmas scaleR_left_imp_eq = real_vector.scale_left_imp_eq
       
   174 lemmas scaleR_right_imp_eq = real_vector.scale_right_imp_eq
       
   175 lemmas scaleR_cancel_left = real_vector.scale_cancel_left
       
   176 lemmas scaleR_cancel_right = real_vector.scale_cancel_right
       
   177 
       
   178 class real_algebra = real_vector + ring +
       
   179   assumes mult_scaleR_left [simp]: "scaleR a x * y = scaleR a (x * y)"
       
   180   and mult_scaleR_right [simp]: "x * scaleR a y = scaleR a (x * y)"
       
   181 
       
   182 class real_algebra_1 = real_algebra + ring_1
       
   183 
       
   184 class real_div_algebra = real_algebra_1 + division_ring
       
   185 
       
   186 class real_field = real_div_algebra + field
       
   187 
       
   188 instance real :: real_field
       
   189 apply (intro_classes, unfold real_scaleR_def)
       
   190 apply (rule right_distrib)
       
   191 apply (rule left_distrib)
       
   192 apply (rule mult_assoc [symmetric])
       
   193 apply (rule mult_1_left)
       
   194 apply (rule mult_assoc)
       
   195 apply (rule mult_left_commute)
       
   196 done
       
   197 
       
   198 interpretation scaleR_left: additive ["(\<lambda>a. scaleR a x::'a::real_vector)"]
       
   199 proof qed (rule scaleR_left_distrib)
       
   200 
       
   201 interpretation scaleR_right: additive ["(\<lambda>x. scaleR a x::'a::real_vector)"]
       
   202 proof qed (rule scaleR_right_distrib)
       
   203 
       
   204 lemma nonzero_inverse_scaleR_distrib:
       
   205   fixes x :: "'a::real_div_algebra" shows
       
   206   "\<lbrakk>a \<noteq> 0; x \<noteq> 0\<rbrakk> \<Longrightarrow> inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
       
   207 by (rule inverse_unique, simp)
       
   208 
       
   209 lemma inverse_scaleR_distrib:
       
   210   fixes x :: "'a::{real_div_algebra,division_by_zero}"
       
   211   shows "inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
       
   212 apply (case_tac "a = 0", simp)
       
   213 apply (case_tac "x = 0", simp)
       
   214 apply (erule (1) nonzero_inverse_scaleR_distrib)
       
   215 done
       
   216 
       
   217 
       
   218 subsection {* Embedding of the Reals into any @{text real_algebra_1}:
       
   219 @{term of_real} *}
       
   220 
       
   221 definition
       
   222   of_real :: "real \<Rightarrow> 'a::real_algebra_1" where
       
   223   "of_real r = scaleR r 1"
       
   224 
       
   225 lemma scaleR_conv_of_real: "scaleR r x = of_real r * x"
       
   226 by (simp add: of_real_def)
       
   227 
       
   228 lemma of_real_0 [simp]: "of_real 0 = 0"
       
   229 by (simp add: of_real_def)
       
   230 
       
   231 lemma of_real_1 [simp]: "of_real 1 = 1"
       
   232 by (simp add: of_real_def)
       
   233 
       
   234 lemma of_real_add [simp]: "of_real (x + y) = of_real x + of_real y"
       
   235 by (simp add: of_real_def scaleR_left_distrib)
       
   236 
       
   237 lemma of_real_minus [simp]: "of_real (- x) = - of_real x"
       
   238 by (simp add: of_real_def)
       
   239 
       
   240 lemma of_real_diff [simp]: "of_real (x - y) = of_real x - of_real y"
       
   241 by (simp add: of_real_def scaleR_left_diff_distrib)
       
   242 
       
   243 lemma of_real_mult [simp]: "of_real (x * y) = of_real x * of_real y"
       
   244 by (simp add: of_real_def mult_commute)
       
   245 
       
   246 lemma nonzero_of_real_inverse:
       
   247   "x \<noteq> 0 \<Longrightarrow> of_real (inverse x) =
       
   248    inverse (of_real x :: 'a::real_div_algebra)"
       
   249 by (simp add: of_real_def nonzero_inverse_scaleR_distrib)
       
   250 
       
   251 lemma of_real_inverse [simp]:
       
   252   "of_real (inverse x) =
       
   253    inverse (of_real x :: 'a::{real_div_algebra,division_by_zero})"
       
   254 by (simp add: of_real_def inverse_scaleR_distrib)
       
   255 
       
   256 lemma nonzero_of_real_divide:
       
   257   "y \<noteq> 0 \<Longrightarrow> of_real (x / y) =
       
   258    (of_real x / of_real y :: 'a::real_field)"
       
   259 by (simp add: divide_inverse nonzero_of_real_inverse)
       
   260 
       
   261 lemma of_real_divide [simp]:
       
   262   "of_real (x / y) =
       
   263    (of_real x / of_real y :: 'a::{real_field,division_by_zero})"
       
   264 by (simp add: divide_inverse)
       
   265 
       
   266 lemma of_real_power [simp]:
       
   267   "of_real (x ^ n) = (of_real x :: 'a::{real_algebra_1,recpower}) ^ n"
       
   268 by (induct n) (simp_all add: power_Suc)
       
   269 
       
   270 lemma of_real_eq_iff [simp]: "(of_real x = of_real y) = (x = y)"
       
   271 by (simp add: of_real_def scaleR_cancel_right)
       
   272 
       
   273 lemmas of_real_eq_0_iff [simp] = of_real_eq_iff [of _ 0, simplified]
       
   274 
       
   275 lemma of_real_eq_id [simp]: "of_real = (id :: real \<Rightarrow> real)"
       
   276 proof
       
   277   fix r
       
   278   show "of_real r = id r"
       
   279     by (simp add: of_real_def)
       
   280 qed
       
   281 
       
   282 text{*Collapse nested embeddings*}
       
   283 lemma of_real_of_nat_eq [simp]: "of_real (of_nat n) = of_nat n"
       
   284 by (induct n) auto
       
   285 
       
   286 lemma of_real_of_int_eq [simp]: "of_real (of_int z) = of_int z"
       
   287 by (cases z rule: int_diff_cases, simp)
       
   288 
       
   289 lemma of_real_number_of_eq:
       
   290   "of_real (number_of w) = (number_of w :: 'a::{number_ring,real_algebra_1})"
       
   291 by (simp add: number_of_eq)
       
   292 
       
   293 text{*Every real algebra has characteristic zero*}
       
   294 instance real_algebra_1 < ring_char_0
       
   295 proof
       
   296   fix m n :: nat
       
   297   have "(of_real (of_nat m) = (of_real (of_nat n)::'a)) = (m = n)"
       
   298     by (simp only: of_real_eq_iff of_nat_eq_iff)
       
   299   thus "(of_nat m = (of_nat n::'a)) = (m = n)"
       
   300     by (simp only: of_real_of_nat_eq)
       
   301 qed
       
   302 
       
   303 instance real_field < field_char_0 ..
       
   304 
       
   305 
       
   306 subsection {* The Set of Real Numbers *}
       
   307 
       
   308 definition
       
   309   Reals :: "'a::real_algebra_1 set" where
       
   310   [code del]: "Reals \<equiv> range of_real"
       
   311 
       
   312 notation (xsymbols)
       
   313   Reals  ("\<real>")
       
   314 
       
   315 lemma Reals_of_real [simp]: "of_real r \<in> Reals"
       
   316 by (simp add: Reals_def)
       
   317 
       
   318 lemma Reals_of_int [simp]: "of_int z \<in> Reals"
       
   319 by (subst of_real_of_int_eq [symmetric], rule Reals_of_real)
       
   320 
       
   321 lemma Reals_of_nat [simp]: "of_nat n \<in> Reals"
       
   322 by (subst of_real_of_nat_eq [symmetric], rule Reals_of_real)
       
   323 
       
   324 lemma Reals_number_of [simp]:
       
   325   "(number_of w::'a::{number_ring,real_algebra_1}) \<in> Reals"
       
   326 by (subst of_real_number_of_eq [symmetric], rule Reals_of_real)
       
   327 
       
   328 lemma Reals_0 [simp]: "0 \<in> Reals"
       
   329 apply (unfold Reals_def)
       
   330 apply (rule range_eqI)
       
   331 apply (rule of_real_0 [symmetric])
       
   332 done
       
   333 
       
   334 lemma Reals_1 [simp]: "1 \<in> Reals"
       
   335 apply (unfold Reals_def)
       
   336 apply (rule range_eqI)
       
   337 apply (rule of_real_1 [symmetric])
       
   338 done
       
   339 
       
   340 lemma Reals_add [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a + b \<in> Reals"
       
   341 apply (auto simp add: Reals_def)
       
   342 apply (rule range_eqI)
       
   343 apply (rule of_real_add [symmetric])
       
   344 done
       
   345 
       
   346 lemma Reals_minus [simp]: "a \<in> Reals \<Longrightarrow> - a \<in> Reals"
       
   347 apply (auto simp add: Reals_def)
       
   348 apply (rule range_eqI)
       
   349 apply (rule of_real_minus [symmetric])
       
   350 done
       
   351 
       
   352 lemma Reals_diff [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a - b \<in> Reals"
       
   353 apply (auto simp add: Reals_def)
       
   354 apply (rule range_eqI)
       
   355 apply (rule of_real_diff [symmetric])
       
   356 done
       
   357 
       
   358 lemma Reals_mult [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a * b \<in> Reals"
       
   359 apply (auto simp add: Reals_def)
       
   360 apply (rule range_eqI)
       
   361 apply (rule of_real_mult [symmetric])
       
   362 done
       
   363 
       
   364 lemma nonzero_Reals_inverse:
       
   365   fixes a :: "'a::real_div_algebra"
       
   366   shows "\<lbrakk>a \<in> Reals; a \<noteq> 0\<rbrakk> \<Longrightarrow> inverse a \<in> Reals"
       
   367 apply (auto simp add: Reals_def)
       
   368 apply (rule range_eqI)
       
   369 apply (erule nonzero_of_real_inverse [symmetric])
       
   370 done
       
   371 
       
   372 lemma Reals_inverse [simp]:
       
   373   fixes a :: "'a::{real_div_algebra,division_by_zero}"
       
   374   shows "a \<in> Reals \<Longrightarrow> inverse a \<in> Reals"
       
   375 apply (auto simp add: Reals_def)
       
   376 apply (rule range_eqI)
       
   377 apply (rule of_real_inverse [symmetric])
       
   378 done
       
   379 
       
   380 lemma nonzero_Reals_divide:
       
   381   fixes a b :: "'a::real_field"
       
   382   shows "\<lbrakk>a \<in> Reals; b \<in> Reals; b \<noteq> 0\<rbrakk> \<Longrightarrow> a / b \<in> Reals"
       
   383 apply (auto simp add: Reals_def)
       
   384 apply (rule range_eqI)
       
   385 apply (erule nonzero_of_real_divide [symmetric])
       
   386 done
       
   387 
       
   388 lemma Reals_divide [simp]:
       
   389   fixes a b :: "'a::{real_field,division_by_zero}"
       
   390   shows "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a / b \<in> Reals"
       
   391 apply (auto simp add: Reals_def)
       
   392 apply (rule range_eqI)
       
   393 apply (rule of_real_divide [symmetric])
       
   394 done
       
   395 
       
   396 lemma Reals_power [simp]:
       
   397   fixes a :: "'a::{real_algebra_1,recpower}"
       
   398   shows "a \<in> Reals \<Longrightarrow> a ^ n \<in> Reals"
       
   399 apply (auto simp add: Reals_def)
       
   400 apply (rule range_eqI)
       
   401 apply (rule of_real_power [symmetric])
       
   402 done
       
   403 
       
   404 lemma Reals_cases [cases set: Reals]:
       
   405   assumes "q \<in> \<real>"
       
   406   obtains (of_real) r where "q = of_real r"
       
   407   unfolding Reals_def
       
   408 proof -
       
   409   from `q \<in> \<real>` have "q \<in> range of_real" unfolding Reals_def .
       
   410   then obtain r where "q = of_real r" ..
       
   411   then show thesis ..
       
   412 qed
       
   413 
       
   414 lemma Reals_induct [case_names of_real, induct set: Reals]:
       
   415   "q \<in> \<real> \<Longrightarrow> (\<And>r. P (of_real r)) \<Longrightarrow> P q"
       
   416   by (rule Reals_cases) auto
       
   417 
       
   418 
       
   419 subsection {* Real normed vector spaces *}
       
   420 
       
   421 class norm = type +
       
   422   fixes norm :: "'a \<Rightarrow> real"
       
   423 
       
   424 instantiation real :: norm
       
   425 begin
       
   426 
       
   427 definition
       
   428   real_norm_def [simp]: "norm r \<equiv> \<bar>r\<bar>"
       
   429 
       
   430 instance ..
       
   431 
       
   432 end
       
   433 
       
   434 class sgn_div_norm = scaleR + norm + sgn +
       
   435   assumes sgn_div_norm: "sgn x = x /\<^sub>R norm x"
       
   436 
       
   437 class real_normed_vector = real_vector + sgn_div_norm +
       
   438   assumes norm_ge_zero [simp]: "0 \<le> norm x"
       
   439   and norm_eq_zero [simp]: "norm x = 0 \<longleftrightarrow> x = 0"
       
   440   and norm_triangle_ineq: "norm (x + y) \<le> norm x + norm y"
       
   441   and norm_scaleR: "norm (scaleR a x) = \<bar>a\<bar> * norm x"
       
   442 
       
   443 class real_normed_algebra = real_algebra + real_normed_vector +
       
   444   assumes norm_mult_ineq: "norm (x * y) \<le> norm x * norm y"
       
   445 
       
   446 class real_normed_algebra_1 = real_algebra_1 + real_normed_algebra +
       
   447   assumes norm_one [simp]: "norm 1 = 1"
       
   448 
       
   449 class real_normed_div_algebra = real_div_algebra + real_normed_vector +
       
   450   assumes norm_mult: "norm (x * y) = norm x * norm y"
       
   451 
       
   452 class real_normed_field = real_field + real_normed_div_algebra
       
   453 
       
   454 instance real_normed_div_algebra < real_normed_algebra_1
       
   455 proof
       
   456   fix x y :: 'a
       
   457   show "norm (x * y) \<le> norm x * norm y"
       
   458     by (simp add: norm_mult)
       
   459 next
       
   460   have "norm (1 * 1::'a) = norm (1::'a) * norm (1::'a)"
       
   461     by (rule norm_mult)
       
   462   thus "norm (1::'a) = 1" by simp
       
   463 qed
       
   464 
       
   465 instance real :: real_normed_field
       
   466 apply (intro_classes, unfold real_norm_def real_scaleR_def)
       
   467 apply (simp add: real_sgn_def)
       
   468 apply (rule abs_ge_zero)
       
   469 apply (rule abs_eq_0)
       
   470 apply (rule abs_triangle_ineq)
       
   471 apply (rule abs_mult)
       
   472 apply (rule abs_mult)
       
   473 done
       
   474 
       
   475 lemma norm_zero [simp]: "norm (0::'a::real_normed_vector) = 0"
       
   476 by simp
       
   477 
       
   478 lemma zero_less_norm_iff [simp]:
       
   479   fixes x :: "'a::real_normed_vector"
       
   480   shows "(0 < norm x) = (x \<noteq> 0)"
       
   481 by (simp add: order_less_le)
       
   482 
       
   483 lemma norm_not_less_zero [simp]:
       
   484   fixes x :: "'a::real_normed_vector"
       
   485   shows "\<not> norm x < 0"
       
   486 by (simp add: linorder_not_less)
       
   487 
       
   488 lemma norm_le_zero_iff [simp]:
       
   489   fixes x :: "'a::real_normed_vector"
       
   490   shows "(norm x \<le> 0) = (x = 0)"
       
   491 by (simp add: order_le_less)
       
   492 
       
   493 lemma norm_minus_cancel [simp]:
       
   494   fixes x :: "'a::real_normed_vector"
       
   495   shows "norm (- x) = norm x"
       
   496 proof -
       
   497   have "norm (- x) = norm (scaleR (- 1) x)"
       
   498     by (simp only: scaleR_minus_left scaleR_one)
       
   499   also have "\<dots> = \<bar>- 1\<bar> * norm x"
       
   500     by (rule norm_scaleR)
       
   501   finally show ?thesis by simp
       
   502 qed
       
   503 
       
   504 lemma norm_minus_commute:
       
   505   fixes a b :: "'a::real_normed_vector"
       
   506   shows "norm (a - b) = norm (b - a)"
       
   507 proof -
       
   508   have "norm (- (b - a)) = norm (b - a)"
       
   509     by (rule norm_minus_cancel)
       
   510   thus ?thesis by simp
       
   511 qed
       
   512 
       
   513 lemma norm_triangle_ineq2:
       
   514   fixes a b :: "'a::real_normed_vector"
       
   515   shows "norm a - norm b \<le> norm (a - b)"
       
   516 proof -
       
   517   have "norm (a - b + b) \<le> norm (a - b) + norm b"
       
   518     by (rule norm_triangle_ineq)
       
   519   thus ?thesis by simp
       
   520 qed
       
   521 
       
   522 lemma norm_triangle_ineq3:
       
   523   fixes a b :: "'a::real_normed_vector"
       
   524   shows "\<bar>norm a - norm b\<bar> \<le> norm (a - b)"
       
   525 apply (subst abs_le_iff)
       
   526 apply auto
       
   527 apply (rule norm_triangle_ineq2)
       
   528 apply (subst norm_minus_commute)
       
   529 apply (rule norm_triangle_ineq2)
       
   530 done
       
   531 
       
   532 lemma norm_triangle_ineq4:
       
   533   fixes a b :: "'a::real_normed_vector"
       
   534   shows "norm (a - b) \<le> norm a + norm b"
       
   535 proof -
       
   536   have "norm (a + - b) \<le> norm a + norm (- b)"
       
   537     by (rule norm_triangle_ineq)
       
   538   thus ?thesis
       
   539     by (simp only: diff_minus norm_minus_cancel)
       
   540 qed
       
   541 
       
   542 lemma norm_diff_ineq:
       
   543   fixes a b :: "'a::real_normed_vector"
       
   544   shows "norm a - norm b \<le> norm (a + b)"
       
   545 proof -
       
   546   have "norm a - norm (- b) \<le> norm (a - - b)"
       
   547     by (rule norm_triangle_ineq2)
       
   548   thus ?thesis by simp
       
   549 qed
       
   550 
       
   551 lemma norm_diff_triangle_ineq:
       
   552   fixes a b c d :: "'a::real_normed_vector"
       
   553   shows "norm ((a + b) - (c + d)) \<le> norm (a - c) + norm (b - d)"
       
   554 proof -
       
   555   have "norm ((a + b) - (c + d)) = norm ((a - c) + (b - d))"
       
   556     by (simp add: diff_minus add_ac)
       
   557   also have "\<dots> \<le> norm (a - c) + norm (b - d)"
       
   558     by (rule norm_triangle_ineq)
       
   559   finally show ?thesis .
       
   560 qed
       
   561 
       
   562 lemma abs_norm_cancel [simp]:
       
   563   fixes a :: "'a::real_normed_vector"
       
   564   shows "\<bar>norm a\<bar> = norm a"
       
   565 by (rule abs_of_nonneg [OF norm_ge_zero])
       
   566 
       
   567 lemma norm_add_less:
       
   568   fixes x y :: "'a::real_normed_vector"
       
   569   shows "\<lbrakk>norm x < r; norm y < s\<rbrakk> \<Longrightarrow> norm (x + y) < r + s"
       
   570 by (rule order_le_less_trans [OF norm_triangle_ineq add_strict_mono])
       
   571 
       
   572 lemma norm_mult_less:
       
   573   fixes x y :: "'a::real_normed_algebra"
       
   574   shows "\<lbrakk>norm x < r; norm y < s\<rbrakk> \<Longrightarrow> norm (x * y) < r * s"
       
   575 apply (rule order_le_less_trans [OF norm_mult_ineq])
       
   576 apply (simp add: mult_strict_mono')
       
   577 done
       
   578 
       
   579 lemma norm_of_real [simp]:
       
   580   "norm (of_real r :: 'a::real_normed_algebra_1) = \<bar>r\<bar>"
       
   581 unfolding of_real_def by (simp add: norm_scaleR)
       
   582 
       
   583 lemma norm_number_of [simp]:
       
   584   "norm (number_of w::'a::{number_ring,real_normed_algebra_1})
       
   585     = \<bar>number_of w\<bar>"
       
   586 by (subst of_real_number_of_eq [symmetric], rule norm_of_real)
       
   587 
       
   588 lemma norm_of_int [simp]:
       
   589   "norm (of_int z::'a::real_normed_algebra_1) = \<bar>of_int z\<bar>"
       
   590 by (subst of_real_of_int_eq [symmetric], rule norm_of_real)
       
   591 
       
   592 lemma norm_of_nat [simp]:
       
   593   "norm (of_nat n::'a::real_normed_algebra_1) = of_nat n"
       
   594 apply (subst of_real_of_nat_eq [symmetric])
       
   595 apply (subst norm_of_real, simp)
       
   596 done
       
   597 
       
   598 lemma nonzero_norm_inverse:
       
   599   fixes a :: "'a::real_normed_div_algebra"
       
   600   shows "a \<noteq> 0 \<Longrightarrow> norm (inverse a) = inverse (norm a)"
       
   601 apply (rule inverse_unique [symmetric])
       
   602 apply (simp add: norm_mult [symmetric])
       
   603 done
       
   604 
       
   605 lemma norm_inverse:
       
   606   fixes a :: "'a::{real_normed_div_algebra,division_by_zero}"
       
   607   shows "norm (inverse a) = inverse (norm a)"
       
   608 apply (case_tac "a = 0", simp)
       
   609 apply (erule nonzero_norm_inverse)
       
   610 done
       
   611 
       
   612 lemma nonzero_norm_divide:
       
   613   fixes a b :: "'a::real_normed_field"
       
   614   shows "b \<noteq> 0 \<Longrightarrow> norm (a / b) = norm a / norm b"
       
   615 by (simp add: divide_inverse norm_mult nonzero_norm_inverse)
       
   616 
       
   617 lemma norm_divide:
       
   618   fixes a b :: "'a::{real_normed_field,division_by_zero}"
       
   619   shows "norm (a / b) = norm a / norm b"
       
   620 by (simp add: divide_inverse norm_mult norm_inverse)
       
   621 
       
   622 lemma norm_power_ineq:
       
   623   fixes x :: "'a::{real_normed_algebra_1,recpower}"
       
   624   shows "norm (x ^ n) \<le> norm x ^ n"
       
   625 proof (induct n)
       
   626   case 0 show "norm (x ^ 0) \<le> norm x ^ 0" by simp
       
   627 next
       
   628   case (Suc n)
       
   629   have "norm (x * x ^ n) \<le> norm x * norm (x ^ n)"
       
   630     by (rule norm_mult_ineq)
       
   631   also from Suc have "\<dots> \<le> norm x * norm x ^ n"
       
   632     using norm_ge_zero by (rule mult_left_mono)
       
   633   finally show "norm (x ^ Suc n) \<le> norm x ^ Suc n"
       
   634     by (simp add: power_Suc)
       
   635 qed
       
   636 
       
   637 lemma norm_power:
       
   638   fixes x :: "'a::{real_normed_div_algebra,recpower}"
       
   639   shows "norm (x ^ n) = norm x ^ n"
       
   640 by (induct n) (simp_all add: power_Suc norm_mult)
       
   641 
       
   642 
       
   643 subsection {* Sign function *}
       
   644 
       
   645 lemma norm_sgn:
       
   646   "norm (sgn(x::'a::real_normed_vector)) = (if x = 0 then 0 else 1)"
       
   647 by (simp add: sgn_div_norm norm_scaleR)
       
   648 
       
   649 lemma sgn_zero [simp]: "sgn(0::'a::real_normed_vector) = 0"
       
   650 by (simp add: sgn_div_norm)
       
   651 
       
   652 lemma sgn_zero_iff: "(sgn(x::'a::real_normed_vector) = 0) = (x = 0)"
       
   653 by (simp add: sgn_div_norm)
       
   654 
       
   655 lemma sgn_minus: "sgn (- x) = - sgn(x::'a::real_normed_vector)"
       
   656 by (simp add: sgn_div_norm)
       
   657 
       
   658 lemma sgn_scaleR:
       
   659   "sgn (scaleR r x) = scaleR (sgn r) (sgn(x::'a::real_normed_vector))"
       
   660 by (simp add: sgn_div_norm norm_scaleR mult_ac)
       
   661 
       
   662 lemma sgn_one [simp]: "sgn (1::'a::real_normed_algebra_1) = 1"
       
   663 by (simp add: sgn_div_norm)
       
   664 
       
   665 lemma sgn_of_real:
       
   666   "sgn (of_real r::'a::real_normed_algebra_1) = of_real (sgn r)"
       
   667 unfolding of_real_def by (simp only: sgn_scaleR sgn_one)
       
   668 
       
   669 lemma sgn_mult:
       
   670   fixes x y :: "'a::real_normed_div_algebra"
       
   671   shows "sgn (x * y) = sgn x * sgn y"
       
   672 by (simp add: sgn_div_norm norm_mult mult_commute)
       
   673 
       
   674 lemma real_sgn_eq: "sgn (x::real) = x / \<bar>x\<bar>"
       
   675 by (simp add: sgn_div_norm divide_inverse)
       
   676 
       
   677 lemma real_sgn_pos: "0 < (x::real) \<Longrightarrow> sgn x = 1"
       
   678 unfolding real_sgn_eq by simp
       
   679 
       
   680 lemma real_sgn_neg: "(x::real) < 0 \<Longrightarrow> sgn x = -1"
       
   681 unfolding real_sgn_eq by simp
       
   682 
       
   683 
       
   684 subsection {* Bounded Linear and Bilinear Operators *}
       
   685 
       
   686 locale bounded_linear = additive +
       
   687   constrains f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
       
   688   assumes scaleR: "f (scaleR r x) = scaleR r (f x)"
       
   689   assumes bounded: "\<exists>K. \<forall>x. norm (f x) \<le> norm x * K"
       
   690 begin
       
   691 
       
   692 lemma pos_bounded:
       
   693   "\<exists>K>0. \<forall>x. norm (f x) \<le> norm x * K"
       
   694 proof -
       
   695   obtain K where K: "\<And>x. norm (f x) \<le> norm x * K"
       
   696     using bounded by fast
       
   697   show ?thesis
       
   698   proof (intro exI impI conjI allI)
       
   699     show "0 < max 1 K"
       
   700       by (rule order_less_le_trans [OF zero_less_one le_maxI1])
       
   701   next
       
   702     fix x
       
   703     have "norm (f x) \<le> norm x * K" using K .
       
   704     also have "\<dots> \<le> norm x * max 1 K"
       
   705       by (rule mult_left_mono [OF le_maxI2 norm_ge_zero])
       
   706     finally show "norm (f x) \<le> norm x * max 1 K" .
       
   707   qed
       
   708 qed
       
   709 
       
   710 lemma nonneg_bounded:
       
   711   "\<exists>K\<ge>0. \<forall>x. norm (f x) \<le> norm x * K"
       
   712 proof -
       
   713   from pos_bounded
       
   714   show ?thesis by (auto intro: order_less_imp_le)
       
   715 qed
       
   716 
       
   717 end
       
   718 
       
   719 locale bounded_bilinear =
       
   720   fixes prod :: "['a::real_normed_vector, 'b::real_normed_vector]
       
   721                  \<Rightarrow> 'c::real_normed_vector"
       
   722     (infixl "**" 70)
       
   723   assumes add_left: "prod (a + a') b = prod a b + prod a' b"
       
   724   assumes add_right: "prod a (b + b') = prod a b + prod a b'"
       
   725   assumes scaleR_left: "prod (scaleR r a) b = scaleR r (prod a b)"
       
   726   assumes scaleR_right: "prod a (scaleR r b) = scaleR r (prod a b)"
       
   727   assumes bounded: "\<exists>K. \<forall>a b. norm (prod a b) \<le> norm a * norm b * K"
       
   728 begin
       
   729 
       
   730 lemma pos_bounded:
       
   731   "\<exists>K>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
       
   732 apply (cut_tac bounded, erule exE)
       
   733 apply (rule_tac x="max 1 K" in exI, safe)
       
   734 apply (rule order_less_le_trans [OF zero_less_one le_maxI1])
       
   735 apply (drule spec, drule spec, erule order_trans)
       
   736 apply (rule mult_left_mono [OF le_maxI2])
       
   737 apply (intro mult_nonneg_nonneg norm_ge_zero)
       
   738 done
       
   739 
       
   740 lemma nonneg_bounded:
       
   741   "\<exists>K\<ge>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
       
   742 proof -
       
   743   from pos_bounded
       
   744   show ?thesis by (auto intro: order_less_imp_le)
       
   745 qed
       
   746 
       
   747 lemma additive_right: "additive (\<lambda>b. prod a b)"
       
   748 by (rule additive.intro, rule add_right)
       
   749 
       
   750 lemma additive_left: "additive (\<lambda>a. prod a b)"
       
   751 by (rule additive.intro, rule add_left)
       
   752 
       
   753 lemma zero_left: "prod 0 b = 0"
       
   754 by (rule additive.zero [OF additive_left])
       
   755 
       
   756 lemma zero_right: "prod a 0 = 0"
       
   757 by (rule additive.zero [OF additive_right])
       
   758 
       
   759 lemma minus_left: "prod (- a) b = - prod a b"
       
   760 by (rule additive.minus [OF additive_left])
       
   761 
       
   762 lemma minus_right: "prod a (- b) = - prod a b"
       
   763 by (rule additive.minus [OF additive_right])
       
   764 
       
   765 lemma diff_left:
       
   766   "prod (a - a') b = prod a b - prod a' b"
       
   767 by (rule additive.diff [OF additive_left])
       
   768 
       
   769 lemma diff_right:
       
   770   "prod a (b - b') = prod a b - prod a b'"
       
   771 by (rule additive.diff [OF additive_right])
       
   772 
       
   773 lemma bounded_linear_left:
       
   774   "bounded_linear (\<lambda>a. a ** b)"
       
   775 apply (unfold_locales)
       
   776 apply (rule add_left)
       
   777 apply (rule scaleR_left)
       
   778 apply (cut_tac bounded, safe)
       
   779 apply (rule_tac x="norm b * K" in exI)
       
   780 apply (simp add: mult_ac)
       
   781 done
       
   782 
       
   783 lemma bounded_linear_right:
       
   784   "bounded_linear (\<lambda>b. a ** b)"
       
   785 apply (unfold_locales)
       
   786 apply (rule add_right)
       
   787 apply (rule scaleR_right)
       
   788 apply (cut_tac bounded, safe)
       
   789 apply (rule_tac x="norm a * K" in exI)
       
   790 apply (simp add: mult_ac)
       
   791 done
       
   792 
       
   793 lemma prod_diff_prod:
       
   794   "(x ** y - a ** b) = (x - a) ** (y - b) + (x - a) ** b + a ** (y - b)"
       
   795 by (simp add: diff_left diff_right)
       
   796 
       
   797 end
       
   798 
       
   799 interpretation mult:
       
   800   bounded_bilinear ["op * :: 'a \<Rightarrow> 'a \<Rightarrow> 'a::real_normed_algebra"]
       
   801 apply (rule bounded_bilinear.intro)
       
   802 apply (rule left_distrib)
       
   803 apply (rule right_distrib)
       
   804 apply (rule mult_scaleR_left)
       
   805 apply (rule mult_scaleR_right)
       
   806 apply (rule_tac x="1" in exI)
       
   807 apply (simp add: norm_mult_ineq)
       
   808 done
       
   809 
       
   810 interpretation mult_left:
       
   811   bounded_linear ["(\<lambda>x::'a::real_normed_algebra. x * y)"]
       
   812 by (rule mult.bounded_linear_left)
       
   813 
       
   814 interpretation mult_right:
       
   815   bounded_linear ["(\<lambda>y::'a::real_normed_algebra. x * y)"]
       
   816 by (rule mult.bounded_linear_right)
       
   817 
       
   818 interpretation divide:
       
   819   bounded_linear ["(\<lambda>x::'a::real_normed_field. x / y)"]
       
   820 unfolding divide_inverse by (rule mult.bounded_linear_left)
       
   821 
       
   822 interpretation scaleR: bounded_bilinear ["scaleR"]
       
   823 apply (rule bounded_bilinear.intro)
       
   824 apply (rule scaleR_left_distrib)
       
   825 apply (rule scaleR_right_distrib)
       
   826 apply simp
       
   827 apply (rule scaleR_left_commute)
       
   828 apply (rule_tac x="1" in exI)
       
   829 apply (simp add: norm_scaleR)
       
   830 done
       
   831 
       
   832 interpretation scaleR_left: bounded_linear ["\<lambda>r. scaleR r x"]
       
   833 by (rule scaleR.bounded_linear_left)
       
   834 
       
   835 interpretation scaleR_right: bounded_linear ["\<lambda>x. scaleR r x"]
       
   836 by (rule scaleR.bounded_linear_right)
       
   837 
       
   838 interpretation of_real: bounded_linear ["\<lambda>r. of_real r"]
       
   839 unfolding of_real_def by (rule scaleR.bounded_linear_left)
       
   840 
       
   841 end