diff -r 4dc65845eab3 -r d8d7d1b785af doc-src/TutorialI/Advanced/Partial.thy --- a/doc-src/TutorialI/Advanced/Partial.thy Wed Feb 24 11:55:52 2010 +0100 +++ b/doc-src/TutorialI/Advanced/Partial.thy Mon Mar 01 13:40:23 2010 +0100 @@ -34,7 +34,7 @@ preconditions: *} -constdefs subtract :: "nat \ nat \ nat" +definition subtract :: "nat \ nat \ nat" where "n \ m \ subtract m n \ m - n" text{* @@ -85,7 +85,7 @@ Phrased differently, the relation *} -constdefs step1 :: "('a \ 'a) \ ('a \ 'a)set" +definition step1 :: "('a \ 'a) \ ('a \ 'a)set" where "step1 f \ {(y,x). y = f x \ y \ x}" text{*\noindent @@ -160,7 +160,7 @@ consider the following definition of function @{const find}: *} -constdefs find2 :: "('a \ 'a) \ 'a \ 'a" +definition find2 :: "('a \ 'a) \ 'a \ 'a" where "find2 f x \ fst(while (\(x,x'). x' \ x) (\(x,x'). (x',f x')) (x,f x))"