remame ASeries to Arithmetic_Series
authorkleing
Fri, 07 Apr 2006 12:48:10 +0200
changeset 19358 9cd12369e753
parent 19357 dade85a75c9f
child 19359 5d523a1b6ddc
remame ASeries to Arithmetic_Series
src/HOL/Complex/ex/ASeries_Complex.thy
src/HOL/Complex/ex/Arithmetic_Series_Complex.thy
src/HOL/Complex/ex/ROOT.ML
--- a/src/HOL/Complex/ex/ASeries_Complex.thy	Fri Apr 07 11:17:44 2006 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,24 +0,0 @@
-(*  Title:      HOL/Library/ASeries.thy
-    ID:         $Id$
-    Author:     Benjamin Porter, 2006
-*)
-
-
-header {* Arithmetic Series for Reals *}
-
-theory ASeries_Complex
-imports Complex_Main ASeries
-begin
-
-lemma arith_series_real:
-  "(2::real) * (\<Sum>i\<in>{..<n}. a + of_nat i * d) =
-  of_nat n * (a + (a + of_nat(n - 1)*d))"
-proof -
-  have
-    "((1::real) + 1) * (\<Sum>i\<in>{..<n}. a + of_nat(i)*d) =
-    of_nat(n) * (a + (a + of_nat(n - 1)*d))"
-    by (rule arith_series_general)
-  thus ?thesis by simp
-qed
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Complex/ex/Arithmetic_Series_Complex.thy	Fri Apr 07 12:48:10 2006 +0200
@@ -0,0 +1,24 @@
+(*  Title:      HOL/Complex/ex/Arithmetic_Series_Complex
+    ID:         $Id$
+    Author:     Benjamin Porter, 2006
+*)
+
+
+header {* Arithmetic Series for Reals *}
+
+theory Arithmetic_Series_Complex
+imports Complex_Main Arithmetic_Series
+begin
+
+lemma arith_series_real:
+  "(2::real) * (\<Sum>i\<in>{..<n}. a + of_nat i * d) =
+  of_nat n * (a + (a + of_nat(n - 1)*d))"
+proof -
+  have
+    "((1::real) + 1) * (\<Sum>i\<in>{..<n}. a + of_nat(i)*d) =
+    of_nat(n) * (a + (a + of_nat(n - 1)*d))"
+    by (rule arith_series_general)
+  thus ?thesis by simp
+qed
+
+end
--- a/src/HOL/Complex/ex/ROOT.ML	Fri Apr 07 11:17:44 2006 +0200
+++ b/src/HOL/Complex/ex/ROOT.ML	Fri Apr 07 12:48:10 2006 +0200
@@ -15,8 +15,8 @@
 no_document use_thy "BigO";
 use_thy "BigO_Complex";
 
-no_document use_thy "ASeries";
-use_thy "ASeries_Complex";
+no_document use_thy "Arithmetic_Series";
+use_thy "Arithmetic_Series_Complex";
 use_thy "HarmonicSeries";
 
 no_document use_thy "NatPair";