| author | wenzelm | 
| Tue, 08 Mar 2016 18:15:16 +0100 | |
| changeset 62559 | 83e815849a91 | 
| parent 62202 | e5bc7cbb0bcc | 
| child 62650 | 7e6bb43e7217 | 
| permissions | -rw-r--r-- | 
| 57250 | 1 | (* Author: Tobias Nipkow *) | 
| 2 | ||
| 60500 | 3 | section \<open>Binary Tree\<close> | 
| 57250 | 4 | |
| 5 | theory Tree | |
| 6 | imports Main | |
| 7 | begin | |
| 8 | ||
| 58424 | 9 | datatype 'a tree = | 
| 62160 | 10 |   is_Leaf: Leaf ("\<langle>\<rangle>") |
 | 
| 11 |   Node (left: "'a tree") (val: 'a) (right: "'a tree") ("(1\<langle>_,/ _,/ _\<rangle>)")
 | |
| 57449 
f81da03b9ebd
Library/Tree: use datatype_new, bst is an inductive predicate
 hoelzl parents: 
57250diff
changeset | 12 | where | 
| 
f81da03b9ebd
Library/Tree: use datatype_new, bst is an inductive predicate
 hoelzl parents: 
57250diff
changeset | 13 | "left Leaf = Leaf" | 
| 
f81da03b9ebd
Library/Tree: use datatype_new, bst is an inductive predicate
 hoelzl parents: 
57250diff
changeset | 14 | | "right Leaf = Leaf" | 
| 57569 
e20a999f7161
register tree with datatype_compat ot support QuickCheck
 hoelzl parents: 
57530diff
changeset | 15 | datatype_compat tree | 
| 57250 | 16 | |
| 60500 | 17 | text\<open>Can be seen as counting the number of leaves rather than nodes:\<close> | 
| 58438 | 18 | |
| 19 | definition size1 :: "'a tree \<Rightarrow> nat" where | |
| 20 | "size1 t = size t + 1" | |
| 21 | ||
| 22 | lemma size1_simps[simp]: | |
| 23 | "size1 \<langle>\<rangle> = 1" | |
| 24 | "size1 \<langle>l, x, r\<rangle> = size1 l + size1 r" | |
| 25 | by (simp_all add: size1_def) | |
| 26 | ||
| 60507 | 27 | lemma size_0_iff_Leaf: "size t = 0 \<longleftrightarrow> t = Leaf" | 
| 60505 | 28 | by(cases t) auto | 
| 29 | ||
| 58424 | 30 | lemma neq_Leaf_iff: "(t \<noteq> \<langle>\<rangle>) = (\<exists>l a r. t = \<langle>l, a, r\<rangle>)" | 
| 31 | by (cases t) auto | |
| 57530 | 32 | |
| 57687 | 33 | lemma finite_set_tree[simp]: "finite(set_tree t)" | 
| 34 | by(induction t) auto | |
| 35 | ||
| 59776 | 36 | lemma size_map_tree[simp]: "size (map_tree f t) = size t" | 
| 37 | by (induction t) auto | |
| 38 | ||
| 39 | lemma size1_map_tree[simp]: "size1 (map_tree f t) = size1 t" | |
| 40 | by (simp add: size1_def) | |
| 41 | ||
| 42 | ||
| 60808 
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
 nipkow parents: 
60507diff
changeset | 43 | subsection "The Height" | 
| 
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
 nipkow parents: 
60507diff
changeset | 44 | |
| 
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
 nipkow parents: 
60507diff
changeset | 45 | class height = fixes height :: "'a \<Rightarrow> nat" | 
| 
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
 nipkow parents: 
60507diff
changeset | 46 | |
| 
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
 nipkow parents: 
60507diff
changeset | 47 | instantiation tree :: (type)height | 
| 
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
 nipkow parents: 
60507diff
changeset | 48 | begin | 
| 59776 | 49 | |
| 60808 
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
 nipkow parents: 
60507diff
changeset | 50 | fun height_tree :: "'a tree => nat" where | 
| 
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
 nipkow parents: 
60507diff
changeset | 51 | "height Leaf = 0" | | 
| 
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
 nipkow parents: 
60507diff
changeset | 52 | "height (Node t1 a t2) = max (height t1) (height t2) + 1" | 
| 59776 | 53 | |
| 60808 
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
 nipkow parents: 
60507diff
changeset | 54 | instance .. | 
| 
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
 nipkow parents: 
60507diff
changeset | 55 | |
| 
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
 nipkow parents: 
60507diff
changeset | 56 | end | 
| 
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
 nipkow parents: 
60507diff
changeset | 57 | |
| 
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
 nipkow parents: 
60507diff
changeset | 58 | lemma height_map_tree[simp]: "height (map_tree f t) = height t" | 
| 59776 | 59 | by (induction t) auto | 
| 60 | ||
| 62202 | 61 | lemma size1_height: "size t + 1 \<le> 2 ^ height (t::'a tree)" | 
| 62 | proof(induction t) | |
| 63 | case (Node l a r) | |
| 64 | show ?case | |
| 65 | proof (cases "height l \<le> height r") | |
| 66 | case True | |
| 67 | have "size(Node l a r) + 1 = (size l + 1) + (size r + 1)" by simp | |
| 68 | also have "size l + 1 \<le> 2 ^ height l" by(rule Node.IH(1)) | |
| 69 | also have "size r + 1 \<le> 2 ^ height r" by(rule Node.IH(2)) | |
| 70 | also have "(2::nat) ^ height l \<le> 2 ^ height r" using True by simp | |
| 71 | finally show ?thesis using True by (auto simp: max_def mult_2) | |
| 72 | next | |
| 73 | case False | |
| 74 | have "size(Node l a r) + 1 = (size l + 1) + (size r + 1)" by simp | |
| 75 | also have "size l + 1 \<le> 2 ^ height l" by(rule Node.IH(1)) | |
| 76 | also have "size r + 1 \<le> 2 ^ height r" by(rule Node.IH(2)) | |
| 77 | also have "(2::nat) ^ height r \<le> 2 ^ height l" using False by simp | |
| 78 | finally show ?thesis using False by (auto simp: max_def mult_2) | |
| 79 | qed | |
| 80 | qed simp | |
| 81 | ||
| 57687 | 82 | |
| 83 | subsection "The set of subtrees" | |
| 84 | ||
| 57250 | 85 | fun subtrees :: "'a tree \<Rightarrow> 'a tree set" where | 
| 60808 
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
 nipkow parents: 
60507diff
changeset | 86 | "subtrees \<langle>\<rangle> = {\<langle>\<rangle>}" |
 | 
| 
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
 nipkow parents: 
60507diff
changeset | 87 | "subtrees (\<langle>l, a, r\<rangle>) = insert \<langle>l, a, r\<rangle> (subtrees l \<union> subtrees r)" | 
| 57250 | 88 | |
| 58424 | 89 | lemma set_treeE: "a \<in> set_tree t \<Longrightarrow> \<exists>l r. \<langle>l, a, r\<rangle> \<in> subtrees t" | 
| 90 | by (induction t)(auto) | |
| 57449 
f81da03b9ebd
Library/Tree: use datatype_new, bst is an inductive predicate
 hoelzl parents: 
57250diff
changeset | 91 | |
| 57450 | 92 | lemma Node_notin_subtrees_if[simp]: "a \<notin> set_tree t \<Longrightarrow> Node l a r \<notin> subtrees t" | 
| 58424 | 93 | by (induction t) auto | 
| 57449 
f81da03b9ebd
Library/Tree: use datatype_new, bst is an inductive predicate
 hoelzl parents: 
57250diff
changeset | 94 | |
| 58424 | 95 | lemma in_set_tree_if: "\<langle>l, a, r\<rangle> \<in> subtrees t \<Longrightarrow> a \<in> set_tree t" | 
| 96 | by (metis Node_notin_subtrees_if) | |
| 57449 
f81da03b9ebd
Library/Tree: use datatype_new, bst is an inductive predicate
 hoelzl parents: 
57250diff
changeset | 97 | |
| 57687 | 98 | |
| 59776 | 99 | subsection "List of entries" | 
| 100 | ||
| 101 | fun preorder :: "'a tree \<Rightarrow> 'a list" where | |
| 102 | "preorder \<langle>\<rangle> = []" | | |
| 103 | "preorder \<langle>l, x, r\<rangle> = x # preorder l @ preorder r" | |
| 57687 | 104 | |
| 57250 | 105 | fun inorder :: "'a tree \<Rightarrow> 'a list" where | 
| 58424 | 106 | "inorder \<langle>\<rangle> = []" | | 
| 107 | "inorder \<langle>l, x, r\<rangle> = inorder l @ [x] @ inorder r" | |
| 57250 | 108 | |
| 57449 
f81da03b9ebd
Library/Tree: use datatype_new, bst is an inductive predicate
 hoelzl parents: 
57250diff
changeset | 109 | lemma set_inorder[simp]: "set (inorder t) = set_tree t" | 
| 58424 | 110 | by (induction t) auto | 
| 57250 | 111 | |
| 59776 | 112 | lemma set_preorder[simp]: "set (preorder t) = set_tree t" | 
| 113 | by (induction t) auto | |
| 114 | ||
| 115 | lemma length_preorder[simp]: "length (preorder t) = size t" | |
| 116 | by (induction t) auto | |
| 117 | ||
| 118 | lemma length_inorder[simp]: "length (inorder t) = size t" | |
| 119 | by (induction t) auto | |
| 120 | ||
| 121 | lemma preorder_map: "preorder (map_tree f t) = map f (preorder t)" | |
| 122 | by (induction t) auto | |
| 123 | ||
| 124 | lemma inorder_map: "inorder (map_tree f t) = map f (inorder t)" | |
| 125 | by (induction t) auto | |
| 126 | ||
| 57687 | 127 | |
| 60500 | 128 | subsection \<open>Binary Search Tree predicate\<close> | 
| 57250 | 129 | |
| 57450 | 130 | fun (in linorder) bst :: "'a tree \<Rightarrow> bool" where | 
| 58424 | 131 | "bst \<langle>\<rangle> \<longleftrightarrow> True" | | 
| 132 | "bst \<langle>l, a, r\<rangle> \<longleftrightarrow> bst l \<and> bst r \<and> (\<forall>x\<in>set_tree l. x < a) \<and> (\<forall>x\<in>set_tree r. a < x)" | |
| 57250 | 133 | |
| 60500 | 134 | text\<open>In case there are duplicates:\<close> | 
| 59561 | 135 | |
| 136 | fun (in linorder) bst_eq :: "'a tree \<Rightarrow> bool" where | |
| 137 | "bst_eq \<langle>\<rangle> \<longleftrightarrow> True" | | |
| 138 | "bst_eq \<langle>l,a,r\<rangle> \<longleftrightarrow> | |
| 139 | bst_eq l \<and> bst_eq r \<and> (\<forall>x\<in>set_tree l. x \<le> a) \<and> (\<forall>x\<in>set_tree r. a \<le> x)" | |
| 140 | ||
| 59928 | 141 | lemma (in linorder) bst_eq_if_bst: "bst t \<Longrightarrow> bst_eq t" | 
| 142 | by (induction t) (auto) | |
| 143 | ||
| 59561 | 144 | lemma (in linorder) bst_eq_imp_sorted: "bst_eq t \<Longrightarrow> sorted (inorder t)" | 
| 145 | apply (induction t) | |
| 146 | apply(simp) | |
| 147 | by (fastforce simp: sorted_append sorted_Cons intro: less_imp_le less_trans) | |
| 148 | ||
| 59928 | 149 | lemma (in linorder) distinct_preorder_if_bst: "bst t \<Longrightarrow> distinct (preorder t)" | 
| 150 | apply (induction t) | |
| 151 | apply simp | |
| 152 | apply(fastforce elim: order.asym) | |
| 153 | done | |
| 154 | ||
| 155 | lemma (in linorder) distinct_inorder_if_bst: "bst t \<Longrightarrow> distinct (inorder t)" | |
| 156 | apply (induction t) | |
| 157 | apply simp | |
| 158 | apply(fastforce elim: order.asym) | |
| 159 | done | |
| 160 | ||
| 59776 | 161 | |
| 60505 | 162 | subsection "The heap predicate" | 
| 163 | ||
| 164 | fun heap :: "'a::linorder tree \<Rightarrow> bool" where | |
| 165 | "heap Leaf = True" | | |
| 166 | "heap (Node l m r) = | |
| 167 | (heap l \<and> heap r \<and> (\<forall>x \<in> set_tree l \<union> set_tree r. m \<le> x))" | |
| 168 | ||
| 169 | ||
| 61585 | 170 | subsection "Function \<open>mirror\<close>" | 
| 59561 | 171 | |
| 172 | fun mirror :: "'a tree \<Rightarrow> 'a tree" where | |
| 173 | "mirror \<langle>\<rangle> = Leaf" | | |
| 174 | "mirror \<langle>l,x,r\<rangle> = \<langle>mirror r, x, mirror l\<rangle>" | |
| 175 | ||
| 176 | lemma mirror_Leaf[simp]: "mirror t = \<langle>\<rangle> \<longleftrightarrow> t = \<langle>\<rangle>" | |
| 177 | by (induction t) simp_all | |
| 178 | ||
| 179 | lemma size_mirror[simp]: "size(mirror t) = size t" | |
| 180 | by (induction t) simp_all | |
| 181 | ||
| 182 | lemma size1_mirror[simp]: "size1(mirror t) = size1 t" | |
| 183 | by (simp add: size1_def) | |
| 184 | ||
| 60808 
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
 nipkow parents: 
60507diff
changeset | 185 | lemma height_mirror[simp]: "height(mirror t) = height t" | 
| 59776 | 186 | by (induction t) simp_all | 
| 187 | ||
| 188 | lemma inorder_mirror: "inorder(mirror t) = rev(inorder t)" | |
| 189 | by (induction t) simp_all | |
| 190 | ||
| 191 | lemma map_mirror: "map_tree f (mirror t) = mirror (map_tree f t)" | |
| 192 | by (induction t) simp_all | |
| 193 | ||
| 59561 | 194 | lemma mirror_mirror[simp]: "mirror(mirror t) = t" | 
| 195 | by (induction t) simp_all | |
| 196 | ||
| 57250 | 197 | end |