doc-src/TutorialI/Misc/Tree2.thy
changeset 9792 bbefb6ce5cb2
parent 9493 494f8cd34df7
child 13305 f88d0c363582
--- a/doc-src/TutorialI/Misc/Tree2.thy	Fri Sep 01 18:29:52 2000 +0200
+++ b/doc-src/TutorialI/Misc/Tree2.thy	Fri Sep 01 19:09:44 2000 +0200
@@ -3,9 +3,9 @@
 (*>*)
 
 text{*\noindent In Exercise~\ref{ex:Tree} we defined a function
-\isa{flatten} from trees to lists. The straightforward version of
-\isa{flatten} is based on \isa{\at} and is thus, like \isa{rev}, quadratic.
-A linear time version of \isa{flatten} again reqires an extra
+@{term"flatten"} from trees to lists. The straightforward version of
+@{term"flatten"} is based on @{text"@"} and is thus, like @{term"rev"},
+quadratic. A linear time version of @{term"flatten"} again reqires an extra
 argument, the accumulator: *}
 
 consts flatten2 :: "'a tree => 'a list => 'a list"
@@ -15,7 +15,7 @@
 "flatten2 (Node l x r) xs = flatten2 l (x#(flatten2 r xs))"
 (*>*)
 
-text{*\noindent Define \isa{flatten2} and prove
+text{*\noindent Define @{term"flatten2"} and prove
 *}
 (*<*)
 lemma [simp]: "!xs. flatten2 t xs = flatten t @ xs";