--- a/src/HOL/Library/Euclidean_Space.thy Sat Feb 21 09:55:32 2009 -0800
+++ b/src/HOL/Library/Euclidean_Space.thy Sat Feb 21 10:58:25 2009 -0800
@@ -344,6 +344,209 @@
apply (auto simp add: vec_def Cart_eq vec_component Cart_lambda_beta )
using dimindex_ge_1 apply auto done
+subsection {* Square root of sum of squares *}
+
+definition
+ "setL2 f A = sqrt (\<Sum>i\<in>A. (f i)\<twosuperior>)"
+
+lemma setL2_cong:
+ "\<lbrakk>A = B; \<And>x. x \<in> B \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> setL2 f A = setL2 g B"
+ unfolding setL2_def by simp
+
+lemma strong_setL2_cong:
+ "\<lbrakk>A = B; \<And>x. x \<in> B =simp=> f x = g x\<rbrakk> \<Longrightarrow> setL2 f A = setL2 g B"
+ unfolding setL2_def simp_implies_def by simp
+
+lemma setL2_infinite [simp]: "\<not> finite A \<Longrightarrow> setL2 f A = 0"
+ unfolding setL2_def by simp
+
+lemma setL2_empty [simp]: "setL2 f {} = 0"
+ unfolding setL2_def by simp
+
+lemma setL2_insert [simp]:
+ "\<lbrakk>finite F; a \<notin> F\<rbrakk> \<Longrightarrow>
+ setL2 f (insert a F) = sqrt ((f a)\<twosuperior> + (setL2 f F)\<twosuperior>)"
+ unfolding setL2_def by (simp add: setsum_nonneg)
+
+lemma setL2_nonneg [simp]: "0 \<le> setL2 f A"
+ unfolding setL2_def by (simp add: setsum_nonneg)
+
+lemma setL2_0': "\<forall>a\<in>A. f a = 0 \<Longrightarrow> setL2 f A = 0"
+ unfolding setL2_def by simp
+
+lemma setL2_mono:
+ assumes "\<And>i. i \<in> K \<Longrightarrow> f i \<le> g i"
+ assumes "\<And>i. i \<in> K \<Longrightarrow> 0 \<le> f i"
+ shows "setL2 f K \<le> setL2 g K"
+ unfolding setL2_def
+ by (simp add: setsum_nonneg setsum_mono power_mono prems)
+
+lemma setL2_right_distrib:
+ "0 \<le> r \<Longrightarrow> r * setL2 f A = setL2 (\<lambda>x. r * f x) A"
+ unfolding setL2_def
+ apply (simp add: power_mult_distrib)
+ apply (simp add: setsum_right_distrib [symmetric])
+ apply (simp add: real_sqrt_mult setsum_nonneg)
+ done
+
+lemma setL2_left_distrib:
+ "0 \<le> r \<Longrightarrow> setL2 f A * r = setL2 (\<lambda>x. f x * r) A"
+ unfolding setL2_def
+ apply (simp add: power_mult_distrib)
+ apply (simp add: setsum_left_distrib [symmetric])
+ apply (simp add: real_sqrt_mult setsum_nonneg)
+ done
+
+lemma setsum_nonneg_eq_0_iff:
+ fixes f :: "'a \<Rightarrow> 'b::pordered_ab_group_add"
+ 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)"
+ apply (induct set: finite, simp)
+ apply (simp add: add_nonneg_eq_0_iff setsum_nonneg)
+ done
+
+lemma setL2_eq_0_iff: "finite A \<Longrightarrow> setL2 f A = 0 \<longleftrightarrow> (\<forall>x\<in>A. f x = 0)"
+ unfolding setL2_def
+ by (simp add: setsum_nonneg setsum_nonneg_eq_0_iff)
+
+lemma setL2_triangle_ineq:
+ shows "setL2 (\<lambda>i. f i + g i) A \<le> setL2 f A + setL2 g A"
+proof (cases "finite A")
+ case False
+ thus ?thesis by simp
+next
+ case True
+ thus ?thesis
+ proof (induct set: finite)
+ case empty
+ show ?case by simp
+ next
+ case (insert x F)
+ hence "sqrt ((f x + g x)\<twosuperior> + (setL2 (\<lambda>i. f i + g i) F)\<twosuperior>) \<le>
+ sqrt ((f x + g x)\<twosuperior> + (setL2 f F + setL2 g F)\<twosuperior>)"
+ by (intro real_sqrt_le_mono add_left_mono power_mono insert
+ setL2_nonneg add_increasing zero_le_power2)
+ also have
+ "\<dots> \<le> sqrt ((f x)\<twosuperior> + (setL2 f F)\<twosuperior>) + sqrt ((g x)\<twosuperior> + (setL2 g F)\<twosuperior>)"
+ by (rule real_sqrt_sum_squares_triangle_ineq)
+ finally show ?case
+ using insert by simp
+ qed
+qed
+
+lemma sqrt_sum_squares_le_sum:
+ "\<lbrakk>0 \<le> x; 0 \<le> y\<rbrakk> \<Longrightarrow> sqrt (x\<twosuperior> + y\<twosuperior>) \<le> x + y"
+ apply (rule power2_le_imp_le)
+ apply (simp add: power2_sum)
+ apply (simp add: mult_nonneg_nonneg)
+ apply (simp add: add_nonneg_nonneg)
+ done
+
+lemma setL2_le_setsum [rule_format]:
+ "(\<forall>i\<in>A. 0 \<le> f i) \<longrightarrow> setL2 f A \<le> setsum f A"
+ apply (cases "finite A")
+ apply (induct set: finite)
+ apply simp
+ apply clarsimp
+ apply (erule order_trans [OF sqrt_sum_squares_le_sum])
+ apply simp
+ apply simp
+ apply simp
+ done
+
+lemma sqrt_sum_squares_le_sum_abs: "sqrt (x\<twosuperior> + y\<twosuperior>) \<le> \<bar>x\<bar> + \<bar>y\<bar>"
+ apply (rule power2_le_imp_le)
+ apply (simp add: power2_sum)
+ apply (simp add: mult_nonneg_nonneg)
+ apply (simp add: add_nonneg_nonneg)
+ done
+
+lemma setL2_le_setsum_abs: "setL2 f A \<le> (\<Sum>i\<in>A. \<bar>f i\<bar>)"
+ apply (cases "finite A")
+ apply (induct set: finite)
+ apply simp
+ apply simp
+ apply (rule order_trans [OF sqrt_sum_squares_le_sum_abs])
+ apply simp
+ apply simp
+ done
+
+lemma setL2_mult_ineq_lemma:
+ fixes a b c d :: real
+ shows "2 * (a * c) * (b * d) \<le> a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior>"
+proof -
+ have "0 \<le> (a * d - b * c)\<twosuperior>" by simp
+ also have "\<dots> = a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior> - 2 * (a * d) * (b * c)"
+ by (simp only: power2_diff power_mult_distrib)
+ also have "\<dots> = a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior> - 2 * (a * c) * (b * d)"
+ by simp
+ finally show "2 * (a * c) * (b * d) \<le> a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior>"
+ by simp
+qed
+
+lemma setL2_mult_ineq: "(\<Sum>i\<in>A. \<bar>f i\<bar> * \<bar>g i\<bar>) \<le> setL2 f A * setL2 g A"
+ apply (cases "finite A")
+ apply (induct set: finite)
+ apply simp
+ apply (rule power2_le_imp_le, simp)
+ apply (rule order_trans)
+ apply (rule power_mono)
+ apply (erule add_left_mono)
+ apply (simp add: add_nonneg_nonneg mult_nonneg_nonneg setsum_nonneg)
+ apply (simp add: power2_sum)
+ apply (simp add: power_mult_distrib)
+ apply (simp add: right_distrib left_distrib)
+ apply (rule ord_le_eq_trans)
+ apply (rule setL2_mult_ineq_lemma)
+ apply simp
+ apply (intro mult_nonneg_nonneg setL2_nonneg)
+ apply simp
+ done
+
+lemma member_le_setL2: "\<lbrakk>finite A; i \<in> A\<rbrakk> \<Longrightarrow> f i \<le> setL2 f A"
+ apply (rule_tac s="insert i (A - {i})" and t="A" in subst)
+ apply fast
+ apply (subst setL2_insert)
+ apply simp
+ apply simp
+ apply simp
+ done
+
+subsection {* Norms *}
+
+instantiation "^" :: (real_normed_vector, type) real_normed_vector
+begin
+
+definition vector_norm_def:
+ "norm (x::'a^'b) = setL2 (\<lambda>i. norm (x$i)) {1 .. dimindex (UNIV:: 'b set)}"
+
+definition vector_sgn_def:
+ "sgn (x::'a^'b) = scaleR (inverse (norm x)) x"
+
+instance proof
+ fix a :: real and x y :: "'a ^ 'b"
+ show "0 \<le> norm x"
+ unfolding vector_norm_def
+ by (rule setL2_nonneg)
+ show "norm x = 0 \<longleftrightarrow> x = 0"
+ unfolding vector_norm_def
+ by (simp add: setL2_eq_0_iff Cart_eq)
+ show "norm (x + y) \<le> norm x + norm y"
+ unfolding vector_norm_def
+ apply (rule order_trans [OF _ setL2_triangle_ineq])
+ apply (rule setL2_mono)
+ apply (simp add: vector_component norm_triangle_ineq)
+ apply simp
+ done
+ show "norm (scaleR a x) = \<bar>a\<bar> * norm x"
+ unfolding vector_norm_def
+ by (simp add: vector_component norm_scaleR setL2_right_distrib
+ cong: strong_setL2_cong)
+ show "sgn x = scaleR (inverse (norm x)) x"
+ by (rule vector_sgn_def)
+qed
+
+end
+
subsection{* Properties of the dot product. *}
lemma dot_sym: "(x::'a:: {comm_monoid_add, ab_semigroup_mult} ^ 'n) \<bullet> y = y \<bullet> x"
@@ -393,18 +596,7 @@
lemma dot_pos_lt: "(0 < x \<bullet> x) \<longleftrightarrow> (x::'a::{ordered_ring_strict,ring_no_zero_divisors} ^ 'n) \<noteq> 0" using dot_eq_0[of x] dot_pos_le[of x]
by (auto simp add: le_less)
-subsection {* Introduce norms, but defer many properties till we get square roots. *}
-text{* FIXME : This is ugly *}
-defs (overloaded)
- real_of_real_def [code inline, simp]: "real == id"
-
-instantiation "^" :: ("{times, comm_monoid_add}", type) norm begin
-definition real_vector_norm_def: "norm \<equiv> (\<lambda>x. sqrt (real (x \<bullet> x)))"
-instance ..
-end
-
-
-subsection{* The collapse of the general concepts to dimention one. *}
+subsection{* The collapse of the general concepts to dimension one. *}
lemma vector_one: "(x::'a ^1) = (\<chi> i. (x$1))"
by (vector dimindex_def)
@@ -415,11 +607,15 @@
apply (simp only: vector_one[symmetric])
done
+lemma norm_vector_1: "norm (x :: _^1) = norm (x$1)"
+ by (simp add: vector_norm_def dimindex_def)
+
lemma norm_real: "norm(x::real ^ 1) = abs(x$1)"
- by (simp add: real_vector_norm_def)
+ by (simp add: norm_vector_1)
text{* Metric *}
+text {* FIXME: generalize to arbitrary @{text real_normed_vector} types *}
definition dist:: "real ^ 'n \<Rightarrow> real ^ 'n \<Rightarrow> real" where
"dist x y = norm (x - y)"
@@ -531,27 +727,30 @@
text{* Hence derive more interesting properties of the norm. *}
lemma norm_0: "norm (0::real ^ 'n) = 0"
- by (simp add: real_vector_norm_def dot_eq_0)
-
-lemma norm_pos_le: "0 <= norm (x::real^'n)"
- by (simp add: real_vector_norm_def dot_pos_le)
+ by (rule norm_zero)
+
+lemma norm_pos_le: "0 <= norm (x::real^'n)"
+ by (rule norm_ge_zero)
lemma norm_neg: " norm(-x) = norm (x:: real ^ 'n)"
- by (simp add: real_vector_norm_def dot_lneg dot_rneg)
+ by (rule norm_minus_cancel)
lemma norm_sub: "norm(x - y) = norm(y - (x::real ^ 'n))"
- by (metis norm_neg minus_diff_eq)
+ by (rule norm_minus_commute)
lemma norm_mul: "norm(a *s x) = abs(a) * norm x"
- by (simp add: real_vector_norm_def dot_lmult dot_rmult mult_assoc[symmetric] real_sqrt_mult)
+ by (simp add: vector_norm_def vector_component setL2_right_distrib
+ abs_mult cong: strong_setL2_cong)
lemma norm_eq_0_dot: "(norm x = 0) \<longleftrightarrow> (x \<bullet> x = (0::real))"
- by (simp add: real_vector_norm_def)
+ by (simp add: vector_norm_def dot_def setL2_def power2_eq_square)
lemma norm_eq_0: "norm x = 0 \<longleftrightarrow> x = (0::real ^ 'n)"
- by (simp add: real_vector_norm_def dot_eq_0)
+ by (rule norm_eq_zero)
lemma norm_pos_lt: "0 < norm x \<longleftrightarrow> x \<noteq> (0::real ^ 'n)"
- by (metis less_le real_vector_norm_def norm_pos_le norm_eq_0)
+ by (rule zero_less_norm_iff)
+lemma real_vector_norm_def: "norm x = sqrt (x \<bullet> x)"
+ by (simp add: vector_norm_def setL2_def dot_def power2_eq_square)
lemma norm_pow_2: "norm x ^ 2 = x \<bullet> x"
- by (simp add: real_vector_norm_def dot_pos_le)
+ by (simp add: real_vector_norm_def)
lemma norm_eq_0_imp: "norm x = 0 ==> x = (0::real ^'n)" by (metis norm_eq_0)
lemma norm_le_0: "norm x <= 0 \<longleftrightarrow> x = (0::real ^'n)"
- by (metis norm_eq_0 norm_pos_le order_antisym)
+ by (rule norm_le_zero_iff)
lemma vector_mul_eq_0: "(a *s x = 0) \<longleftrightarrow> a = (0::'a::idom) \<or> x = 0"
by vector
lemma vector_mul_lcancel: "a *s x = a *s y \<longleftrightarrow> a = (0::real) \<or> x = y"
@@ -583,19 +782,14 @@
ultimately show ?thesis by metis
qed
-lemma norm_abs[simp]: "abs (norm x) = norm (x::real ^'n)"
- using norm_pos_le[of x] by (simp add: real_abs_def linorder_linear)
+lemma norm_abs: "abs (norm x) = norm (x::real ^'n)"
+ by (rule abs_norm_cancel) (* already declared [simp] *)
lemma norm_cauchy_schwarz_abs: "\<bar>x \<bullet> y\<bar> \<le> norm x * norm y"
using norm_cauchy_schwarz[of x y] norm_cauchy_schwarz[of x "-y"]
by (simp add: real_abs_def dot_rneg norm_neg)
lemma norm_triangle: "norm(x + y) <= norm x + norm (y::real ^'n)"
- unfolding real_vector_norm_def
- apply (rule real_le_lsqrt)
- apply (auto simp add: dot_pos_le real_vector_norm_def[symmetric] norm_pos_le norm_pow_2[symmetric] intro: add_nonneg_nonneg)[1]
- apply (auto simp add: dot_pos_le real_vector_norm_def[symmetric] norm_pos_le norm_pow_2[symmetric] intro: add_nonneg_nonneg)[1]
- apply (simp add: dot_ladd dot_radd dot_sym )
- by (simp add: norm_pow_2[symmetric] power2_eq_square ring_simps norm_cauchy_schwarz)
+ by (rule norm_triangle_ineq)
lemma norm_triangle_sub: "norm (x::real ^'n) <= norm(y) + norm(x - y)"
using norm_triangle[of "y" "x - y"] by (simp add: ring_simps)
@@ -627,19 +821,10 @@
qed
lemma component_le_norm: "i \<in> {1 .. dimindex(UNIV :: 'n set)} ==> \<bar>x$i\<bar> <= norm (x::real ^ 'n)"
-proof(simp add: real_vector_norm_def, rule real_le_rsqrt, clarsimp)
- assume i: "Suc 0 \<le> i" "i \<le> dimindex (UNIV :: 'n set)"
- let ?S = "{1 .. dimindex(UNIV :: 'n set)}"
- let ?f = "(\<lambda>k. if k = i then x$i ^2 else 0)"
- have fS: "finite ?S" by simp
- from i setsum_delta[OF fS, of i "\<lambda>k. x$i ^ 2"]
- have th: "x$i^2 = setsum ?f ?S" by simp
- let ?g = "\<lambda>k. x$k * x$k"
- {fix x assume x: "x \<in> ?S" have "?f x \<le> ?g x" by (simp add: power2_eq_square)}
- with setsum_mono[of ?S ?f ?g]
- have "setsum ?f ?S \<le> setsum ?g ?S" by blast
- then show "x$i ^2 \<le> x \<bullet> (x:: real ^ 'n)" unfolding dot_def th[symmetric] .
-qed
+ apply (simp add: vector_norm_def)
+ apply (rule member_le_setL2, simp_all)
+ done
+
lemma norm_bound_component_le: "norm(x::real ^ 'n) <= e
==> \<forall>i \<in> {1 .. dimindex(UNIV:: 'n set)}. \<bar>x$i\<bar> <= e"
by (metis component_le_norm order_trans)
@@ -649,24 +834,12 @@
by (metis component_le_norm basic_trans_rules(21))
lemma norm_le_l1: "norm (x:: real ^'n) <= setsum(\<lambda>i. \<bar>x$i\<bar>) {1..dimindex(UNIV::'n set)}"
-proof (simp add: real_vector_norm_def, rule real_le_lsqrt,simp add: dot_pos_le, simp add: setsum_mono, simp add: dot_def, induct "dimindex(UNIV::'n set)")
- case 0 thus ?case by simp
-next
- case (Suc n)
- have th: "2 * (\<bar>x$(Suc n)\<bar> * (\<Sum>i = Suc 0..n. \<bar>x$i\<bar>)) \<ge> 0"
- apply simp
- apply (rule mult_nonneg_nonneg)
- by (simp_all add: setsum_abs_ge_zero)
-
- from Suc
- show ?case using th by (simp add: power2_eq_square ring_simps)
-qed
+ by (simp add: vector_norm_def setL2_le_setsum)
lemma real_abs_norm: "\<bar> norm x\<bar> = norm (x :: real ^'n)"
- by (simp add: norm_pos_le)
+ by (rule abs_norm_cancel)
lemma real_abs_sub_norm: "\<bar>norm(x::real ^'n) - norm y\<bar> <= norm(x - y)"
- apply (simp add: abs_le_iff ring_simps)
- by (metis norm_triangle_sub norm_sub)
+ by (rule norm_triangle_ineq3)
lemma norm_le: "norm(x::real ^ 'n) <= norm(y) \<longleftrightarrow> x \<bullet> x <= y \<bullet> y"
by (simp add: real_vector_norm_def)
lemma norm_lt: "norm(x::real ^'n) < norm(y) \<longleftrightarrow> x \<bullet> x < y \<bullet> y"
@@ -682,13 +855,7 @@
by (simp add: real_vector_norm_def dot_pos_le )
lemma norm_eq_square: "norm(x) = a \<longleftrightarrow> 0 <= a \<and> x \<bullet> x = a^2"
-proof-
- have th: "\<And>x y::real. x^2 = y^2 \<longleftrightarrow> x = y \<or> x = -y" by algebra
- show ?thesis using norm_pos_le[of x]
- apply (simp add: dot_square_norm th)
- apply arith
- done
-qed
+ by (auto simp add: real_vector_norm_def)
lemma real_abs_le_square_iff: "\<bar>x\<bar> \<le> \<bar>y\<bar> \<longleftrightarrow> (x::real)^2 \<le> y^2"
proof-
@@ -698,14 +865,14 @@
qed
lemma norm_le_square: "norm(x) <= a \<longleftrightarrow> 0 <= a \<and> x \<bullet> x <= a^2"
+ apply (simp add: dot_square_norm real_abs_le_square_iff[symmetric])
using norm_pos_le[of x]
- apply (simp add: dot_square_norm real_abs_le_square_iff[symmetric])
apply arith
done
lemma norm_ge_square: "norm(x) >= a \<longleftrightarrow> a <= 0 \<or> x \<bullet> x >= a ^ 2"
+ apply (simp add: dot_square_norm real_abs_le_square_iff[symmetric])
using norm_pos_le[of x]
- apply (simp add: dot_square_norm real_abs_le_square_iff[symmetric])
apply arith
done
@@ -917,10 +1084,10 @@
assumes fS: "finite S"
shows "norm (setsum f S) <= setsum (\<lambda>x. norm(f x)) S"
proof(induct rule: finite_induct[OF fS])
- case 1 thus ?case by simp norm
+ case 1 thus ?case by simp
next
case (2 x S)
- from "2.hyps" have "norm (setsum f (insert x S)) \<le> norm (f x) + norm (setsum f S)" apply (simp add: norm_triangle_ineq) by norm
+ from "2.hyps" have "norm (setsum f (insert x S)) \<le> norm (f x) + norm (setsum f S)" by (simp add: norm_triangle_ineq)
also have "\<dots> \<le> norm (f x) + setsum (\<lambda>x. norm(f x)) S"
using "2.hyps" by simp
finally show ?case using "2.hyps" by simp
@@ -1552,7 +1719,9 @@
{fix x::"real ^ 'n"
have "norm (f x) \<le> ?K * norm x"
using B[rule_format, of x] norm_pos_le[of x] norm_pos_le[of "f x"] Bp
- by (auto simp add: ring_simps split add: abs_split)
+ apply (auto simp add: ring_simps split add: abs_split)
+ apply (erule order_trans, simp)
+ done
}
then show ?thesis using Kp by blast
qed
@@ -2268,7 +2437,7 @@
{assume H: ?lhs
from H[rule_format, of "basis 1"]
have bp: "b \<ge> 0" using norm_pos_le[of "f (basis 1)"] dimindex_ge_1[of "UNIV:: 'n set"]
- by (auto simp add: norm_basis)
+ by (auto simp add: norm_basis elim: order_trans [OF norm_ge_zero])
{fix x :: "real ^'n"
{assume "x = 0"
then have "norm (f x) \<le> b * norm x" by (simp add: linear_0[OF lf] norm_0 bp)}
@@ -2276,7 +2445,7 @@
{assume x0: "x \<noteq> 0"
hence n0: "norm x \<noteq> 0" by (metis norm_eq_0)
let ?c = "1/ norm x"
- have "norm (?c*s x) = 1" by (simp add: n0 norm_mul)
+ have "norm (?c*s x) = 1" using x0 by (simp add: n0 norm_mul)
with H have "norm (f(?c*s x)) \<le> b" by blast
hence "?c * norm (f x) \<le> b"
by (simp add: linear_cmul[OF lf] norm_mul)
@@ -2585,7 +2754,7 @@
by (simp add: dot_def setsum_add_split[OF th_0, of _ ?m] pastecart_def dimindex_finite_sum Cart_lambda_beta setsum_nonneg zero_le_square del: One_nat_def)
then show ?thesis
unfolding th0
- unfolding real_vector_norm_def real_sqrt_le_iff real_of_real_def id_def
+ unfolding real_vector_norm_def real_sqrt_le_iff id_def
by (simp add: dot_def dimindex_finite_sum Cart_lambda_beta)
qed
@@ -2617,7 +2786,7 @@
by (simp add: dot_def setsum_add_split[OF th_0, of _ ?m] pastecart_def dimindex_finite_sum Cart_lambda_beta setsum_nonneg zero_le_square setsum_reindex[OF finj, unfolded fS] del: One_nat_def)
then show ?thesis
unfolding th0
- unfolding real_vector_norm_def real_sqrt_le_iff real_of_real_def id_def
+ unfolding real_vector_norm_def real_sqrt_le_iff id_def
by (simp add: dot_def dimindex_finite_sum Cart_lambda_beta)
qed
@@ -2674,7 +2843,7 @@
qed
lemma norm_pastecart: "norm(pastecart x y) <= norm(x :: real ^ _) + norm(y)"
- unfolding real_vector_norm_def dot_pastecart real_sqrt_le_iff real_of_real_def id_def
+ unfolding real_vector_norm_def dot_pastecart real_sqrt_le_iff id_def
apply (rule power2_le_imp_le)
apply (simp add: real_sqrt_pow2[OF add_nonneg_nonneg[OF dot_pos_le[of x] dot_pos_le[of y]]])
apply (auto simp add: power2_eq_square ring_simps)
@@ -4998,7 +5167,7 @@
apply blast
by (rule abs_ge_zero)
from real_le_lsqrt[OF dot_pos_le th th1]
- show ?thesis unfolding real_vector_norm_def real_of_real_def id_def .
+ show ?thesis unfolding real_vector_norm_def id_def .
qed
(* Equality in Cauchy-Schwarz and triangle inequalities. *)