author | nipkow |
Fri, 02 Jun 2000 15:19:18 +0200 | |
changeset 9016 | d61c76716984 |
parent 8745 | 13b32661dde4 |
child 9458 | c613cd06d5cf |
permissions | -rw-r--r-- |
(*<*) 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 (*>*)