# HG changeset patch # User paulson # Date 1427816928 -3600 # Node ID 8a20dd967385afb2917e0b3d9b85da119e07d4ff # Parent 30519ff3dffb45ebe02864f8f12409c8510e904a rationalised and generalised some theorems concerning abs and x^2. diff -r 30519ff3dffb -r 8a20dd967385 src/HOL/Multivariate_Analysis/Bounded_Continuous_Function.thy --- a/src/HOL/Multivariate_Analysis/Bounded_Continuous_Function.thy Tue Mar 31 15:01:06 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/Bounded_Continuous_Function.thy Tue Mar 31 16:48:48 2015 +0100 @@ -407,7 +407,7 @@ hence "(\i\Basis. (dist (clamp a b y \ i) (clamp a b x \ i))\<^sup>2) \ (\i\Basis. (dist (y \ i) (x \ i))\<^sup>2)" by (auto intro!: setsum_mono - simp add: clamp_def dist_real_def real_abs_le_square_iff[symmetric]) + simp add: clamp_def dist_real_def abs_le_square_iff[symmetric]) thus ?thesis by (auto intro: real_sqrt_le_mono simp add: euclidean_dist_l2[where y=x] euclidean_dist_l2[where y="clamp a b x"] setL2_def) diff -r 30519ff3dffb -r 8a20dd967385 src/HOL/Multivariate_Analysis/Linear_Algebra.thy --- a/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Tue Mar 31 15:01:06 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Tue Mar 31 16:48:48 2015 +0100 @@ -68,25 +68,14 @@ lemma norm_eq_square: "norm x = a \ 0 \ a \ x \ x = a\<^sup>2" by (auto simp add: norm_eq_sqrt_inner) -lemma real_abs_le_square_iff: "\x\ \ \y\ \ (x::real)\<^sup>2 \ y\<^sup>2" -proof - assume "\x\ \ \y\" - then have "\x\\<^sup>2 \ \y\\<^sup>2" by (rule power_mono, simp) - then show "x\<^sup>2 \ y\<^sup>2" by simp -next - assume "x\<^sup>2 \ y\<^sup>2" - then have "sqrt (x\<^sup>2) \ sqrt (y\<^sup>2)" by (rule real_sqrt_le_mono) - then show "\x\ \ \y\" by simp -qed - lemma norm_le_square: "norm x \ a \ 0 \ a \ x \ x \ a\<^sup>2" - apply (simp add: dot_square_norm real_abs_le_square_iff[symmetric]) + apply (simp add: dot_square_norm abs_le_square_iff[symmetric]) using norm_ge_zero[of x] apply arith done lemma norm_ge_square: "norm x \ a \ a \ 0 \ x \ x \ a\<^sup>2" - apply (simp add: dot_square_norm real_abs_le_square_iff[symmetric]) + apply (simp add: dot_square_norm abs_le_square_iff[symmetric]) using norm_ge_zero[of x] apply arith done diff -r 30519ff3dffb -r 8a20dd967385 src/HOL/Power.thy --- a/src/HOL/Power.thy Tue Mar 31 15:01:06 2015 +0100 +++ b/src/HOL/Power.thy Tue Mar 31 16:48:48 2015 +0100 @@ -645,6 +645,29 @@ "0 < x\<^sup>2 + y\<^sup>2 \ x \ 0 \ y \ 0" unfolding not_le [symmetric] by (simp add: sum_power2_le_zero_iff) +lemma abs_le_square_iff: + "\x\ \ \y\ \ x\<^sup>2 \ y\<^sup>2" +proof + assume "\x\ \ \y\" + then have "\x\\<^sup>2 \ \y\\<^sup>2" by (rule power_mono, simp) + then show "x\<^sup>2 \ y\<^sup>2" by simp +next + assume "x\<^sup>2 \ y\<^sup>2" + then show "\x\ \ \y\" + by (auto intro!: power2_le_imp_le [OF _ abs_ge_zero]) +qed + +lemma abs_square_le_1:"x\<^sup>2 \ 1 \ abs(x) \ 1" + using abs_le_square_iff [of x 1] + by simp + +lemma abs_square_eq_1: "x\<^sup>2 = 1 \ abs(x) = 1" + by (auto simp add: abs_if power2_eq_1_iff) + +lemma abs_square_less_1: "x\<^sup>2 < 1 \ abs(x) < 1" + using abs_square_eq_1 [of x] abs_square_le_1 [of x] + by (auto simp add: le_less) + end diff -r 30519ff3dffb -r 8a20dd967385 src/HOL/Rings.thy --- a/src/HOL/Rings.thy Tue Mar 31 15:01:06 2015 +0100 +++ b/src/HOL/Rings.thy Tue Mar 31 16:48:48 2015 +0100 @@ -1260,6 +1260,10 @@ "\x - a\ < r \ a - r < x \ x < a + r" by (auto simp add: diff_less_eq ac_simps abs_less_iff) +lemma abs_diff_le_iff: + "\x - a\ \ r \ a - r \ x \ x \ a + r" + by (auto simp add: diff_le_eq ac_simps abs_le_iff) + end hide_fact (open) comm_mult_left_mono comm_mult_strict_left_mono distrib diff -r 30519ff3dffb -r 8a20dd967385 src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy Tue Mar 31 15:01:06 2015 +0100 +++ b/src/HOL/Transcendental.thy Tue Mar 31 16:48:48 2015 +0100 @@ -4542,16 +4542,6 @@ (is "summable (?c x)") by (rule summable_Leibniz(1), rule zeroseq_arctan_series[OF assms], rule monoseq_arctan_series[OF assms]) -lemma less_one_imp_sqr_less_one: - fixes x :: real - assumes "\x\ < 1" - shows "x\<^sup>2 < 1" -proof - - have "\x\<^sup>2\ < 1" - by (metis abs_power2 assms pos2 power2_abs power_0 power_strict_decreasing zero_eq_power2 zero_less_abs_iff) - thus ?thesis using zero_le_power2 by auto -qed - lemma DERIV_arctan_series: assumes "\ x \ < 1" shows "DERIV (\ x'. \ k. (-1)^k * (1 / real (k*2+1) * x' ^ (k*2+1))) x :> (\ k. (-1)^k * x^(k*2))" @@ -4568,7 +4558,7 @@ { fix x :: real assume "\x\ < 1" - hence "x\<^sup>2 < 1" by (rule less_one_imp_sqr_less_one) + hence "x\<^sup>2 < 1" by (simp add: abs_square_less_1) have "summable (\ n. (- 1) ^ n * (x\<^sup>2) ^n)" by (rule summable_Leibniz(1), auto intro!: LIMSEQ_realpow_zero monoseq_realpow `x\<^sup>2 < 1` order_less_imp_le[OF `x\<^sup>2 < 1`]) hence "summable (\ n. (- 1) ^ n * x^(2*n))" unfolding power_mult . @@ -4676,7 +4666,7 @@ hence "\x\ < r" by auto hence "\x\ < 1" using `r < 1` by auto have "\ - (x\<^sup>2) \ < 1" - using less_one_imp_sqr_less_one[OF `\x\ < 1`] by auto + using abs_square_less_1 `\x\ < 1` by auto hence "(\ n. (- (x\<^sup>2)) ^ n) sums (1 / (1 - (- (x\<^sup>2))))" unfolding real_norm_def[symmetric] by (rule geometric_sums) hence "(?c' x) sums (1 / (1 - (- (x\<^sup>2))))"