| 8745 |      1 | (*<*)
 | 
|  |      2 | theory Tree = Main:
 | 
|  |      3 | (*>*)
 | 
|  |      4 | 
 | 
|  |      5 | text{*\noindent
 | 
| 11456 |      6 | Define the datatype of \rmindex{binary trees}:
 | 
| 8745 |      7 | *}
 | 
|  |      8 | 
 | 
|  |      9 | datatype 'a tree = Tip | Node "'a tree" 'a "'a tree";(*<*)
 | 
|  |     10 | 
 | 
| 10795 |     11 | consts mirror :: "'a tree \<Rightarrow> 'a tree";
 | 
| 8745 |     12 | primrec
 | 
|  |     13 | "mirror Tip = Tip"
 | 
| 9493 |     14 | "mirror (Node l x r) = Node (mirror r) x (mirror l)";(*>*)
 | 
| 8745 |     15 | 
 | 
|  |     16 | text{*\noindent
 | 
| 11456 |     17 | Define a function @{term"mirror"} that mirrors a binary tree
 | 
| 10795 |     18 | by swapping subtrees recursively. Prove
 | 
| 8745 |     19 | *}
 | 
|  |     20 | 
 | 
|  |     21 | lemma mirror_mirror: "mirror(mirror t) = t";
 | 
|  |     22 | (*<*)
 | 
|  |     23 | apply(induct_tac t);
 | 
| 9458 |     24 | by(auto);
 | 
| 8745 |     25 | 
 | 
| 9493 |     26 | consts flatten :: "'a tree => 'a list"
 | 
|  |     27 | primrec
 | 
|  |     28 | "flatten Tip = []"
 | 
|  |     29 | "flatten (Node l x r) = flatten l @ [x] @ flatten r";
 | 
|  |     30 | (*>*)
 | 
|  |     31 | 
 | 
|  |     32 | text{*\noindent
 | 
| 9792 |     33 | Define a function @{term"flatten"} that flattens a tree into a list
 | 
| 9493 |     34 | by traversing it in infix order. Prove
 | 
|  |     35 | *}
 | 
|  |     36 | 
 | 
|  |     37 | lemma "flatten(mirror t) = rev(flatten t)";
 | 
|  |     38 | (*<*)
 | 
|  |     39 | apply(induct_tac t);
 | 
|  |     40 | by(auto);
 | 
|  |     41 | 
 | 
| 8745 |     42 | end
 | 
|  |     43 | (*>*)
 |