doc-src/TutorialI/Misc/Tree2.thy
changeset 9493 494f8cd34df7
child 9792 bbefb6ce5cb2
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/TutorialI/Misc/Tree2.thy	Wed Aug 02 11:30:38 2000 +0200
@@ -0,0 +1,30 @@
+(*<*)
+theory Tree2 = Tree:
+(*>*)
+
+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
+argument, the accumulator: *}
+
+consts flatten2 :: "'a tree => 'a list => 'a list"
+(*<*)
+primrec
+"flatten2 Tip xs = xs"
+"flatten2 (Node l x r) xs = flatten2 l (x#(flatten2 r xs))"
+(*>*)
+
+text{*\noindent Define \isa{flatten2} and prove
+*}
+(*<*)
+lemma [simp]: "!xs. flatten2 t xs = flatten t @ xs";
+apply(induct_tac t);
+by(auto);
+(*>*)
+lemma "flatten2 t [] = flatten t";
+(*<*)
+by(simp);
+
+end
+(*>*)