author | nipkow |
Tue, 17 Jun 2025 06:29:55 +0200 | |
changeset 82732 | 71574900b6ba |
parent 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
61640 | 1 |
(* Author: Tobias Nipkow *) |
2 |
||
3 |
section \<open>2-3 Trees\<close> |
|
4 |
||
5 |
theory Tree23 |
|
6 |
imports Main |
|
7 |
begin |
|
8 |
||
9 |
class height = |
|
10 |
fixes height :: "'a \<Rightarrow> nat" |
|
11 |
||
12 |
datatype 'a tree23 = |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
72586
diff
changeset
|
13 |
Leaf (\<open>\<langle>\<rangle>\<close>) | |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
72586
diff
changeset
|
14 |
Node2 "'a tree23" 'a "'a tree23" (\<open>\<langle>_, _, _\<rangle>\<close>) | |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
72586
diff
changeset
|
15 |
Node3 "'a tree23" 'a "'a tree23" 'a "'a tree23" (\<open>\<langle>_, _, _, _, _\<rangle>\<close>) |
61640 | 16 |
|
17 |
fun inorder :: "'a tree23 \<Rightarrow> 'a list" where |
|
18 |
"inorder Leaf = []" | |
|
19 |
"inorder(Node2 l a r) = inorder l @ a # inorder r" | |
|
20 |
"inorder(Node3 l a m b r) = inorder l @ a # inorder m @ b # inorder r" |
|
21 |
||
22 |
||
23 |
instantiation tree23 :: (type)height |
|
24 |
begin |
|
25 |
||
26 |
fun height_tree23 :: "'a tree23 \<Rightarrow> nat" where |
|
27 |
"height Leaf = 0" | |
|
28 |
"height (Node2 l _ r) = Suc(max (height l) (height r))" | |
|
29 |
"height (Node3 l _ m _ r) = Suc(max (height l) (max (height m) (height r)))" |
|
30 |
||
31 |
instance .. |
|
32 |
||
33 |
end |
|
34 |
||
72566
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents:
70273
diff
changeset
|
35 |
text \<open>Completeness:\<close> |
61640 | 36 |
|
70273 | 37 |
fun complete :: "'a tree23 \<Rightarrow> bool" where |
38 |
"complete Leaf = True" | |
|
72586 | 39 |
"complete (Node2 l _ r) = (height l = height r \<and> complete l & complete r)" | |
70273 | 40 |
"complete (Node3 l _ m _ r) = |
72586 | 41 |
(height l = height m & height m = height r & complete l & complete m & complete r)" |
61640 | 42 |
|
70273 | 43 |
lemma ht_sz_if_complete: "complete t \<Longrightarrow> 2 ^ height t \<le> size t + 1" |
66302 | 44 |
by (induction t) auto |
45 |
||
61640 | 46 |
end |