doc-src/TutorialI/Misc/Tree.thy
author krauss
Tue, 02 Aug 2011 11:52:57 +0200
changeset 44014 88bd7d74a2c1
parent 27015 f8537d69f514
permissions -rw-r--r--
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef

(*<*)
theory Tree imports Main begin
(*>*)

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

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

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

text{*\noindent
Define a function @{term"mirror"} that mirrors a binary tree
by swapping subtrees recursively. Prove
*}

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

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

text{*\noindent
Define a function @{term"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
(*>*)