doc-src/TutorialI/Misc/Tree.thy
author paulson
Fri, 26 May 2000 18:04:17 +0200
changeset 8983 15bd7d568fb2
parent 8745 13b32661dde4
child 9458 c613cd06d5cf
permissions -rw-r--r--
named the primrec clauses of upt

(*<*)
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 l) x (mirror r)";(*>*)

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);
apply(auto).;

end
(*>*)