--- a/doc-src/IsarOverview/Isar/Induction.thy Wed Jun 22 07:54:01 2005 +0200
+++ b/doc-src/IsarOverview/Isar/Induction.thy Wed Jun 22 07:54:13 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 < n+1. i) = n*(n+1)"
+lemma "2 * (\<Sum>i::nat\<le>n. i) = n*(n+1)"
by (induct n, simp_all)
text{*\noindent The constraint @{text"::nat"} is needed because all of
@@ -86,7 +86,7 @@
If we want to expose more of the structure of the
proof, we can use pattern matching to avoid having to repeat the goal
statement: *}
-lemma "2 * (\<Sum>i::nat = 0..<n+1. i) = n*(n+1)" (is "?P n")
+lemma "2 * (\<Sum>i::nat\<le>n. i) = n*(n+1)" (is "?P n")
proof (induct n)
show "?P 0" by simp
next
@@ -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 < n+1. i) = n*(n+1)"
+lemma "2 * (\<Sum>i::nat \<le> n. i) = n*(n+1)"
proof (induct n)
case 0 show ?case by simp
next