doc-src/TutorialI/Misc/natsum.thy
author nipkow
Mon, 28 Aug 2000 09:32:51 +0200
changeset 9689 751fde5307e4
parent 9541 d17c0b34d5c8
child 9792 bbefb6ce5cb2
permissions -rw-r--r--
*** empty log message ***

(*<*)
theory natsum = Main:;
(*>*)
text{*\noindent
In particular, there are \isa{case}-expressions, for example
\begin{quote}
@{term[display]"case n of 0 => 0 | Suc m => m"}
\end{quote}
primitive recursion, for example
*}

consts sum :: "nat \\<Rightarrow> nat";
primrec "sum 0 = 0"
        "sum (Suc n) = Suc n + sum n";

text{*\noindent
and induction, for example
*}

lemma "sum n + sum n = n*(Suc n)";
apply(induct_tac n);
by(auto);

(*<*)
end
(*>*)