doc-src/TutorialI/Misc/Plus.thy
changeset 19249 86c73395c99b
parent 16417 9bc16273c2d4
child 27015 f8537d69f514
--- a/doc-src/TutorialI/Misc/Plus.thy	Sat Mar 11 17:24:37 2006 +0100
+++ b/doc-src/TutorialI/Misc/Plus.thy	Sat Mar 11 17:30:35 2006 +0100
@@ -4,18 +4,18 @@
 
 text{*\noindent Define the following addition function *}
 
-consts plus :: "nat \<Rightarrow> nat \<Rightarrow> nat"
+consts add :: "nat \<Rightarrow> nat \<Rightarrow> nat"
 primrec
-"plus m 0 = m"
-"plus m (Suc n) = plus (Suc m) n"
+"add m 0 = m"
+"add m (Suc n) = add (Suc m) n"
 
 text{*\noindent and prove*}
 (*<*)
-lemma [simp]: "!m. plus m n = m+n"
+lemma [simp]: "!m. add m n = m+n"
 apply(induct_tac n)
 by(auto)
 (*>*)
-lemma "plus m n = m+n"
+lemma "add m n = m+n"
 (*<*)
 by(simp)