doc-src/TutorialI/Misc/Tree.thy
changeset 8745 13b32661dde4
child 9458 c613cd06d5cf
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/TutorialI/Misc/Tree.thy	Wed Apr 19 11:56:31 2000 +0200
@@ -0,0 +1,27 @@
+(*<*)
+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
+(*>*)