(* Title: HOL/ex/AVL_completion2.thy ID: $Id: AVL.thy,v 1.1 2010/03/28 16:36:37 smaus Exp $ Author: Cornelia Pusch and Tobias Nipkow modified by B. Wolff Copyright 1998 TUM Adapted for Aquila 2010 course by Jan-Georg Smaus *) AVL = Main + datatype 'a tree = Leaf | Node 'a ('a tree) ('a tree) consts height :: "'a tree => nat" isin :: "'a => 'a tree => bool" isord :: "('a::order) tree => bool" isbal :: "'a tree => bool" isin_eff :: "('a::order) => 'a tree => bool" primrec height_empty "height Leaf = 0" height_branch "height (Node n l r) = Suc(max (height l) (height r))" primrec isin_empty "isin k Leaf = False" isin_branch "isin k (Node n l r) = (k=n | isin k l | isin k r)" primrec isord_empty "isord Leaf = True" isord_branch "isord (Node n l r) = ((! n'. isin n' l --> n' < n) & (! n'. isin n' r --> n < n') & isord l & isord r)" primrec isbal_empty "isbal Leaf = True" isbal_branch "isbal (Node n l r) = ((height l = height r | height l = Suc(height r) | Suc(height l) = height r) & isbal l & isbal r)" (*It was discovered by Christoph Sprunk that it is crucial for the automatic proofs that all above equations have the "l" part on the lhs and the "r" part on the rhs.*) primrec isin_eff_empty "isin_eff k Leaf = False" isin_eff_branch "isin_eff k (Node n l r) = (if k = n then True else (if k bal" "bal t == case t of Leaf => Just | (Node n l r) => if height l = height r then Just else if height l < height r then Right else Left" consts r_rot,l_rot,lr_rot,rl_rot :: "'a * 'a tree * 'a tree => 'a tree" recdef r_rot "{}" "r_rot (n, Node ln ll lr, r) = Node ln ll (Node n lr r)" recdef l_rot "{}" "l_rot(n, l, Node rn rl rr) = Node rn (Node n l rl) rr" recdef lr_rot "{}" "lr_rot(n, Node ln ll (Node lrn lrl lrr), r) = Node lrn (Node ln ll lrl) (Node n lrr r)" recdef rl_rot "{}" "rl_rot(n, l, Node rn (Node rln rll rlr) rr) = Node rln (Node n l rll) (Node rn rlr rr)" constdefs l_bal :: "'a => 'a tree => 'a tree => 'a tree" "l_bal n l r == if bal l = Right then lr_rot (n, l, r) else r_rot (n, l, r)" r_bal :: "'a => 'a tree => 'a tree => 'a tree" "r_bal n l r == if bal r = Left then rl_rot (n, l, r) else l_rot (n, l, r)" consts insert :: "'a::order => 'a tree => 'a tree" primrec "insert x Leaf = Node x Leaf Leaf" "insert x (Node n l r) = (if x=n then Node n l r else if x 'a tree" (*prune means: throw away the balancing labels*) correct :: "'a etree => bool" (*correct means: the balancing arguments are correct*) primrec prune_empty "prune ELeaf = Leaf" prune_branch "prune (ENode b n l r) = Node n (prune l) (prune r)" primrec correct_empty "correct ELeaf = True" correct_branch "correct (ENode b n l r) = (b = bal (Node n (prune l) (prune r)) & correct l & correct r)" consts label :: "'a tree => 'a etree" (*label means: add correct balancing labels*) primrec label_empty "label Leaf = ELeaf" label_branch "label (Node n l r) = ENode (bal (Node n l r)) n (label l) (label r)" end