src/HOL/Analysis/Inner_Product.thy
author paulson <lp15@cam.ac.uk>
Sat, 12 Apr 2025 22:42:32 +0100
changeset 82489 d35d355f7330
parent 81116 0fb1e2dd4122
child 82522 62afd98e3f3e
permissions -rw-r--r--
Tidied more messy old proofs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63971
da89140186e2 HOL-Analysis: move Product_Vector and Inner_Product from Library
hoelzl
parents: 63886
diff changeset
     1
(*  Title:      HOL/Analysis/Inner_Product.thy
41959
b460124855b8 tuned headers;
wenzelm
parents: 39302
diff changeset
     2
    Author:     Brian Huffman
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
     3
*)
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
     4
69600
86e8e7347ac0 typed definitions
nipkow
parents: 69597
diff changeset
     5
section \<open>Inner Product Spaces and Gradient Derivative\<close>
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
     6
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
     7
theory Inner_Product
65513
587433a18053 tuned imports;
wenzelm
parents: 65064
diff changeset
     8
imports Complex_Main
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
     9
begin
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    10
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    11
subsection \<open>Real inner product spaces\<close>
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    12
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    13
text \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69513
diff changeset
    14
  Temporarily relax type constraints for \<^term>\<open>open\<close>, \<^term>\<open>uniformity\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69513
diff changeset
    15
  \<^term>\<open>dist\<close>, and \<^term>\<open>norm\<close>.
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    16
\<close>
31492
5400beeddb55 replace 'topo' with 'open'; add extra type constraint for 'open'
huffman
parents: 31446
diff changeset
    17
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    18
setup \<open>Sign.add_const_constraint
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69513
diff changeset
    19
  (\<^const_name>\<open>open\<close>, SOME \<^typ>\<open>'a::open set \<Rightarrow> bool\<close>)\<close>
31446
2d91b2416de8 add extra type constraints for dist, norm
huffman
parents: 31417
diff changeset
    20
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    21
setup \<open>Sign.add_const_constraint
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69513
diff changeset
    22
  (\<^const_name>\<open>dist\<close>, SOME \<^typ>\<open>'a::dist \<Rightarrow> 'a \<Rightarrow> real\<close>)\<close>
31446
2d91b2416de8 add extra type constraints for dist, norm
huffman
parents: 31417
diff changeset
    23
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    24
setup \<open>Sign.add_const_constraint
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69513
diff changeset
    25
  (\<^const_name>\<open>uniformity\<close>, SOME \<^typ>\<open>('a::uniformity \<times> 'a) filter\<close>)\<close>
62101
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61915
diff changeset
    26
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61915
diff changeset
    27
setup \<open>Sign.add_const_constraint
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69513
diff changeset
    28
  (\<^const_name>\<open>norm\<close>, SOME \<^typ>\<open>'a::norm \<Rightarrow> real\<close>)\<close>
31446
2d91b2416de8 add extra type constraints for dist, norm
huffman
parents: 31417
diff changeset
    29
68623
b942da0962c2 correct import
nipkow
parents: 68617
diff changeset
    30
class real_inner = real_vector + sgn_div_norm + dist_norm + uniformity_dist + open_uniformity +
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    31
  fixes inner :: "'a \<Rightarrow> 'a \<Rightarrow> real"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    32
  assumes inner_commute: "inner x y = inner y x"
31590
776d6a4c1327 declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents: 31492
diff changeset
    33
  and inner_add_left: "inner (x + y) z = inner x z + inner y z"
776d6a4c1327 declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents: 31492
diff changeset
    34
  and inner_scaleR_left [simp]: "inner (scaleR r x) y = r * (inner x y)"
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    35
  and inner_ge_zero [simp]: "0 \<le> inner x x"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    36
  and inner_eq_zero_iff [simp]: "inner x x = 0 \<longleftrightarrow> x = 0"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    37
  and norm_eq_sqrt_inner: "norm x = sqrt (inner x x)"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    38
begin
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    39
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    40
lemma inner_zero_left [simp]: "inner 0 x = 0"
31590
776d6a4c1327 declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents: 31492
diff changeset
    41
  using inner_add_left [of 0 0 x] by simp
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    42
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    43
lemma inner_minus_left [simp]: "inner (- x) y = - inner x y"
31590
776d6a4c1327 declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents: 31492
diff changeset
    44
  using inner_add_left [of x "- x" y] by simp
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    45
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    46
lemma inner_diff_left: "inner (x - y) z = inner x z - inner y z"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53938
diff changeset
    47
  using inner_add_left [of x "- y" z] by simp
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    48
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63971
diff changeset
    49
lemma inner_sum_left: "inner (\<Sum>x\<in>A. f x) y = (\<Sum>x\<in>A. inner (f x) y)"
44282
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44233
diff changeset
    50
  by (cases "finite A", induct set: finite, simp_all add: inner_add_left)
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44233
diff changeset
    51
67986
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67962
diff changeset
    52
lemma all_zero_iff [simp]: "(\<forall>u. inner x u = 0) \<longleftrightarrow> (x = 0)"
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67962
diff changeset
    53
  by auto (use inner_eq_zero_iff in blast)
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67962
diff changeset
    54
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    55
text \<open>Transfer distributivity rules to right argument.\<close>
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    56
31590
776d6a4c1327 declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents: 31492
diff changeset
    57
lemma inner_add_right: "inner x (y + z) = inner x y + inner x z"
776d6a4c1327 declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents: 31492
diff changeset
    58
  using inner_add_left [of y z x] by (simp only: inner_commute)
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    59
31590
776d6a4c1327 declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents: 31492
diff changeset
    60
lemma inner_scaleR_right [simp]: "inner x (scaleR r y) = r * (inner x y)"
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    61
  using inner_scaleR_left [of r y x] by (simp only: inner_commute)
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    62
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    63
lemma inner_zero_right [simp]: "inner x 0 = 0"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    64
  using inner_zero_left [of x] by (simp only: inner_commute)
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    65
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    66
lemma inner_minus_right [simp]: "inner x (- y) = - inner x y"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    67
  using inner_minus_left [of y x] by (simp only: inner_commute)
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    68
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    69
lemma inner_diff_right: "inner x (y - z) = inner x y - inner x z"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    70
  using inner_diff_left [of y z x] by (simp only: inner_commute)
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    71
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63971
diff changeset
    72
lemma inner_sum_right: "inner x (\<Sum>y\<in>A. f y) = (\<Sum>y\<in>A. inner x (f y))"
b9a1486e79be setsum -> sum
nipkow
parents: 63971
diff changeset
    73
  using inner_sum_left [of f A x] by (simp only: inner_commute)
44282
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44233
diff changeset
    74
31590
776d6a4c1327 declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents: 31492
diff changeset
    75
lemmas inner_add [algebra_simps] = inner_add_left inner_add_right
776d6a4c1327 declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents: 31492
diff changeset
    76
lemmas inner_diff [algebra_simps]  = inner_diff_left inner_diff_right
776d6a4c1327 declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents: 31492
diff changeset
    77
lemmas inner_scaleR = inner_scaleR_left inner_scaleR_right
776d6a4c1327 declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents: 31492
diff changeset
    78
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    79
text \<open>Legacy theorem names\<close>
31590
776d6a4c1327 declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents: 31492
diff changeset
    80
lemmas inner_left_distrib = inner_add_left
776d6a4c1327 declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents: 31492
diff changeset
    81
lemmas inner_right_distrib = inner_add_right
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    82
lemmas inner_distrib = inner_left_distrib inner_right_distrib
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    83
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    84
lemma inner_gt_zero_iff [simp]: "0 < inner x x \<longleftrightarrow> x \<noteq> 0"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    85
  by (simp add: order_less_le)
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    86
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51642
diff changeset
    87
lemma power2_norm_eq_inner: "(norm x)\<^sup>2 = inner x x"
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    88
  by (simp add: norm_eq_sqrt_inner)
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
    89
82489
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
    90
lemma dot_square_norm: "inner x x = (norm x)\<^sup>2"
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
    91
  by (metis power2_norm_eq_inner) 
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
    92
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
    93
text \<open>Identities involving real multiplication and division.\<close>
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
    94
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
    95
lemma inner_mult_left: "inner (of_real m * a) b = m * (inner a b)"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
    96
  by (metis real_inner_class.inner_scaleR_left scaleR_conv_of_real)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
    97
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
    98
lemma inner_mult_right: "inner a (of_real m * b) = m * (inner a b)"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
    99
  by (metis real_inner_class.inner_scaleR_right scaleR_conv_of_real)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
   100
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
   101
lemma inner_mult_left': "inner (a * of_real m) b = m * (inner a b)"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
   102
  by (simp add: of_real_def)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
   103
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
   104
lemma inner_mult_right': "inner a (b * of_real m) = (inner a b) * m"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
   105
  by (simp add: of_real_def real_inner_class.inner_scaleR_right)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
   106
30046
49f603f92c47 fix spelling
huffman
parents: 29993
diff changeset
   107
lemma Cauchy_Schwarz_ineq:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51642
diff changeset
   108
  "(inner x y)\<^sup>2 \<le> inner x x * inner y y"
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   109
proof (cases)
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   110
  assume "y = 0"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   111
  thus ?thesis by simp
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   112
next
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   113
  assume y: "y \<noteq> 0"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   114
  let ?r = "inner x y / inner y y"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   115
  have "0 \<le> inner (x - scaleR ?r y) (x - scaleR ?r y)"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   116
    by (rule inner_ge_zero)
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   117
  also have "\<dots> = inner x x - inner y x * ?r"
31590
776d6a4c1327 declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents: 31492
diff changeset
   118
    by (simp add: inner_diff)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51642
diff changeset
   119
  also have "\<dots> = inner x x - (inner x y)\<^sup>2 / inner y y"
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   120
    by (simp add: power2_eq_square inner_commute)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51642
diff changeset
   121
  finally have "0 \<le> inner x x - (inner x y)\<^sup>2 / inner y y" .
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51642
diff changeset
   122
  hence "(inner x y)\<^sup>2 / inner y y \<le> inner x x"
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   123
    by (simp add: le_diff_eq)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51642
diff changeset
   124
  thus "(inner x y)\<^sup>2 \<le> inner x x * inner y y"
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   125
    by (simp add: pos_divide_le_eq y)
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   126
qed
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   127
30046
49f603f92c47 fix spelling
huffman
parents: 29993
diff changeset
   128
lemma Cauchy_Schwarz_ineq2:
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   129
  "\<bar>inner x y\<bar> \<le> norm x * norm y"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   130
proof (rule power2_le_imp_le)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51642
diff changeset
   131
  have "(inner x y)\<^sup>2 \<le> inner x x * inner y y"
30046
49f603f92c47 fix spelling
huffman
parents: 29993
diff changeset
   132
    using Cauchy_Schwarz_ineq .
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51642
diff changeset
   133
  thus "\<bar>inner x y\<bar>\<^sup>2 \<le> (norm x * norm y)\<^sup>2"
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   134
    by (simp add: power_mult_distrib power2_norm_eq_inner)
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   135
  show "0 \<le> norm x * norm y"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   136
    unfolding norm_eq_sqrt_inner
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   137
    by (intro mult_nonneg_nonneg real_sqrt_ge_zero inner_ge_zero)
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   138
qed
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   139
53938
eb93cca4589a moved lemma
huffman
parents: 53015
diff changeset
   140
lemma norm_cauchy_schwarz: "inner x y \<le> norm x * norm y"
eb93cca4589a moved lemma
huffman
parents: 53015
diff changeset
   141
  using Cauchy_Schwarz_ineq2 [of x y] by auto
eb93cca4589a moved lemma
huffman
parents: 53015
diff changeset
   142
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   143
subclass real_normed_vector
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   144
proof
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   145
  fix a :: real and x y :: 'a
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   146
  show "norm x = 0 \<longleftrightarrow> x = 0"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   147
    unfolding norm_eq_sqrt_inner by simp
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   148
  show "norm (x + y) \<le> norm x + norm y"
82489
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
   149
  proof (rule power2_le_imp_le)
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
   150
    have "inner x y \<le> norm x * norm y"
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
   151
      by (rule norm_cauchy_schwarz)
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
   152
    thus "(norm (x + y))\<^sup>2 \<le> (norm x + norm y)\<^sup>2"
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
   153
      unfolding power2_sum power2_norm_eq_inner
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
   154
      by (simp add: inner_add inner_commute)
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
   155
    show "0 \<le> norm x + norm y"
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
   156
      unfolding norm_eq_sqrt_inner by simp
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
   157
  qed
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51642
diff changeset
   158
  have "sqrt (a\<^sup>2 * inner x x) = \<bar>a\<bar> * sqrt (inner x x)"
68611
4bc4b5c0ccfc de-applying, etc.
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   159
    by (simp add: real_sqrt_mult)
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   160
  then show "norm (a *\<^sub>R x) = \<bar>a\<bar> * norm x"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   161
    unfolding norm_eq_sqrt_inner
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56381
diff changeset
   162
    by (simp add: power2_eq_square mult.assoc)
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   163
qed
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   164
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   165
end
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   166
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   167
lemma square_bound_lemma:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   168
  fixes x :: real
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   169
  shows "x < (1 + x) * (1 + x)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   170
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   171
  have "(x + 1/2)\<^sup>2 + 3/4 > 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   172
    using zero_le_power2[of "x+1/2"] by arith
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   173
  then show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   174
    by (simp add: field_simps power2_eq_square)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   175
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   176
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   177
lemma square_continuous:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   178
  fixes e :: real
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   179
  shows "e > 0 \<Longrightarrow> \<exists>d. 0 < d \<and> (\<forall>y. \<bar>y - x\<bar> < d \<longrightarrow> \<bar>y * y - x * x\<bar> < e)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   180
  using isCont_power[OF continuous_ident, of x, unfolded isCont_def LIM_eq, rule_format, of e 2]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   181
  by (force simp add: power2_eq_square)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   182
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   183
lemma norm_le: "norm x \<le> norm y \<longleftrightarrow> inner x x \<le> inner y y"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   184
  by (simp add: norm_eq_sqrt_inner)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   185
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   186
lemma norm_lt: "norm x < norm y \<longleftrightarrow> inner x x < inner y y"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   187
  by (simp add: norm_eq_sqrt_inner)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   188
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   189
lemma norm_eq: "norm x = norm y \<longleftrightarrow> inner x x = inner y y"
82489
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
   190
  by (simp add: norm_eq_sqrt_inner)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   191
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   192
lemma norm_eq_1: "norm x = 1 \<longleftrightarrow> inner x x = 1"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   193
  by (simp add: norm_eq_sqrt_inner)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67986
diff changeset
   194
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
   195
lemma inner_divide_left:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
   196
  fixes a :: "'a :: {real_inner,real_div_algebra}"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
   197
  shows "inner (a / of_real m) b = (inner a b) / m"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
   198
  by (metis (no_types) divide_inverse inner_commute inner_scaleR_right mult.left_neutral mult.right_neutral mult_scaleR_right of_real_inverse scaleR_conv_of_real times_divide_eq_left)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
   199
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
   200
lemma inner_divide_right:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
   201
  fixes a :: "'a :: {real_inner,real_div_algebra}"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
   202
  shows "inner a (b / of_real m) = (inner a b) / m"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
   203
  by (metis inner_commute inner_divide_left)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 60679
diff changeset
   204
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   205
text \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69513
diff changeset
   206
  Re-enable constraints for \<^term>\<open>open\<close>, \<^term>\<open>uniformity\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69513
diff changeset
   207
  \<^term>\<open>dist\<close>, and \<^term>\<open>norm\<close>.
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   208
\<close>
31492
5400beeddb55 replace 'topo' with 'open'; add extra type constraint for 'open'
huffman
parents: 31446
diff changeset
   209
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   210
setup \<open>Sign.add_const_constraint
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69513
diff changeset
   211
  (\<^const_name>\<open>open\<close>, SOME \<^typ>\<open>'a::topological_space set \<Rightarrow> bool\<close>)\<close>
31446
2d91b2416de8 add extra type constraints for dist, norm
huffman
parents: 31417
diff changeset
   212
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   213
setup \<open>Sign.add_const_constraint
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69513
diff changeset
   214
  (\<^const_name>\<open>uniformity\<close>, SOME \<^typ>\<open>('a::uniform_space \<times> 'a) filter\<close>)\<close>
62101
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61915
diff changeset
   215
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61915
diff changeset
   216
setup \<open>Sign.add_const_constraint
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69513
diff changeset
   217
  (\<^const_name>\<open>dist\<close>, SOME \<^typ>\<open>'a::metric_space \<Rightarrow> 'a \<Rightarrow> real\<close>)\<close>
31446
2d91b2416de8 add extra type constraints for dist, norm
huffman
parents: 31417
diff changeset
   218
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   219
setup \<open>Sign.add_const_constraint
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69513
diff changeset
   220
  (\<^const_name>\<open>norm\<close>, SOME \<^typ>\<open>'a::real_normed_vector \<Rightarrow> real\<close>)\<close>
31446
2d91b2416de8 add extra type constraints for dist, norm
huffman
parents: 31417
diff changeset
   221
44282
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44233
diff changeset
   222
lemma bounded_bilinear_inner:
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44233
diff changeset
   223
  "bounded_bilinear (inner::'a::real_inner \<Rightarrow> 'a \<Rightarrow> real)"
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   224
proof
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   225
  fix x y z :: 'a and r :: real
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   226
  show "inner (x + y) z = inner x z + inner y z"
31590
776d6a4c1327 declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents: 31492
diff changeset
   227
    by (rule inner_add_left)
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   228
  show "inner x (y + z) = inner x y + inner x z"
31590
776d6a4c1327 declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents: 31492
diff changeset
   229
    by (rule inner_add_right)
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   230
  show "inner (scaleR r x) y = scaleR r (inner x y)"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   231
    unfolding real_scaleR_def by (rule inner_scaleR_left)
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   232
  show "inner x (scaleR r y) = scaleR r (inner x y)"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   233
    unfolding real_scaleR_def by (rule inner_scaleR_right)
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   234
  show "\<exists>K. \<forall>x y::'a. norm (inner x y) \<le> norm x * norm y * K"
82489
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
   235
    by (metis Cauchy_Schwarz_ineq2 mult.commute mult_1 real_norm_def)
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   236
qed
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   237
44282
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44233
diff changeset
   238
lemmas tendsto_inner [tendsto_intros] =
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44233
diff changeset
   239
  bounded_bilinear.tendsto [OF bounded_bilinear_inner]
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44233
diff changeset
   240
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44233
diff changeset
   241
lemmas isCont_inner [simp] =
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44233
diff changeset
   242
  bounded_bilinear.isCont [OF bounded_bilinear_inner]
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   243
56381
0556204bc230 merged DERIV_intros, has_derivative_intros into derivative_intros
hoelzl
parents: 56181
diff changeset
   244
lemmas has_derivative_inner [derivative_intros] =
44282
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44233
diff changeset
   245
  bounded_bilinear.FDERIV [OF bounded_bilinear_inner]
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   246
44282
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44233
diff changeset
   247
lemmas bounded_linear_inner_left =
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44233
diff changeset
   248
  bounded_bilinear.bounded_linear_left [OF bounded_bilinear_inner]
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44233
diff changeset
   249
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44233
diff changeset
   250
lemmas bounded_linear_inner_right =
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44233
diff changeset
   251
  bounded_bilinear.bounded_linear_right [OF bounded_bilinear_inner]
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44126
diff changeset
   252
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61518
diff changeset
   253
lemmas bounded_linear_inner_left_comp = bounded_linear_inner_left[THEN bounded_linear_compose]
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61518
diff changeset
   254
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61518
diff changeset
   255
lemmas bounded_linear_inner_right_comp = bounded_linear_inner_right[THEN bounded_linear_compose]
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61518
diff changeset
   256
56381
0556204bc230 merged DERIV_intros, has_derivative_intros into derivative_intros
hoelzl
parents: 56181
diff changeset
   257
lemmas has_derivative_inner_right [derivative_intros] =
56181
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 54230
diff changeset
   258
  bounded_linear.has_derivative [OF bounded_linear_inner_right]
51642
400ec5ae7f8f move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
hoelzl
parents: 51002
diff changeset
   259
56381
0556204bc230 merged DERIV_intros, has_derivative_intros into derivative_intros
hoelzl
parents: 56181
diff changeset
   260
lemmas has_derivative_inner_left [derivative_intros] =
56181
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 54230
diff changeset
   261
  bounded_linear.has_derivative [OF bounded_linear_inner_left]
51642
400ec5ae7f8f move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
hoelzl
parents: 51002
diff changeset
   262
400ec5ae7f8f move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
hoelzl
parents: 51002
diff changeset
   263
lemma differentiable_inner [simp]:
56181
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 54230
diff changeset
   264
  "f differentiable (at x within s) \<Longrightarrow> g differentiable at x within s \<Longrightarrow> (\<lambda>x. inner (f x) (g x)) differentiable at x within s"
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 54230
diff changeset
   265
  unfolding differentiable_def by (blast intro: has_derivative_inner)
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   266
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   267
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   268
subsection \<open>Class instances\<close>
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   269
68617
75129a73aca3 more economic tagging
nipkow
parents: 68611
diff changeset
   270
instantiation real :: real_inner
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   271
begin
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   272
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68623
diff changeset
   273
definition inner_real_def [simp]: "inner = (*)"
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   274
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   275
instance
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   276
proof
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   277
  fix x y z r :: real
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   278
  show "inner x y = inner y x"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56381
diff changeset
   279
    unfolding inner_real_def by (rule mult.commute)
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   280
  show "inner (x + y) z = inner x z + inner y z"
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 44902
diff changeset
   281
    unfolding inner_real_def by (rule distrib_right)
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   282
  show "inner (scaleR r x) y = r * inner x y"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56381
diff changeset
   283
    unfolding inner_real_def real_scaleR_def by (rule mult.assoc)
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   284
  show "0 \<le> inner x x"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   285
    unfolding inner_real_def by simp
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   286
  show "inner x x = 0 \<longleftrightarrow> x = 0"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   287
    unfolding inner_real_def by simp
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   288
  show "norm x = sqrt (inner x x)"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   289
    unfolding inner_real_def by simp
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   290
qed
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   291
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   292
end
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   293
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63589
diff changeset
   294
lemma
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63589
diff changeset
   295
  shows real_inner_1_left[simp]: "inner 1 x = x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63589
diff changeset
   296
    and real_inner_1_right[simp]: "inner x 1 = x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63589
diff changeset
   297
  by simp_all
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63589
diff changeset
   298
68617
75129a73aca3 more economic tagging
nipkow
parents: 68611
diff changeset
   299
instantiation complex :: real_inner
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   300
begin
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   301
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   302
definition inner_complex_def:
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   303
  "inner x y = Re x * Re y + Im x * Im y"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   304
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   305
instance
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   306
proof
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   307
  fix x y z :: complex and r :: real
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   308
  show "inner x y = inner y x"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56381
diff changeset
   309
    unfolding inner_complex_def by (simp add: mult.commute)
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   310
  show "inner (x + y) z = inner x z + inner y z"
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 44902
diff changeset
   311
    unfolding inner_complex_def by (simp add: distrib_right)
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   312
  show "inner (scaleR r x) y = r * inner x y"
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 44902
diff changeset
   313
    unfolding inner_complex_def by (simp add: distrib_left)
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   314
  show "0 \<le> inner x x"
44126
ce44e70d0c47 avoid duplicate rewrite warnings
huffman
parents: 41959
diff changeset
   315
    unfolding inner_complex_def by simp
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   316
  show "inner x x = 0 \<longleftrightarrow> x = 0"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   317
    unfolding inner_complex_def
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68073
diff changeset
   318
    by (simp add: add_nonneg_eq_0_iff complex_eq_iff)
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   319
  show "norm x = sqrt (inner x x)"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68073
diff changeset
   320
    unfolding inner_complex_def norm_complex_def
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   321
    by (simp add: power2_eq_square)
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   322
qed
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   323
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   324
end
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   325
44902
9ba11d41cd1f move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
huffman
parents: 44282
diff changeset
   326
lemma complex_inner_1 [simp]: "inner 1 x = Re x"
9ba11d41cd1f move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
huffman
parents: 44282
diff changeset
   327
  unfolding inner_complex_def by simp
9ba11d41cd1f move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
huffman
parents: 44282
diff changeset
   328
9ba11d41cd1f move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
huffman
parents: 44282
diff changeset
   329
lemma complex_inner_1_right [simp]: "inner x 1 = Re x"
9ba11d41cd1f move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
huffman
parents: 44282
diff changeset
   330
  unfolding inner_complex_def by simp
9ba11d41cd1f move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
huffman
parents: 44282
diff changeset
   331
65064
a4abec71279a Renamed ii to imaginary_unit in order to free up ii as a variable name. Also replaced some legacy def commands
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   332
lemma complex_inner_i_left [simp]: "inner \<i> x = Im x"
44902
9ba11d41cd1f move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
huffman
parents: 44282
diff changeset
   333
  unfolding inner_complex_def by simp
9ba11d41cd1f move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
huffman
parents: 44282
diff changeset
   334
65064
a4abec71279a Renamed ii to imaginary_unit in order to free up ii as a variable name. Also replaced some legacy def commands
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   335
lemma complex_inner_i_right [simp]: "inner x \<i> = Im x"
44902
9ba11d41cd1f move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
huffman
parents: 44282
diff changeset
   336
  unfolding inner_complex_def by simp
9ba11d41cd1f move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
huffman
parents: 44282
diff changeset
   337
66486
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   338
lemma dot_square_norm: "inner x x = (norm x)\<^sup>2"
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   339
  by (simp only: power2_norm_eq_inner) (* TODO: move? *)
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   340
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   341
lemma norm_eq_square: "norm x = a \<longleftrightarrow> 0 \<le> a \<and> inner x x = a\<^sup>2"
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   342
  by (auto simp add: norm_eq_sqrt_inner)
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   343
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   344
lemma norm_le_square: "norm x \<le> a \<longleftrightarrow> 0 \<le> a \<and> inner x x \<le> a\<^sup>2"
82489
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
   345
  by (metis norm_eq_sqrt_inner norm_ge_zero order_trans real_le_lsqrt sqrt_le_D)
66486
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   346
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   347
lemma norm_ge_square: "norm x \<ge> a \<longleftrightarrow> a \<le> 0 \<or> inner x x \<ge> a\<^sup>2"
82489
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
   348
  by (metis nle_le norm_eq_square norm_le_square)
66486
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   349
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   350
lemma norm_lt_square: "norm x < a \<longleftrightarrow> 0 < a \<and> inner x x < a\<^sup>2"
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   351
  by (metis not_le norm_ge_square)
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   352
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   353
lemma norm_gt_square: "norm x > a \<longleftrightarrow> a < 0 \<or> inner x x > a\<^sup>2"
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   354
  by (metis norm_le_square not_less)
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   355
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   356
text\<open>Dot product in terms of the norm rather than conversely.\<close>
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   357
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   358
lemmas inner_simps = inner_add_left inner_add_right inner_diff_right inner_diff_left
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   359
  inner_scaleR_left inner_scaleR_right
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   360
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   361
lemma dot_norm: "inner x y = ((norm (x + y))\<^sup>2 - (norm x)\<^sup>2 - (norm y)\<^sup>2) / 2"
82489
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
   362
  by (auto simp: power2_norm_eq_inner inner_simps inner_commute) 
66486
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   363
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   364
lemma dot_norm_neg: "inner x y = (((norm x)\<^sup>2 + (norm y)\<^sup>2) - (norm (x - y))\<^sup>2) / 2"
82489
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
   365
  by (auto simp: power2_norm_eq_inner inner_simps inner_commute)
66486
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   366
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   367
lemma of_real_inner_1 [simp]: 
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   368
  "inner (of_real x) (1 :: 'a :: {real_inner, real_normed_algebra_1}) = x"
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   369
  by (simp add: of_real_def dot_square_norm)
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   370
  
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   371
lemma summable_of_real_iff: 
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   372
  "summable (\<lambda>x. of_real (f x) :: 'a :: {real_normed_algebra_1,real_inner}) \<longleftrightarrow> summable f"
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   373
proof
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   374
  assume *: "summable (\<lambda>x. of_real (f x) :: 'a)"
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   375
  interpret bounded_linear "\<lambda>x::'a. inner x 1"
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   376
    by (rule bounded_linear_inner_left)
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   377
  from summable [OF *] show "summable f" by simp
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   378
qed (auto intro: summable_of_real)
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   379
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 65513
diff changeset
   380
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   381
subsection \<open>Gradient derivative\<close>
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   382
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69674
diff changeset
   383
definition\<^marker>\<open>tag important\<close>
81097
6c81cdca5b82 more inner syntax markup: HOL-Analysis;
wenzelm
parents: 80914
diff changeset
   384
  gderiv :: "['a::real_inner \<Rightarrow> real, 'a, 'a] \<Rightarrow> bool"
6c81cdca5b82 more inner syntax markup: HOL-Analysis;
wenzelm
parents: 80914
diff changeset
   385
    (\<open>(\<open>notation=\<open>mixfix GDERIV\<close>\<close>GDERIV (_)/ (_)/ :> (_))\<close> [1000, 1000, 60] 60)
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   386
where
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   387
  "GDERIV f x :> D \<longleftrightarrow> FDERIV f x :> (\<lambda>h. inner h D)"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   388
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   389
lemma gderiv_deriv [simp]: "GDERIV f x :> D \<longleftrightarrow> DERIV f x :> D"
56181
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 54230
diff changeset
   390
  by (simp only: gderiv_def has_field_derivative_def inner_real_def mult_commute_abs)
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   391
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   392
lemma GDERIV_DERIV_compose:
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   393
    "\<lbrakk>GDERIV f x :> df; DERIV g (f x) :> dg\<rbrakk>
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   394
     \<Longrightarrow> GDERIV (\<lambda>x. g (f x)) x :> scaleR dg df"
56181
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 54230
diff changeset
   395
  unfolding gderiv_def has_field_derivative_def
82489
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
   396
  using has_derivative_compose by fastforce
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   397
56181
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 54230
diff changeset
   398
lemma has_derivative_subst: "\<lbrakk>FDERIV f x :> df; df = d\<rbrakk> \<Longrightarrow> FDERIV f x :> d"
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   399
  by simp
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   400
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   401
lemma GDERIV_subst: "\<lbrakk>GDERIV f x :> df; df = d\<rbrakk> \<Longrightarrow> GDERIV f x :> d"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   402
  by simp
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   403
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   404
lemma GDERIV_const: "GDERIV (\<lambda>x. k) x :> 0"
56181
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 54230
diff changeset
   405
  unfolding gderiv_def inner_zero_right by (rule has_derivative_const)
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   406
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   407
lemma GDERIV_add:
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   408
    "\<lbrakk>GDERIV f x :> df; GDERIV g x :> dg\<rbrakk>
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   409
     \<Longrightarrow> GDERIV (\<lambda>x. f x + g x) x :> df + dg"
56181
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 54230
diff changeset
   410
  unfolding gderiv_def inner_add_right by (rule has_derivative_add)
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   411
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   412
lemma GDERIV_minus:
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   413
    "GDERIV f x :> df \<Longrightarrow> GDERIV (\<lambda>x. - f x) x :> - df"
56181
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 54230
diff changeset
   414
  unfolding gderiv_def inner_minus_right by (rule has_derivative_minus)
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   415
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   416
lemma GDERIV_diff:
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   417
    "\<lbrakk>GDERIV f x :> df; GDERIV g x :> dg\<rbrakk>
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   418
     \<Longrightarrow> GDERIV (\<lambda>x. f x - g x) x :> df - dg"
56181
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 54230
diff changeset
   419
  unfolding gderiv_def inner_diff_right by (rule has_derivative_diff)
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   420
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   421
lemma GDERIV_scaleR:
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   422
    "\<lbrakk>DERIV f x :> df; GDERIV g x :> dg\<rbrakk>
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   423
     \<Longrightarrow> GDERIV (\<lambda>x. scaleR (f x) (g x)) x
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   424
      :> (scaleR (f x) dg + scaleR df (g x))"
82489
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 81116
diff changeset
   425
  by (simp add: DERIV_mult')
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   426
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   427
lemma GDERIV_mult:
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   428
    "\<lbrakk>GDERIV f x :> df; GDERIV g x :> dg\<rbrakk>
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   429
     \<Longrightarrow> GDERIV (\<lambda>x. f x * g x) x :> scaleR (f x) dg + scaleR (g x) df"
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   430
  unfolding gderiv_def
56181
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 54230
diff changeset
   431
  apply (rule has_derivative_subst)
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 54230
diff changeset
   432
  apply (erule (1) has_derivative_mult)
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
   433
  apply (simp add: inner_add ac_simps)
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   434
  done
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   435
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   436
lemma GDERIV_inverse:
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   437
    "\<lbrakk>GDERIV f x :> df; f x \<noteq> 0\<rbrakk>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51642
diff changeset
   438
     \<Longrightarrow> GDERIV (\<lambda>x. inverse (f x)) x :> - (inverse (f x))\<^sup>2 *\<^sub>R df"
67986
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67962
diff changeset
   439
  by (metis DERIV_inverse GDERIV_DERIV_compose numerals(2))
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67962
diff changeset
   440
  
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   441
lemma GDERIV_norm:
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   442
  assumes "x \<noteq> 0" shows "GDERIV (\<lambda>x. norm x) x :> sgn x"
67986
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67962
diff changeset
   443
    unfolding gderiv_def norm_eq_sqrt_inner
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67962
diff changeset
   444
    by (rule derivative_eq_intros | force simp add: inner_commute sgn_div_norm norm_eq_sqrt_inner assms)+
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   445
56181
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 54230
diff changeset
   446
lemmas has_derivative_norm = GDERIV_norm [unfolded gderiv_def]
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   447
81100
6ae3d0b2b8ad tuned whitespace;
wenzelm
parents: 81097
diff changeset
   448
bundle inner_syntax
6ae3d0b2b8ad tuned whitespace;
wenzelm
parents: 81097
diff changeset
   449
begin
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 70136
diff changeset
   450
notation inner (infix \<open>\<bullet>\<close> 70)
29993
84b2c432b94a new theory of real inner product spaces
huffman
parents:
diff changeset
   451
end
69674
fc252acb7100 bundle syntax for inner
immler
parents: 69600
diff changeset
   452
fc252acb7100 bundle syntax for inner
immler
parents: 69600
diff changeset
   453
end