src/HOL/Analysis/L2_Norm.thy
author paulson <lp15@cam.ac.uk>
Sat, 12 Apr 2025 22:42:32 +0100 (2 weeks ago)
changeset 82489 d35d355f7330
parent 70136 f03a01a18c6e
child 82522 62afd98e3f3e
permissions -rw-r--r--
Tidied more messy old proofs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63627
6ddb43c6b711 rename HOL-Multivariate_Analysis to HOL-Analysis.
hoelzl
parents: 60974
diff changeset
     1
(*  Title:      HOL/Analysis/L2_Norm.thy
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
     2
    Author:     Brian Huffman, Portland State University
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
     3
*)
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
     4
69676
56acd449da41 chapters for analysis manual
immler
parents: 69600
diff changeset
     5
chapter \<open>Linear Algebra\<close>
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
     6
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
     7
theory L2_Norm
67006
b1278ed3cd46 prefer main entry points of HOL;
wenzelm
parents: 66453
diff changeset
     8
imports Complex_Main
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
     9
begin
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    10
69676
56acd449da41 chapters for analysis manual
immler
parents: 69600
diff changeset
    11
section \<open>L2 Norm\<close>
56acd449da41 chapters for analysis manual
immler
parents: 69600
diff changeset
    12
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69676
diff changeset
    13
definition\<^marker>\<open>tag important\<close> L2_set :: "('a \<Rightarrow> real) \<Rightarrow> 'a set \<Rightarrow> real" where
69600
86e8e7347ac0 typed definitions
nipkow
parents: 69546
diff changeset
    14
"L2_set f A = sqrt (\<Sum>i\<in>A. (f i)\<^sup>2)"
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    15
67155
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    16
lemma L2_set_cong:
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    17
  "\<lbrakk>A = B; \<And>x. x \<in> B \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> L2_set f A = L2_set g B"
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    18
  unfolding L2_set_def by simp
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    19
69546
27dae626822b prefer naming convention from datatype package for strong congruence rules
haftmann
parents: 69164
diff changeset
    20
lemma L2_set_cong_simp:
67155
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    21
  "\<lbrakk>A = B; \<And>x. x \<in> B =simp=> f x = g x\<rbrakk> \<Longrightarrow> L2_set f A = L2_set g B"
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    22
  unfolding L2_set_def simp_implies_def by simp
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    23
67155
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    24
lemma L2_set_infinite [simp]: "\<not> finite A \<Longrightarrow> L2_set f A = 0"
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    25
  unfolding L2_set_def by simp
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    26
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    27
lemma L2_set_empty [simp]: "L2_set f {} = 0"
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    28
  unfolding L2_set_def by simp
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    29
67155
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    30
lemma L2_set_insert [simp]:
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    31
  "\<lbrakk>finite F; a \<notin> F\<rbrakk> \<Longrightarrow>
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    32
    L2_set f (insert a F) = sqrt ((f a)\<^sup>2 + (L2_set f F)\<^sup>2)"
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    33
  unfolding L2_set_def by (simp add: sum_nonneg)
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    34
67155
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    35
lemma L2_set_nonneg [simp]: "0 \<le> L2_set f A"
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    36
  unfolding L2_set_def by (simp add: sum_nonneg)
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    37
67155
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    38
lemma L2_set_0': "\<forall>a\<in>A. f a = 0 \<Longrightarrow> L2_set f A = 0"
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    39
  unfolding L2_set_def by simp
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    40
67155
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    41
lemma L2_set_constant: "L2_set (\<lambda>x. y) A = sqrt (of_nat (card A)) * \<bar>y\<bar>"
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    42
  unfolding L2_set_def by (simp add: real_sqrt_mult)
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    43
67155
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    44
lemma L2_set_mono:
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    45
  assumes "\<And>i. i \<in> K \<Longrightarrow> f i \<le> g i"
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    46
  assumes "\<And>i. i \<in> K \<Longrightarrow> 0 \<le> f i"
67155
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    47
  shows "L2_set f K \<le> L2_set g K"
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    48
  unfolding L2_set_def
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63938
diff changeset
    49
  by (simp add: sum_nonneg sum_mono power_mono assms)
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    50
67155
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    51
lemma L2_set_strict_mono:
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    52
  assumes "finite K" and "K \<noteq> {}"
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    53
  assumes "\<And>i. i \<in> K \<Longrightarrow> f i < g i"
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    54
  assumes "\<And>i. i \<in> K \<Longrightarrow> 0 \<le> f i"
67155
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    55
  shows "L2_set f K < L2_set g K"
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    56
  unfolding L2_set_def
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63938
diff changeset
    57
  by (simp add: sum_strict_mono power_strict_mono assms)
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    58
67155
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    59
lemma L2_set_right_distrib:
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    60
  "0 \<le> r \<Longrightarrow> r * L2_set f A = L2_set (\<lambda>x. r * f x) A"
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    61
  unfolding L2_set_def
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    62
  apply (simp add: power_mult_distrib)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63938
diff changeset
    63
  apply (simp add: sum_distrib_left [symmetric])
b9a1486e79be setsum -> sum
nipkow
parents: 63938
diff changeset
    64
  apply (simp add: real_sqrt_mult sum_nonneg)
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    65
  done
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    66
67155
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    67
lemma L2_set_left_distrib:
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    68
  "0 \<le> r \<Longrightarrow> L2_set f A * r = L2_set (\<lambda>x. f x * r) A"
82489
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
    69
  unfolding L2_set_def power_mult_distrib
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
    70
  by (simp add: real_sqrt_mult sum_nonneg flip: sum_distrib_right)
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    71
67155
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    72
lemma L2_set_eq_0_iff: "finite A \<Longrightarrow> L2_set f A = 0 \<longleftrightarrow> (\<forall>x\<in>A. f x = 0)"
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    73
  unfolding L2_set_def
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63938
diff changeset
    74
  by (simp add: sum_nonneg sum_nonneg_eq_0_iff)
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    75
68607
67bb59e49834 make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents: 68465
diff changeset
    76
proposition L2_set_triangle_ineq:
67156
3a9966b88a50 "important" annotations
nipkow
parents: 67155
diff changeset
    77
  "L2_set (\<lambda>i. f i + g i) A \<le> L2_set f A + L2_set g A"
68607
67bb59e49834 make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents: 68465
diff changeset
    78
proof (cases "finite A")
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    79
  case False
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    80
  thus ?thesis by simp
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    81
next
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    82
  case True
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    83
  thus ?thesis
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    84
  proof (induct set: finite)
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    85
    case empty
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    86
    show ?case by simp
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    87
  next
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    88
    case (insert x F)
67155
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    89
    hence "sqrt ((f x + g x)\<^sup>2 + (L2_set (\<lambda>i. f i + g i) F)\<^sup>2) \<le>
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    90
           sqrt ((f x + g x)\<^sup>2 + (L2_set f F + L2_set g F)\<^sup>2)"
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    91
      by (intro real_sqrt_le_mono add_left_mono power_mono insert
67155
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    92
                L2_set_nonneg add_increasing zero_le_power2)
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    93
    also have
67155
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
    94
      "\<dots> \<le> sqrt ((f x)\<^sup>2 + (L2_set f F)\<^sup>2) + sqrt ((g x)\<^sup>2 + (L2_set g F)\<^sup>2)"
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    95
      by (rule real_sqrt_sum_squares_triangle_ineq)
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    96
    finally show ?case
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    97
      using insert by simp
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    98
  qed
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
    99
qed
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
   100
82489
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   101
lemma L2_set_le_sum:
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   102
  "(\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> L2_set f A \<le> sum f A"
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   103
proof (induction A rule: infinite_finite_induct)
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   104
  case (insert a A)
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   105
  with order_trans [OF sqrt_sum_squares_le_sum] show ?case by force
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   106
qed auto
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
   107
67155
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
   108
lemma L2_set_le_sum_abs: "L2_set f A \<le> (\<Sum>i\<in>A. \<bar>f i\<bar>)"
82489
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   109
proof (induction A rule: infinite_finite_induct)
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   110
  case (insert a A)
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   111
  with order_trans [OF sqrt_sum_squares_le_sum_abs] show ?case by force
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   112
qed auto
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
   113
67155
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
   114
lemma L2_set_mult_ineq: "(\<Sum>i\<in>A. \<bar>f i\<bar> * \<bar>g i\<bar>) \<le> L2_set f A * L2_set g A"
82489
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   115
proof (induction A rule: infinite_finite_induct)
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   116
  case (insert a A)
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   117
  have "(\<bar>f a\<bar> * \<bar>g a\<bar> + (\<Sum>i\<in>A. \<bar>f i\<bar> * \<bar>g i\<bar>))\<^sup>2
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   118
      \<le> (\<bar>f a\<bar> * \<bar>g a\<bar> + L2_set f A * L2_set g A)\<^sup>2"
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   119
    by (simp add: insert.IH sum_nonneg)
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   120
  also have "... \<le> ((f a)\<^sup>2 + (L2_set f A)\<^sup>2) * ((g a)\<^sup>2 + (L2_set g A)\<^sup>2)"
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   121
    using L2_set_mult_ineq_lemma [of "L2_set f A" "L2_set g A" "\<bar>f a\<bar>" "\<bar>g a\<bar>"]
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   122
    by (simp add: power2_eq_square algebra_simps)
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   123
  also have "... = (sqrt ((f a)\<^sup>2 + (L2_set f A)\<^sup>2) * sqrt ((g a)\<^sup>2 + (L2_set g A)\<^sup>2))\<^sup>2"
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   124
    using real_sqrt_mult real_sqrt_sum_squares_mult_squared_eq by presburger
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   125
  finally have "(\<bar>f a\<bar> * \<bar>g a\<bar> + (\<Sum>i\<in>A. \<bar>f i\<bar> * \<bar>g i\<bar>))\<^sup>2
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   126
              \<le> (sqrt ((f a)\<^sup>2 + (L2_set f A)\<^sup>2) * sqrt ((g a)\<^sup>2 + (L2_set g A)\<^sup>2))\<^sup>2" .
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   127
  then
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   128
  show ?case
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   129
    using power2_le_imp_le insert.hyps by fastforce
d35d355f7330 Tidied more messy old proofs
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   130
qed auto
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
   131
67155
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
   132
lemma member_le_L2_set: "\<lbrakk>finite A; i \<in> A\<rbrakk> \<Longrightarrow> f i \<le> L2_set f A"
9e5b05d54f9d canonical name
nipkow
parents: 67144
diff changeset
   133
  unfolding L2_set_def
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63938
diff changeset
   134
  by (auto intro!: member_le_sum real_le_rsqrt)
36333
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
   135
82356c9e218a move l2-norm stuff into separate theory file
huffman
parents:
diff changeset
   136
end