| author | nipkow |
| Wed, 16 Jan 2019 17:03:31 +0100 | |
| changeset 69669 | de2f0a24b0f0 |
| parent 69600 | 86e8e7347ac0 |
| child 69676 | 56acd449da41 |
| permissions | -rw-r--r-- |
| 63627 | 1 |
(* Title: HOL/Analysis/L2_Norm.thy |
| 36333 | 2 |
Author: Brian Huffman, Portland State University |
3 |
*) |
|
4 |
||
| 67143 | 5 |
section \<open>L2 Norm\<close> |
| 36333 | 6 |
|
7 |
theory L2_Norm |
|
| 67006 | 8 |
imports Complex_Main |
| 36333 | 9 |
begin |
10 |
||
| 69600 | 11 |
definition%important L2_set :: "('a \<Rightarrow> real) \<Rightarrow> 'a set \<Rightarrow> real" where
|
12 |
"L2_set f A = sqrt (\<Sum>i\<in>A. (f i)\<^sup>2)" |
|
| 36333 | 13 |
|
| 67155 | 14 |
lemma L2_set_cong: |
15 |
"\<lbrakk>A = B; \<And>x. x \<in> B \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> L2_set f A = L2_set g B" |
|
16 |
unfolding L2_set_def by simp |
|
| 36333 | 17 |
|
|
69546
27dae626822b
prefer naming convention from datatype package for strong congruence rules
haftmann
parents:
69164
diff
changeset
|
18 |
lemma L2_set_cong_simp: |
| 67155 | 19 |
"\<lbrakk>A = B; \<And>x. x \<in> B =simp=> f x = g x\<rbrakk> \<Longrightarrow> L2_set f A = L2_set g B" |
20 |
unfolding L2_set_def simp_implies_def by simp |
|
| 36333 | 21 |
|
| 67155 | 22 |
lemma L2_set_infinite [simp]: "\<not> finite A \<Longrightarrow> L2_set f A = 0" |
23 |
unfolding L2_set_def by simp |
|
24 |
||
25 |
lemma L2_set_empty [simp]: "L2_set f {} = 0"
|
|
26 |
unfolding L2_set_def by simp |
|
| 36333 | 27 |
|
| 67155 | 28 |
lemma L2_set_insert [simp]: |
29 |
"\<lbrakk>finite F; a \<notin> F\<rbrakk> \<Longrightarrow> |
|
30 |
L2_set f (insert a F) = sqrt ((f a)\<^sup>2 + (L2_set f F)\<^sup>2)" |
|
31 |
unfolding L2_set_def by (simp add: sum_nonneg) |
|
| 36333 | 32 |
|
| 67155 | 33 |
lemma L2_set_nonneg [simp]: "0 \<le> L2_set f A" |
34 |
unfolding L2_set_def by (simp add: sum_nonneg) |
|
| 36333 | 35 |
|
| 67155 | 36 |
lemma L2_set_0': "\<forall>a\<in>A. f a = 0 \<Longrightarrow> L2_set f A = 0" |
37 |
unfolding L2_set_def by simp |
|
| 36333 | 38 |
|
| 67155 | 39 |
lemma L2_set_constant: "L2_set (\<lambda>x. y) A = sqrt (of_nat (card A)) * \<bar>y\<bar>" |
40 |
unfolding L2_set_def by (simp add: real_sqrt_mult) |
|
| 36333 | 41 |
|
| 67155 | 42 |
lemma L2_set_mono: |
| 36333 | 43 |
assumes "\<And>i. i \<in> K \<Longrightarrow> f i \<le> g i" |
44 |
assumes "\<And>i. i \<in> K \<Longrightarrow> 0 \<le> f i" |
|
| 67155 | 45 |
shows "L2_set f K \<le> L2_set g K" |
46 |
unfolding L2_set_def |
|
| 64267 | 47 |
by (simp add: sum_nonneg sum_mono power_mono assms) |
| 36333 | 48 |
|
| 67155 | 49 |
lemma L2_set_strict_mono: |
| 36333 | 50 |
assumes "finite K" and "K \<noteq> {}"
|
51 |
assumes "\<And>i. i \<in> K \<Longrightarrow> f i < g i" |
|
52 |
assumes "\<And>i. i \<in> K \<Longrightarrow> 0 \<le> f i" |
|
| 67155 | 53 |
shows "L2_set f K < L2_set g K" |
54 |
unfolding L2_set_def |
|
| 64267 | 55 |
by (simp add: sum_strict_mono power_strict_mono assms) |
| 36333 | 56 |
|
| 67155 | 57 |
lemma L2_set_right_distrib: |
58 |
"0 \<le> r \<Longrightarrow> r * L2_set f A = L2_set (\<lambda>x. r * f x) A" |
|
59 |
unfolding L2_set_def |
|
| 36333 | 60 |
apply (simp add: power_mult_distrib) |
| 64267 | 61 |
apply (simp add: sum_distrib_left [symmetric]) |
62 |
apply (simp add: real_sqrt_mult sum_nonneg) |
|
| 36333 | 63 |
done |
64 |
||
| 67155 | 65 |
lemma L2_set_left_distrib: |
66 |
"0 \<le> r \<Longrightarrow> L2_set f A * r = L2_set (\<lambda>x. f x * r) A" |
|
67 |
unfolding L2_set_def |
|
| 36333 | 68 |
apply (simp add: power_mult_distrib) |
| 64267 | 69 |
apply (simp add: sum_distrib_right [symmetric]) |
70 |
apply (simp add: real_sqrt_mult sum_nonneg) |
|
| 36333 | 71 |
done |
72 |
||
| 67155 | 73 |
lemma L2_set_eq_0_iff: "finite A \<Longrightarrow> L2_set f A = 0 \<longleftrightarrow> (\<forall>x\<in>A. f x = 0)" |
74 |
unfolding L2_set_def |
|
| 64267 | 75 |
by (simp add: sum_nonneg sum_nonneg_eq_0_iff) |
| 36333 | 76 |
|
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68465
diff
changeset
|
77 |
proposition L2_set_triangle_ineq: |
| 67156 | 78 |
"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
|
79 |
proof (cases "finite A") |
| 36333 | 80 |
case False |
81 |
thus ?thesis by simp |
|
82 |
next |
|
83 |
case True |
|
84 |
thus ?thesis |
|
85 |
proof (induct set: finite) |
|
86 |
case empty |
|
87 |
show ?case by simp |
|
88 |
next |
|
89 |
case (insert x F) |
|
| 67155 | 90 |
hence "sqrt ((f x + g x)\<^sup>2 + (L2_set (\<lambda>i. f i + g i) F)\<^sup>2) \<le> |
91 |
sqrt ((f x + g x)\<^sup>2 + (L2_set f F + L2_set g F)\<^sup>2)" |
|
| 36333 | 92 |
by (intro real_sqrt_le_mono add_left_mono power_mono insert |
| 67155 | 93 |
L2_set_nonneg add_increasing zero_le_power2) |
| 36333 | 94 |
also have |
| 67155 | 95 |
"\<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 | 96 |
by (rule real_sqrt_sum_squares_triangle_ineq) |
97 |
finally show ?case |
|
98 |
using insert by simp |
|
99 |
qed |
|
100 |
qed |
|
101 |
||
| 67155 | 102 |
lemma L2_set_le_sum [rule_format]: |
103 |
"(\<forall>i\<in>A. 0 \<le> f i) \<longrightarrow> L2_set f A \<le> sum f A" |
|
| 36333 | 104 |
apply (cases "finite A") |
105 |
apply (induct set: finite) |
|
106 |
apply simp |
|
107 |
apply clarsimp |
|
108 |
apply (erule order_trans [OF sqrt_sum_squares_le_sum]) |
|
109 |
apply simp |
|
110 |
apply simp |
|
111 |
apply simp |
|
112 |
done |
|
113 |
||
| 67155 | 114 |
lemma L2_set_le_sum_abs: "L2_set f A \<le> (\<Sum>i\<in>A. \<bar>f i\<bar>)" |
| 36333 | 115 |
apply (cases "finite A") |
116 |
apply (induct set: finite) |
|
117 |
apply simp |
|
118 |
apply simp |
|
119 |
apply (rule order_trans [OF sqrt_sum_squares_le_sum_abs]) |
|
120 |
apply simp |
|
121 |
apply simp |
|
122 |
done |
|
123 |
||
| 67155 | 124 |
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" |
| 36333 | 125 |
apply (cases "finite A") |
126 |
apply (induct set: finite) |
|
127 |
apply simp |
|
128 |
apply (rule power2_le_imp_le, simp) |
|
129 |
apply (rule order_trans) |
|
130 |
apply (rule power_mono) |
|
131 |
apply (erule add_left_mono) |
|
| 64267 | 132 |
apply (simp add: sum_nonneg) |
| 36333 | 133 |
apply (simp add: power2_sum) |
134 |
apply (simp add: power_mult_distrib) |
|
|
49962
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents:
44142
diff
changeset
|
135 |
apply (simp add: distrib_left distrib_right) |
| 36333 | 136 |
apply (rule ord_le_eq_trans) |
| 67155 | 137 |
apply (rule L2_set_mult_ineq_lemma) |
| 56536 | 138 |
apply simp_all |
| 36333 | 139 |
done |
140 |
||
| 67155 | 141 |
lemma member_le_L2_set: "\<lbrakk>finite A; i \<in> A\<rbrakk> \<Longrightarrow> f i \<le> L2_set f A" |
142 |
unfolding L2_set_def |
|
| 64267 | 143 |
by (auto intro!: member_le_sum real_le_rsqrt) |
| 36333 | 144 |
|
145 |
end |