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