src/HOL/Data_Structures/Tree23.thy
author nipkow
Tue, 17 Jun 2025 06:29:55 +0200
changeset 82732 71574900b6ba
parent 80914 d97fdabd9e2b
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
61640
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
     1
(* Author: Tobias Nipkow *)
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
     2
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
     3
section \<open>2-3 Trees\<close>
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
     4
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
     5
theory Tree23
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
     6
imports Main
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
     7
begin
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
     8
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
     9
class height =
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    10
fixes height :: "'a \<Rightarrow> nat"
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    11
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    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
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    16
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    17
fun inorder :: "'a tree23 \<Rightarrow> 'a list" where
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    18
"inorder Leaf = []" |
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    19
"inorder(Node2 l a r) = inorder l @ a # inorder r" |
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    20
"inorder(Node3 l a m b r) = inorder l @ a # inorder m @ b # inorder r"
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    21
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    22
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    23
instantiation tree23 :: (type)height
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    24
begin
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    25
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    26
fun height_tree23 :: "'a tree23 \<Rightarrow> nat" where
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    27
"height Leaf = 0" |
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    28
"height (Node2 l _ r) = Suc(max (height l) (height r))" |
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    29
"height (Node3 l _ m _ r) = Suc(max (height l) (max (height m) (height r)))"
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    30
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    31
instance ..
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    32
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    33
end
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    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
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    36
70273
acc1749c2be9 tuned name
nipkow
parents: 66302
diff changeset
    37
fun complete :: "'a tree23 \<Rightarrow> bool" where
acc1749c2be9 tuned name
nipkow
parents: 66302
diff changeset
    38
"complete Leaf = True" |
72586
nipkow
parents: 72566
diff changeset
    39
"complete (Node2 l _ r) = (height l = height r \<and> complete l & complete r)" |
70273
acc1749c2be9 tuned name
nipkow
parents: 66302
diff changeset
    40
"complete (Node3 l _ m _ r) =
72586
nipkow
parents: 72566
diff changeset
    41
  (height l = height m & height m = height r & complete l & complete m & complete r)"
61640
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    42
70273
acc1749c2be9 tuned name
nipkow
parents: 66302
diff changeset
    43
lemma ht_sz_if_complete: "complete t \<Longrightarrow> 2 ^ height t \<le> size t + 1"
66302
fd89f97c80c2 new lemma
nipkow
parents: 61679
diff changeset
    44
by (induction t) auto
fd89f97c80c2 new lemma
nipkow
parents: 61679
diff changeset
    45
61640
44c9198f210c no CRLF
nipkow
parents: 61469
diff changeset
    46
end