| author | wenzelm | 
| Fri, 17 May 2013 17:45:51 +0200 | |
| changeset 52051 | 9362fcd0318c | 
| parent 49962 | a8cc904a6820 | 
| child 58249 | 180f1b3508ed | 
| permissions | -rw-r--r-- | 
| 1476 | 1 | (* Title: HOL/ex/BT.thy | 
| 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 1167 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 3 | Copyright 1995 University of Cambridge | 
| 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 4 | |
| 20711 | 5 | Binary trees | 
| 1167 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 6 | *) | 
| 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 7 | |
| 11024 | 8 | header {* Binary trees *}
 | 
| 9 | ||
| 16417 | 10 | theory BT imports Main begin | 
| 1167 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 11 | |
| 11024 | 12 | datatype 'a bt = | 
| 13 | Lf | |
| 14 | | Br 'a "'a bt" "'a bt" | |
| 15 | ||
| 39246 | 16 | primrec n_nodes :: "'a bt => nat" where | 
| 17 | "n_nodes Lf = 0" | |
| 18 | | "n_nodes (Br a t1 t2) = Suc (n_nodes t1 + n_nodes t2)" | |
| 19 | ||
| 20 | primrec n_leaves :: "'a bt => nat" where | |
| 21 | "n_leaves Lf = Suc 0" | |
| 22 | | "n_leaves (Br a t1 t2) = n_leaves t1 + n_leaves t2" | |
| 1167 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 23 | |
| 39246 | 24 | primrec depth :: "'a bt => nat" where | 
| 25 | "depth Lf = 0" | |
| 26 | | "depth (Br a t1 t2) = Suc (max (depth t1) (depth t2))" | |
| 1167 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 27 | |
| 39246 | 28 | primrec reflect :: "'a bt => 'a bt" where | 
| 29 | "reflect Lf = Lf" | |
| 30 | | "reflect (Br a t1 t2) = Br a (reflect t2) (reflect t1)" | |
| 19478 | 31 | |
| 39246 | 32 | primrec bt_map :: "('a => 'b) => ('a bt => 'b bt)" where
 | 
| 1896 | 33 | "bt_map f Lf = Lf" | 
| 39246 | 34 | | "bt_map f (Br a t1 t2) = Br (f a) (bt_map f t1) (bt_map f t2)" | 
| 1167 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 35 | |
| 39246 | 36 | primrec preorder :: "'a bt => 'a list" where | 
| 19478 | 37 | "preorder Lf = []" | 
| 39246 | 38 | | "preorder (Br a t1 t2) = [a] @ (preorder t1) @ (preorder t2)" | 
| 1167 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 39 | |
| 39246 | 40 | primrec inorder :: "'a bt => 'a list" where | 
| 19478 | 41 | "inorder Lf = []" | 
| 39246 | 42 | | "inorder (Br a t1 t2) = (inorder t1) @ [a] @ (inorder t2)" | 
| 1167 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 43 | |
| 39246 | 44 | primrec postorder :: "'a bt => 'a list" where | 
| 19478 | 45 | "postorder Lf = []" | 
| 39246 | 46 | | "postorder (Br a t1 t2) = (postorder t1) @ (postorder t2) @ [a]" | 
| 1167 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 47 | |
| 39246 | 48 | primrec append :: "'a bt => 'a bt => 'a bt" where | 
| 23236 | 49 | "append Lf t = t" | 
| 39246 | 50 | | "append (Br a t1 t2) t = Br a (append t1 t) (append t2 t)" | 
| 11024 | 51 | |
| 52 | text {* \medskip BT simplification *}
 | |
| 53 | ||
| 54 | lemma n_leaves_reflect: "n_leaves (reflect t) = n_leaves t" | |
| 55 | apply (induct t) | |
| 56 | apply auto | |
| 57 | done | |
| 58 | ||
| 59 | lemma n_nodes_reflect: "n_nodes (reflect t) = n_nodes t" | |
| 60 | apply (induct t) | |
| 61 | apply auto | |
| 62 | done | |
| 63 | ||
| 19478 | 64 | lemma depth_reflect: "depth (reflect t) = depth t" | 
| 20711 | 65 | apply (induct t) | 
| 66 | apply auto | |
| 19478 | 67 | done | 
| 68 | ||
| 11024 | 69 | text {*
 | 
| 70 | The famous relationship between the numbers of leaves and nodes. | |
| 71 | *} | |
| 72 | ||
| 73 | lemma n_leaves_nodes: "n_leaves t = Suc (n_nodes t)" | |
| 74 | apply (induct t) | |
| 75 | apply auto | |
| 76 | done | |
| 77 | ||
| 78 | lemma reflect_reflect_ident: "reflect (reflect t) = t" | |
| 79 | apply (induct t) | |
| 80 | apply auto | |
| 81 | done | |
| 82 | ||
| 83 | lemma bt_map_reflect: "bt_map f (reflect t) = reflect (bt_map f t)" | |
| 84 | apply (induct t) | |
| 85 | apply simp_all | |
| 86 | done | |
| 87 | ||
| 19526 | 88 | lemma preorder_bt_map: "preorder (bt_map f t) = map f (preorder t)" | 
| 89 | apply (induct t) | |
| 90 | apply simp_all | |
| 91 | done | |
| 92 | ||
| 11024 | 93 | lemma inorder_bt_map: "inorder (bt_map f t) = map f (inorder t)" | 
| 94 | apply (induct t) | |
| 95 | apply simp_all | |
| 96 | done | |
| 97 | ||
| 19526 | 98 | lemma postorder_bt_map: "postorder (bt_map f t) = map f (postorder t)" | 
| 99 | apply (induct t) | |
| 100 | apply simp_all | |
| 101 | done | |
| 102 | ||
| 103 | lemma depth_bt_map [simp]: "depth (bt_map f t) = depth t" | |
| 104 | apply (induct t) | |
| 105 | apply simp_all | |
| 106 | done | |
| 107 | ||
| 108 | lemma n_leaves_bt_map [simp]: "n_leaves (bt_map f t) = n_leaves t" | |
| 109 | apply (induct t) | |
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
39246diff
changeset | 110 | apply (simp_all add: distrib_right) | 
| 19526 | 111 | done | 
| 112 | ||
| 11024 | 113 | lemma preorder_reflect: "preorder (reflect t) = rev (postorder t)" | 
| 114 | apply (induct t) | |
| 115 | apply simp_all | |
| 116 | done | |
| 117 | ||
| 118 | lemma inorder_reflect: "inorder (reflect t) = rev (inorder t)" | |
| 119 | apply (induct t) | |
| 120 | apply simp_all | |
| 121 | done | |
| 122 | ||
| 123 | lemma postorder_reflect: "postorder (reflect t) = rev (preorder t)" | |
| 124 | apply (induct t) | |
| 125 | apply simp_all | |
| 126 | done | |
| 127 | ||
| 19478 | 128 | text {*
 | 
| 129 | Analogues of the standard properties of the append function for lists. | |
| 130 | *} | |
| 131 | ||
| 132 | lemma append_assoc [simp]: | |
| 23236 | 133 | "append (append t1 t2) t3 = append t1 (append t2 t3)" | 
| 19478 | 134 | apply (induct t1) | 
| 135 | apply simp_all | |
| 136 | done | |
| 137 | ||
| 23236 | 138 | lemma append_Lf2 [simp]: "append t Lf = t" | 
| 19478 | 139 | apply (induct t) | 
| 140 | apply simp_all | |
| 141 | done | |
| 142 | ||
| 23236 | 143 | lemma depth_append [simp]: "depth (append t1 t2) = depth t1 + depth t2" | 
| 19478 | 144 | apply (induct t1) | 
| 145 | apply (simp_all add: max_add_distrib_left) | |
| 146 | done | |
| 147 | ||
| 148 | lemma n_leaves_append [simp]: | |
| 23236 | 149 | "n_leaves (append t1 t2) = n_leaves t1 * n_leaves t2" | 
| 19478 | 150 | apply (induct t1) | 
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
39246diff
changeset | 151 | apply (simp_all add: distrib_right) | 
| 19478 | 152 | done | 
| 153 | ||
| 19526 | 154 | lemma bt_map_append: | 
| 23236 | 155 | "bt_map f (append t1 t2) = append (bt_map f t1) (bt_map f t2)" | 
| 19526 | 156 | apply (induct t1) | 
| 157 | apply simp_all | |
| 158 | done | |
| 159 | ||
| 1167 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 160 | end |