doc-src/TutorialI/Misc/Tree2.thy
changeset 9493 494f8cd34df7
child 9792 bbefb6ce5cb2
equal deleted inserted replaced
9492:72e429c66608 9493:494f8cd34df7
       
     1 (*<*)
       
     2 theory Tree2 = Tree:
       
     3 (*>*)
       
     4 
       
     5 text{*\noindent In Exercise~\ref{ex:Tree} we defined a function
       
     6 \isa{flatten} from trees to lists. The straightforward version of
       
     7 \isa{flatten} is based on \isa{\at} and is thus, like \isa{rev}, quadratic.
       
     8 A linear time version of \isa{flatten} again reqires an extra
       
     9 argument, the accumulator: *}
       
    10 
       
    11 consts flatten2 :: "'a tree => 'a list => 'a list"
       
    12 (*<*)
       
    13 primrec
       
    14 "flatten2 Tip xs = xs"
       
    15 "flatten2 (Node l x r) xs = flatten2 l (x#(flatten2 r xs))"
       
    16 (*>*)
       
    17 
       
    18 text{*\noindent Define \isa{flatten2} and prove
       
    19 *}
       
    20 (*<*)
       
    21 lemma [simp]: "!xs. flatten2 t xs = flatten t @ xs";
       
    22 apply(induct_tac t);
       
    23 by(auto);
       
    24 (*>*)
       
    25 lemma "flatten2 t [] = flatten t";
       
    26 (*<*)
       
    27 by(simp);
       
    28 
       
    29 end
       
    30 (*>*)