# HG changeset patch # User huffman # Date 1333110332 -7200 # Node ID 1b7c909a6fadb208245919e7b99b2a61e09193c2 # Parent 7205eb4a0a054f5396f50c68cd9343f38d78a343 rephrase lemmas about arithmetic series using numeral '2' diff -r 7205eb4a0a05 -r 1b7c909a6fad src/HOL/SetInterval.thy --- a/src/HOL/SetInterval.thy Fri Mar 30 14:00:18 2012 +0200 +++ b/src/HOL/SetInterval.thy Fri Mar 30 14:25:32 2012 +0200 @@ -1282,19 +1282,21 @@ subsection {* The formula for arithmetic sums *} -lemma gauss_sum: (* FIXME: rephrase in terms of "2" *) - "((1::'a::comm_semiring_1) + 1)*(\i\{1..n}. of_nat i) = +lemma gauss_sum: + "(2::'a::comm_semiring_1)*(\i\{1..n}. of_nat i) = of_nat n*((of_nat n)+1)" proof (induct n) case 0 show ?case by simp next case (Suc n) - then show ?case by (simp add: algebra_simps del: one_add_one) (* FIXME *) + then show ?case + by (simp add: algebra_simps add: one_add_one [symmetric] del: one_add_one) + (* FIXME: make numeral cancellation simprocs work for semirings *) qed theorem arith_series_general: - "((1::'a::comm_semiring_1) + 1) * (\i\{..i\{.. 1" @@ -1307,26 +1309,27 @@ also from ngt1 have "\ = (?n*a + d*(\i\{1.. = (1+1)*?n*a + d*(1+1)*(\i\{1.. = 2*?n*a + d*2*(\i\{1..i\{1..n - 1}. ?I i) = ((1+1)*?n*a + d*?I (n - 1)*?I n)" + have "2*?n*a + d*2*(\i\{1..n - 1}. ?I i) = (2*?n*a + d*?I (n - 1)*?I n)" by (simp only: mult_ac gauss_sum [of "n - 1"], unfold One_nat_def) (simp add: mult_ac trans [OF add_commute of_nat_Suc [symmetric]]) - finally show ?thesis by (simp add: algebra_simps del: one_add_one) + finally show ?thesis + unfolding mult_2 by (simp add: algebra_simps) next assume "\(n > 1)" hence "n = 1 \ n = 0" by auto - thus ?thesis by (auto simp: algebra_simps mult_2_right) + thus ?thesis by (auto simp: mult_2) qed lemma arith_series_nat: - "Suc (Suc 0) * (\i\{..i\{..i\{..i\{..i\{..i\{..i\{..nat" diff -r 7205eb4a0a05 -r 1b7c909a6fad src/HOL/ex/Arithmetic_Series_Complex.thy --- a/src/HOL/ex/Arithmetic_Series_Complex.thy Fri Mar 30 14:00:18 2012 +0200 +++ b/src/HOL/ex/Arithmetic_Series_Complex.thy Fri Mar 30 14:25:32 2012 +0200 @@ -6,18 +6,12 @@ header {* Arithmetic Series for Reals *} theory Arithmetic_Series_Complex -imports Complex_Main +imports "../RealDef" begin lemma arith_series_real: "(2::real) * (\i\{..i\{..