doc-src/IsarOverview/Isar/Induction.thy
changeset 16044 6887e6d12a94
parent 15909 5f0c8a3f0226
child 16522 f718767efd49
--- a/doc-src/IsarOverview/Isar/Induction.thy	Mon May 23 12:09:30 2005 +0200
+++ b/doc-src/IsarOverview/Isar/Induction.thy	Mon May 23 13:39:45 2005 +0200
@@ -77,7 +77,7 @@
 subsection{*Structural induction*}
 
 text{* We start with an inductive proof where both cases are proved automatically: *}
-lemma "2 * (\<Sum>i::nat = 0..<n+1. i) = n*(n+1)"
+lemma "2 * (\<Sum>i::nat < n+1. i) = n*(n+1)"
 by (induct n, simp_all)
 
 text{*\noindent The constraint @{text"::nat"} is needed because all of
@@ -97,7 +97,7 @@
 text{* \noindent We could refine this further to show more of the equational
 proof. Instead we explore the same avenue as for case distinctions:
 introducing context via the \isakeyword{case} command: *}
-lemma "2 * (\<Sum>i::nat = 0..<n+1. i) = n*(n+1)"
+lemma "2 * (\<Sum>i::nat < n+1. i) = n*(n+1)"
 proof (induct n)
   case 0 show ?case by simp
 next