# HG changeset patch # User huffman # Date 1395169110 25200 # Node ID c7dfd924a1658cd45ec10cd820c0d81c121b70bf # Parent 9ffbb4004c81ac7f43c5d83b43ba71afb71a8ca0 adapt to Isabelle/c726ecfb22b6 diff -r 9ffbb4004c81 -r c7dfd924a165 src/HOL/Decision_Procs/Approximation.thy --- a/src/HOL/Decision_Procs/Approximation.thy Tue Mar 18 16:29:32 2014 +0100 +++ b/src/HOL/Decision_Procs/Approximation.thy Tue Mar 18 11:58:30 2014 -0700 @@ -371,7 +371,7 @@ have "\ real x \ \ 1" using `0 \ real x` `real x \ 1` by auto from mp[OF summable_Leibniz(2)[OF zeroseq_arctan_series[OF this] monoseq_arctan_series[OF this]] prem, THEN spec, of m, unfolded `2 * m = n`] - show ?thesis unfolding arctan_series[OF `\ real x \ \ 1`] Suc_eq_plus1 . + show ?thesis unfolding arctan_series[OF `\ real x \ \ 1`] Suc_eq_plus1 atLeast0LessThan . qed auto note arctan_bounds = this[unfolded atLeastAtMost_iff] @@ -744,7 +744,7 @@ also have "\ = (\ j = 0 ..< n. -1 ^ ((2 * j) div 2) / (real (fact (2 * j))) * x ^(2 * j)) + (\ j = 0 ..< n. 0)" by auto also have "\ = (\ i = 0 ..< 2 * n. if even i then -1 ^ (i div 2) / (real (fact i)) * x ^ i else 0)" - unfolding sum_split_even_odd .. + unfolding sum_split_even_odd atLeast0LessThan .. also have "\ = (\ i = 0 ..< 2 * n. (if even i then -1 ^ (i div 2) / (real (fact i)) else 0) * x ^ i)" by (rule setsum_cong2) auto finally show ?thesis by assumption @@ -758,7 +758,7 @@ + (cos (t + 1/2 * (2 * n) * pi) / real (fact (2*n))) * (real x)^(2*n)" (is "_ = ?SUM + ?rest / ?fact * ?pow") using Maclaurin_cos_expansion2[OF `0 < real x` `0 < 2 * n`] - unfolding cos_coeff_def by auto + unfolding cos_coeff_def atLeast0LessThan by auto have "cos t * -1^n = cos t * cos (n * pi) + sin t * sin (n * pi)" by auto also have "\ = cos (t + n * pi)" using cos_add by auto @@ -860,7 +860,7 @@ have pow: "!!i. x ^ (2 * i + 1) = x * x ^ (2 * i)" by auto have "?SUM = (\ j = 0 ..< n. 0) + ?SUM" by auto also have "\ = (\ i = 0 ..< 2 * n. if even i then 0 else -1 ^ ((i - Suc 0) div 2) / (real (fact i)) * x ^ i)" - unfolding sum_split_even_odd .. + unfolding sum_split_even_odd atLeast0LessThan .. also have "\ = (\ i = 0 ..< 2 * n. (if even i then 0 else -1 ^ ((i - Suc 0) div 2) / (real (fact i))) * x ^ i)" by (rule setsum_cong2) auto finally show ?thesis by assumption @@ -873,7 +873,7 @@ + (sin (t + 1/2 * (2 * n + 1) * pi) / real (fact (2*n + 1))) * (real x)^(2*n + 1)" (is "_ = ?SUM + ?rest / ?fact * ?pow") using Maclaurin_sin_expansion3[OF `0 < 2 * n + 1` `0 < real x`] - unfolding sin_coeff_def by auto + unfolding sin_coeff_def atLeast0LessThan by auto have "?rest = cos t * -1^n" unfolding sin_add cos_add real_of_nat_add distrib_right distrib_left by auto moreover @@ -1344,7 +1344,7 @@ also have "\ \ exp x" proof - obtain t where "\t\ \ \real x\" and "exp x = (\m = 0.. exp t / real (fact (get_even n)) * (real x) ^ (get_even n)" by (auto intro!: mult_nonneg_nonneg divide_nonneg_pos simp add: zero_le_even_power) ultimately show ?thesis @@ -1364,7 +1364,7 @@ qed obtain t where "\t\ \ \real x\" and "exp x = (\m = 0.. 0" by (auto intro!: mult_nonneg_nonpos divide_nonpos_pos simp add: x_less_zero) ultimately have "exp x \ (\j = 0.. x ^ Suc n" by auto qed auto } from summable_Leibniz'(2,4)[OF `?a ----> 0` `\n. 0 \ ?a n`, OF `\n. ?a (Suc n) \ ?a n`, unfolded ln_eq] - show "?lb" and "?ub" by auto + show "?lb" and "?ub" unfolding atLeast0LessThan by auto qed lemma ln_float_bounds: @@ -3031,7 +3031,7 @@ and fl_eq: "interpret_floatarith f (xs[x := xs ! x]) = (\m = 0.. {lx .. ux}" by (cases "xs ! x < c", auto) diff -r 9ffbb4004c81 -r c7dfd924a165 src/HOL/ex/HarmonicSeries.thy --- a/src/HOL/ex/HarmonicSeries.thy Tue Mar 18 16:29:32 2014 +0100 +++ b/src/HOL/ex/HarmonicSeries.thy Tue Mar 18 11:58:30 2014 -0700 @@ -284,7 +284,7 @@ proof - have "\n. 0 \ ?f n" by simp with sf have "?s \ 0" - by - (rule suminf_0_le, simp_all) + by - (rule suminf_ge_zero, simp_all) then have cgt0: "\2*?s\ \ 0" by simp from ndef have "n = nat \(2*?s)\" . @@ -298,7 +298,7 @@ obtain j where jdef: "j = (2::nat)^n" by simp have "\m\j. 0 < ?f m" by simp - with sf have "(\i\{0..i\{0..i\{1..