diff -r caae34eabcef -r 90360390a916 src/HOL/Library/Tree.thy --- a/src/HOL/Library/Tree.thy Tue Sep 13 07:59:20 2016 +0200 +++ b/src/HOL/Library/Tree.thy Tue Sep 13 11:31:30 2016 +0200 @@ -23,6 +23,80 @@ definition size1 :: "'a tree \ nat" where "size1 t = size t + 1" +fun subtrees :: "'a tree \ 'a tree set" where +"subtrees \\ = {\\}" | +"subtrees (\l, a, r\) = insert \l, a, r\ (subtrees l \ subtrees r)" + +fun mirror :: "'a tree \ 'a tree" where +"mirror \\ = Leaf" | +"mirror \l,x,r\ = \mirror r, x, mirror l\" + +class height = fixes height :: "'a \ nat" + +instantiation tree :: (type)height +begin + +fun height_tree :: "'a tree => nat" where +"height Leaf = 0" | +"height (Node t1 a t2) = max (height t1) (height t2) + 1" + +instance .. + +end + +fun min_height :: "'a tree \ nat" where +"min_height Leaf = 0" | +"min_height (Node l _ r) = min (min_height l) (min_height r) + 1" + +fun complete :: "'a tree \ bool" where +"complete Leaf = True" | +"complete (Node l x r) = (complete l \ complete r \ height l = height r)" + +definition balanced :: "'a tree \ bool" where +"balanced t = (height t - min_height t \ 1)" + +text \Weight balanced:\ +fun wbalanced :: "'a tree \ bool" where +"wbalanced Leaf = True" | +"wbalanced (Node l x r) = (abs(int(size l) - int(size r)) \ 1 \ wbalanced l \ wbalanced r)" + +text \Internal path length:\ +fun path_len :: "'a tree \ nat" where +"path_len Leaf = 0 " | +"path_len (Node l _ r) = path_len l + size l + path_len r + size r" + +fun preorder :: "'a tree \ 'a list" where +"preorder \\ = []" | +"preorder \l, x, r\ = x # preorder l @ preorder r" + +fun inorder :: "'a tree \ 'a list" where +"inorder \\ = []" | +"inorder \l, x, r\ = inorder l @ [x] @ inorder r" + +text\A linear version avoiding append:\ +fun inorder2 :: "'a tree \ 'a list \ 'a list" where +"inorder2 \\ xs = xs" | +"inorder2 \l, x, r\ xs = inorder2 l (x # inorder2 r xs)" + +text\Binary Search Tree:\ +fun (in linorder) bst :: "'a tree \ bool" where +"bst \\ \ True" | +"bst \l, a, r\ \ bst l \ bst r \ (\x\set_tree l. x < a) \ (\x\set_tree r. a < x)" + +text\Binary Search Tree with duplicates:\ +fun (in linorder) bst_eq :: "'a tree \ bool" where +"bst_eq \\ \ True" | +"bst_eq \l,a,r\ \ + bst_eq l \ bst_eq r \ (\x\set_tree l. x \ a) \ (\x\set_tree r. a \ x)" + +fun (in linorder) heap :: "'a tree \ bool" where +"heap Leaf = True" | +"heap (Node l m r) = + (heap l \ heap r \ (\x \ set_tree l \ set_tree r. m \ x))" + + +subsection \@{const size}\ + lemma size1_simps[simp]: "size1 \\ = 1" "size1 \l, x, r\ = size1 l + size1 r" @@ -47,20 +121,19 @@ by (simp add: size1_def) -subsection "The Height" +subsection \@{const subtrees}\ -class height = fixes height :: "'a \ nat" - -instantiation tree :: (type)height -begin +lemma set_treeE: "a \ set_tree t \ \l r. \l, a, r\ \ subtrees t" +by (induction t)(auto) -fun height_tree :: "'a tree => nat" where -"height Leaf = 0" | -"height (Node t1 a t2) = max (height t1) (height t2) + 1" +lemma Node_notin_subtrees_if[simp]: "a \ set_tree t \ Node l a r \ subtrees t" +by (induction t) auto -instance .. +lemma in_set_tree_if: "\l, a, r\ \ subtrees t \ a \ set_tree t" +by (metis Node_notin_subtrees_if) -end + +subsection \@{const height} and @{const min_height}\ lemma height_0_iff_Leaf: "height t = 0 \ t = Leaf" by(cases t) auto @@ -92,10 +165,9 @@ corollary size_height: "size t \ 2 ^ height (t::'a tree) - 1" using size1_height[of t] by(arith) +lemma height_subtrees: "s \ subtrees t \ height s \ height t" +by (induction t) auto -fun min_height :: "'a tree \ nat" where -"min_height Leaf = 0" | -"min_height (Node l _ r) = min (min_height l) (min_height r) + 1" lemma min_hight_le_height: "min_height t \ height t" by(induction t) auto @@ -113,11 +185,7 @@ qed simp -subsection \Complete\ - -fun complete :: "'a tree \ bool" where -"complete Leaf = True" | -"complete (Node l x r) = (complete l \ complete r \ height l = height r)" +subsection \@{const complete}\ lemma complete_iff_height: "complete t \ (min_height t = height t)" apply(induction t) @@ -228,10 +296,18 @@ qed -subsection \Balanced\ +subsection \@{const balanced}\ + +lemma balanced_subtreeL: "balanced (Node l x r) \ balanced l" +by(simp add: balanced_def) -definition balanced :: "'a tree \ bool" where -"balanced t = (height t - min_height t \ 1)" +lemma balanced_subtreeR: "balanced (Node l x r) \ balanced r" +by(simp add: balanced_def) + +lemma balanced_subtrees: "\ balanced t; s \ subtrees t \ \ balanced s" +using [[simp_depth_limit=1]] +by(induction t arbitrary: s) + (auto simp add: balanced_subtreeL balanced_subtreeR) text\Balanced trees have optimal height:\ @@ -268,14 +344,18 @@ qed -subsection \Path length\ +subsection \@{const wbalanced}\ + +lemma wbalanced_subtrees: "\ wbalanced t; s \ subtrees t \ \ wbalanced s" +using [[simp_depth_limit=1]] by(induction t arbitrary: s) auto + +(* show wbalanced \ balanced and use that in Balanced.thy *) + + +subsection \@{const path_len}\ text \The internal path length of a tree:\ -fun path_len :: "'a tree \ nat" where -"path_len Leaf = 0 " | -"path_len (Node l _ r) = path_len l + size l + path_len r + size r" - lemma path_len_if_bal: "complete t \ path_len t = (let n = height t in 2 + n*2^n - 2^(n+1))" proof(induction t) @@ -288,38 +368,8 @@ qed simp -subsection "The set of subtrees" - -fun subtrees :: "'a tree \ 'a tree set" where -"subtrees \\ = {\\}" | -"subtrees (\l, a, r\) = insert \l, a, r\ (subtrees l \ subtrees r)" - -lemma set_treeE: "a \ set_tree t \ \l r. \l, a, r\ \ subtrees t" -by (induction t)(auto) - -lemma Node_notin_subtrees_if[simp]: "a \ set_tree t \ Node l a r \ subtrees t" -by (induction t) auto - -lemma in_set_tree_if: "\l, a, r\ \ subtrees t \ a \ set_tree t" -by (metis Node_notin_subtrees_if) - - subsection "List of entries" -fun preorder :: "'a tree \ 'a list" where -"preorder \\ = []" | -"preorder \l, x, r\ = x # preorder l @ preorder r" - -fun inorder :: "'a tree \ 'a list" where -"inorder \\ = []" | -"inorder \l, x, r\ = inorder l @ [x] @ inorder r" - -text\A linear version avoiding append:\ -fun inorder2 :: "'a tree \ 'a list \ 'a list" where -"inorder2 \\ xs = xs" | -"inorder2 \l, x, r\ xs = inorder2 l (x # inorder2 r xs)" - - lemma set_inorder[simp]: "set (inorder t) = set_tree t" by (induction t) auto @@ -342,18 +392,7 @@ by (induction t arbitrary: xs) auto -subsection \Binary Search Tree predicate\ - -fun (in linorder) bst :: "'a tree \ bool" where -"bst \\ \ True" | -"bst \l, a, r\ \ bst l \ bst r \ (\x\set_tree l. x < a) \ (\x\set_tree r. a < x)" - -text\In case there are duplicates:\ - -fun (in linorder) bst_eq :: "'a tree \ bool" where -"bst_eq \\ \ True" | -"bst_eq \l,a,r\ \ - bst_eq l \ bst_eq r \ (\x\set_tree l. x \ a) \ (\x\set_tree r. a \ x)" +subsection \Binary Search Tree\ lemma (in linorder) bst_eq_if_bst: "bst t \ bst_eq t" by (induction t) (auto) @@ -376,19 +415,10 @@ done -subsection "The heap predicate" - -fun heap :: "'a::linorder tree \ bool" where -"heap Leaf = True" | -"heap (Node l m r) = - (heap l \ heap r \ (\x \ set_tree l \ set_tree r. m \ x))" +subsection \@{const heap}\ -subsection "Function \mirror\" - -fun mirror :: "'a tree \ 'a tree" where -"mirror \\ = Leaf" | -"mirror \l,x,r\ = \mirror r, x, mirror l\" +subsection \@{const mirror}\ lemma mirror_Leaf[simp]: "mirror t = \\ \ t = \\" by (induction t) simp_all