9493
|
1 |
(*<*)
|
16417
|
2 |
theory Tree2 imports Tree begin
|
9493
|
3 |
(*>*)
|
|
4 |
|
|
5 |
text{*\noindent In Exercise~\ref{ex:Tree} we defined a function
|
9792
|
6 |
@{term"flatten"} from trees to lists. The straightforward version of
|
|
7 |
@{term"flatten"} is based on @{text"@"} and is thus, like @{term"rev"},
|
|
8 |
quadratic. A linear time version of @{term"flatten"} again reqires an extra
|
27015
|
9 |
argument, the accumulator. Define *}
|
|
10 |
(*<*)primrec(*>*)flatten2 :: "'a tree \<Rightarrow> 'a list \<Rightarrow> 'a list"(*<*)where
|
|
11 |
"flatten2 Tip xs = xs" |
|
9493
|
12 |
"flatten2 (Node l x r) xs = flatten2 l (x#(flatten2 r xs))"
|
|
13 |
(*>*)
|
|
14 |
|
27015
|
15 |
text{*\noindent and prove*}
|
9493
|
16 |
(*<*)
|
27015
|
17 |
lemma [simp]: "!xs. flatten2 t xs = flatten t @ xs"
|
|
18 |
apply(induct_tac t)
|
9493
|
19 |
by(auto);
|
|
20 |
(*>*)
|
27015
|
21 |
lemma "flatten2 t [] = flatten t"
|
9493
|
22 |
(*<*)
|
27015
|
23 |
by(simp)
|
9493
|
24 |
|
|
25 |
end
|
|
26 |
(*>*)
|