doc-src/TutorialI/Misc/Tree.thy
author nipkow
Wed, 30 Aug 2000 14:38:48 +0200
changeset 9742 98d3ca2c18f7
parent 9493 494f8cd34df7
child 9792 bbefb6ce5cb2
permissions -rw-r--r--
*** empty log message ***

(*<*)
theory Tree = Main:
(*>*)

text{*\noindent
Define the datatype of binary trees
*}

datatype 'a tree = Tip | Node "'a tree" 'a "'a tree";(*<*)

consts mirror :: "'a tree \\<Rightarrow> 'a tree";
primrec
"mirror Tip = Tip"
"mirror (Node l x r) = Node (mirror r) x (mirror l)";(*>*)

text{*\noindent
and a function \isa{mirror} that mirrors a binary tree
by swapping subtrees (recursively). Prove
*}

lemma mirror_mirror: "mirror(mirror t) = t";
(*<*)
apply(induct_tac t);
by(auto);

consts flatten :: "'a tree => 'a list"
primrec
"flatten Tip = []"
"flatten (Node l x r) = flatten l @ [x] @ flatten r";
(*>*)

text{*\noindent
Define a function \isa{flatten} that flattens a tree into a list
by traversing it in infix order. Prove
*}

lemma "flatten(mirror t) = rev(flatten t)";
(*<*)
apply(induct_tac t);
by(auto);

end
(*>*)