| author | noschinl | 
| Mon, 13 Apr 2015 10:39:49 +0200 | |
| changeset 60048 | e9c30726ca8e | 
| parent 58877 | 262572d90bc6 | 
| child 60420 | 884f54e01427 | 
| permissions | -rw-r--r-- | 
| 41959 | 1  | 
(* Title: HOL/Multivariate_Analysis/L2_Norm.thy  | 
| 36333 | 2  | 
Author: Brian Huffman, Portland State University  | 
3  | 
*)  | 
|
4  | 
||
| 58877 | 5  | 
section {* Square root of sum of squares *}
 | 
| 36333 | 6  | 
|
7  | 
theory L2_Norm  | 
|
8  | 
imports NthRoot  | 
|
9  | 
begin  | 
|
10  | 
||
11  | 
definition  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
49962 
diff
changeset
 | 
12  | 
"setL2 f A = sqrt (\<Sum>i\<in>A. (f i)\<^sup>2)"  | 
| 36333 | 13  | 
|
14  | 
lemma setL2_cong:  | 
|
15  | 
"\<lbrakk>A = B; \<And>x. x \<in> B \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> setL2 f A = setL2 g B"  | 
|
16  | 
unfolding setL2_def by simp  | 
|
17  | 
||
18  | 
lemma strong_setL2_cong:  | 
|
19  | 
"\<lbrakk>A = B; \<And>x. x \<in> B =simp=> f x = g x\<rbrakk> \<Longrightarrow> setL2 f A = setL2 g B"  | 
|
20  | 
unfolding setL2_def simp_implies_def by simp  | 
|
21  | 
||
22  | 
lemma setL2_infinite [simp]: "\<not> finite A \<Longrightarrow> setL2 f A = 0"  | 
|
23  | 
unfolding setL2_def by simp  | 
|
24  | 
||
25  | 
lemma setL2_empty [simp]: "setL2 f {} = 0"
 | 
|
26  | 
unfolding setL2_def by simp  | 
|
27  | 
||
28  | 
lemma setL2_insert [simp]:  | 
|
29  | 
"\<lbrakk>finite F; a \<notin> F\<rbrakk> \<Longrightarrow>  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
49962 
diff
changeset
 | 
30  | 
setL2 f (insert a F) = sqrt ((f a)\<^sup>2 + (setL2 f F)\<^sup>2)"  | 
| 36333 | 31  | 
unfolding setL2_def by (simp add: setsum_nonneg)  | 
32  | 
||
33  | 
lemma setL2_nonneg [simp]: "0 \<le> setL2 f A"  | 
|
34  | 
unfolding setL2_def by (simp add: setsum_nonneg)  | 
|
35  | 
||
36  | 
lemma setL2_0': "\<forall>a\<in>A. f a = 0 \<Longrightarrow> setL2 f A = 0"  | 
|
37  | 
unfolding setL2_def by simp  | 
|
38  | 
||
39  | 
lemma setL2_constant: "setL2 (\<lambda>x. y) A = sqrt (of_nat (card A)) * \<bar>y\<bar>"  | 
|
40  | 
unfolding setL2_def by (simp add: real_sqrt_mult)  | 
|
41  | 
||
42  | 
lemma setL2_mono:  | 
|
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"  | 
|
45  | 
shows "setL2 f K \<le> setL2 g K"  | 
|
46  | 
unfolding setL2_def  | 
|
| 41891 | 47  | 
by (simp add: setsum_nonneg setsum_mono power_mono assms)  | 
| 36333 | 48  | 
|
49  | 
lemma setL2_strict_mono:  | 
|
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"  | 
|
53  | 
shows "setL2 f K < setL2 g K"  | 
|
54  | 
unfolding setL2_def  | 
|
55  | 
by (simp add: setsum_strict_mono power_strict_mono assms)  | 
|
56  | 
||
57  | 
lemma setL2_right_distrib:  | 
|
58  | 
"0 \<le> r \<Longrightarrow> r * setL2 f A = setL2 (\<lambda>x. r * f x) A"  | 
|
59  | 
unfolding setL2_def  | 
|
60  | 
apply (simp add: power_mult_distrib)  | 
|
61  | 
apply (simp add: setsum_right_distrib [symmetric])  | 
|
62  | 
apply (simp add: real_sqrt_mult setsum_nonneg)  | 
|
63  | 
done  | 
|
64  | 
||
65  | 
lemma setL2_left_distrib:  | 
|
66  | 
"0 \<le> r \<Longrightarrow> setL2 f A * r = setL2 (\<lambda>x. f x * r) A"  | 
|
67  | 
unfolding setL2_def  | 
|
68  | 
apply (simp add: power_mult_distrib)  | 
|
69  | 
apply (simp add: setsum_left_distrib [symmetric])  | 
|
70  | 
apply (simp add: real_sqrt_mult setsum_nonneg)  | 
|
71  | 
done  | 
|
72  | 
||
73  | 
lemma setsum_nonneg_eq_0_iff:  | 
|
74  | 
fixes f :: "'a \<Rightarrow> 'b::ordered_ab_group_add"  | 
|
75  | 
shows "\<lbrakk>finite A; \<forall>x\<in>A. 0 \<le> f x\<rbrakk> \<Longrightarrow> setsum f A = 0 \<longleftrightarrow> (\<forall>x\<in>A. f x = 0)"  | 
|
76  | 
apply (induct set: finite, simp)  | 
|
77  | 
apply (simp add: add_nonneg_eq_0_iff setsum_nonneg)  | 
|
78  | 
done  | 
|
79  | 
||
80  | 
lemma setL2_eq_0_iff: "finite A \<Longrightarrow> setL2 f A = 0 \<longleftrightarrow> (\<forall>x\<in>A. f x = 0)"  | 
|
81  | 
unfolding setL2_def  | 
|
82  | 
by (simp add: setsum_nonneg setsum_nonneg_eq_0_iff)  | 
|
83  | 
||
84  | 
lemma setL2_triangle_ineq:  | 
|
85  | 
shows "setL2 (\<lambda>i. f i + g i) A \<le> setL2 f A + setL2 g A"  | 
|
86  | 
proof (cases "finite A")  | 
|
87  | 
case False  | 
|
88  | 
thus ?thesis by simp  | 
|
89  | 
next  | 
|
90  | 
case True  | 
|
91  | 
thus ?thesis  | 
|
92  | 
proof (induct set: finite)  | 
|
93  | 
case empty  | 
|
94  | 
show ?case by simp  | 
|
95  | 
next  | 
|
96  | 
case (insert x F)  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
49962 
diff
changeset
 | 
97  | 
hence "sqrt ((f x + g x)\<^sup>2 + (setL2 (\<lambda>i. f i + g i) F)\<^sup>2) \<le>  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
49962 
diff
changeset
 | 
98  | 
sqrt ((f x + g x)\<^sup>2 + (setL2 f F + setL2 g F)\<^sup>2)"  | 
| 36333 | 99  | 
by (intro real_sqrt_le_mono add_left_mono power_mono insert  | 
100  | 
setL2_nonneg add_increasing zero_le_power2)  | 
|
101  | 
also have  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
49962 
diff
changeset
 | 
102  | 
"\<dots> \<le> sqrt ((f x)\<^sup>2 + (setL2 f F)\<^sup>2) + sqrt ((g x)\<^sup>2 + (setL2 g F)\<^sup>2)"  | 
| 36333 | 103  | 
by (rule real_sqrt_sum_squares_triangle_ineq)  | 
104  | 
finally show ?case  | 
|
105  | 
using insert by simp  | 
|
106  | 
qed  | 
|
107  | 
qed  | 
|
108  | 
||
109  | 
lemma sqrt_sum_squares_le_sum:  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
49962 
diff
changeset
 | 
110  | 
"\<lbrakk>0 \<le> x; 0 \<le> y\<rbrakk> \<Longrightarrow> sqrt (x\<^sup>2 + y\<^sup>2) \<le> x + y"  | 
| 36333 | 111  | 
apply (rule power2_le_imp_le)  | 
| 56536 | 112  | 
apply (simp add: power2_sum)  | 
| 44142 | 113  | 
apply simp  | 
| 36333 | 114  | 
done  | 
115  | 
||
116  | 
lemma setL2_le_setsum [rule_format]:  | 
|
117  | 
"(\<forall>i\<in>A. 0 \<le> f i) \<longrightarrow> setL2 f A \<le> setsum f A"  | 
|
118  | 
apply (cases "finite A")  | 
|
119  | 
apply (induct set: finite)  | 
|
120  | 
apply simp  | 
|
121  | 
apply clarsimp  | 
|
122  | 
apply (erule order_trans [OF sqrt_sum_squares_le_sum])  | 
|
123  | 
apply simp  | 
|
124  | 
apply simp  | 
|
125  | 
apply simp  | 
|
126  | 
done  | 
|
127  | 
||
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
49962 
diff
changeset
 | 
128  | 
lemma sqrt_sum_squares_le_sum_abs: "sqrt (x\<^sup>2 + y\<^sup>2) \<le> \<bar>x\<bar> + \<bar>y\<bar>"  | 
| 36333 | 129  | 
apply (rule power2_le_imp_le)  | 
| 56536 | 130  | 
apply (simp add: power2_sum)  | 
| 44142 | 131  | 
apply simp  | 
| 36333 | 132  | 
done  | 
133  | 
||
134  | 
lemma setL2_le_setsum_abs: "setL2 f A \<le> (\<Sum>i\<in>A. \<bar>f i\<bar>)"  | 
|
135  | 
apply (cases "finite A")  | 
|
136  | 
apply (induct set: finite)  | 
|
137  | 
apply simp  | 
|
138  | 
apply simp  | 
|
139  | 
apply (rule order_trans [OF sqrt_sum_squares_le_sum_abs])  | 
|
140  | 
apply simp  | 
|
141  | 
apply simp  | 
|
142  | 
done  | 
|
143  | 
||
144  | 
lemma setL2_mult_ineq_lemma:  | 
|
145  | 
fixes a b c d :: real  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
49962 
diff
changeset
 | 
146  | 
shows "2 * (a * c) * (b * d) \<le> a\<^sup>2 * d\<^sup>2 + b\<^sup>2 * c\<^sup>2"  | 
| 36333 | 147  | 
proof -  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
49962 
diff
changeset
 | 
148  | 
have "0 \<le> (a * d - b * c)\<^sup>2" by simp  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
49962 
diff
changeset
 | 
149  | 
also have "\<dots> = a\<^sup>2 * d\<^sup>2 + b\<^sup>2 * c\<^sup>2 - 2 * (a * d) * (b * c)"  | 
| 36333 | 150  | 
by (simp only: power2_diff power_mult_distrib)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
49962 
diff
changeset
 | 
151  | 
also have "\<dots> = a\<^sup>2 * d\<^sup>2 + b\<^sup>2 * c\<^sup>2 - 2 * (a * c) * (b * d)"  | 
| 36333 | 152  | 
by simp  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
49962 
diff
changeset
 | 
153  | 
finally show "2 * (a * c) * (b * d) \<le> a\<^sup>2 * d\<^sup>2 + b\<^sup>2 * c\<^sup>2"  | 
| 36333 | 154  | 
by simp  | 
155  | 
qed  | 
|
156  | 
||
157  | 
lemma setL2_mult_ineq: "(\<Sum>i\<in>A. \<bar>f i\<bar> * \<bar>g i\<bar>) \<le> setL2 f A * setL2 g A"  | 
|
158  | 
apply (cases "finite A")  | 
|
159  | 
apply (induct set: finite)  | 
|
160  | 
apply simp  | 
|
161  | 
apply (rule power2_le_imp_le, simp)  | 
|
162  | 
apply (rule order_trans)  | 
|
163  | 
apply (rule power_mono)  | 
|
164  | 
apply (erule add_left_mono)  | 
|
| 56536 | 165  | 
apply (simp add: setsum_nonneg)  | 
| 36333 | 166  | 
apply (simp add: power2_sum)  | 
167  | 
apply (simp add: power_mult_distrib)  | 
|
| 
49962
 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 
webertj 
parents: 
44142 
diff
changeset
 | 
168  | 
apply (simp add: distrib_left distrib_right)  | 
| 36333 | 169  | 
apply (rule ord_le_eq_trans)  | 
170  | 
apply (rule setL2_mult_ineq_lemma)  | 
|
| 56536 | 171  | 
apply simp_all  | 
| 36333 | 172  | 
done  | 
173  | 
||
174  | 
lemma member_le_setL2: "\<lbrakk>finite A; i \<in> A\<rbrakk> \<Longrightarrow> f i \<le> setL2 f A"  | 
|
175  | 
  apply (rule_tac s="insert i (A - {i})" and t="A" in subst)
 | 
|
176  | 
apply fast  | 
|
177  | 
apply (subst setL2_insert)  | 
|
178  | 
apply simp  | 
|
179  | 
apply simp  | 
|
180  | 
apply simp  | 
|
181  | 
done  | 
|
182  | 
||
183  | 
end  |