src/HOL/Data_Structures/Tree234.thy
author haftmann
Mon, 06 Feb 2017 20:56:34 +0100
changeset 64990 c6a7de505796
parent 61703 1f1354ca7ea7
child 67406 23307fd33906
permissions -rw-r--r--
more explicit errors in pathological cases
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
61640
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
     1
(* Author: Tobias Nipkow *)
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
     2
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
     3
section {* 2-3-4 Trees *}
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
     4
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
     5
theory Tree234
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
     6
imports Main
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
     7
begin
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
     8
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
     9
class height =
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    10
fixes height :: "'a \<Rightarrow> nat"
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    11
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    12
datatype 'a tree234 =
61703
1f1354ca7ea7 tuned and converted to cmp
nipkow
parents: 61640
diff changeset
    13
  Leaf ("\<langle>\<rangle>") |
1f1354ca7ea7 tuned and converted to cmp
nipkow
parents: 61640
diff changeset
    14
  Node2 "'a tree234" 'a "'a tree234"  ("\<langle>_, _, _\<rangle>") |
1f1354ca7ea7 tuned and converted to cmp
nipkow
parents: 61640
diff changeset
    15
  Node3 "'a tree234" 'a "'a tree234" 'a "'a tree234"  ("\<langle>_, _, _, _, _\<rangle>") |
61640
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    16
  Node4 "'a tree234" 'a "'a tree234" 'a "'a tree234" 'a "'a tree234"
61703
1f1354ca7ea7 tuned and converted to cmp
nipkow
parents: 61640
diff changeset
    17
    ("\<langle>_, _, _, _, _, _, _\<rangle>")
61640
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    18
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    19
fun inorder :: "'a tree234 \<Rightarrow> 'a list" where
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    20
"inorder Leaf = []" |
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    21
"inorder(Node2 l a r) = inorder l @ a # inorder r" |
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    22
"inorder(Node3 l a m b r) = inorder l @ a # inorder m @ b # inorder r" |
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    23
"inorder(Node4 l a m b n c r) = inorder l @ a # inorder m @ b # inorder n @ c # inorder r"
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    24
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    25
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    26
instantiation tree234 :: (type)height
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    27
begin
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    28
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    29
fun height_tree234 :: "'a tree234 \<Rightarrow> nat" where
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    30
"height Leaf = 0" |
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    31
"height (Node2 l _ r) = Suc(max (height l) (height r))" |
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    32
"height (Node3 l _ m _ r) = Suc(max (height l) (max (height m) (height r)))" |
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    33
"height (Node4 l _ m _ n _ r) = Suc(max (height l) (max (height m) (max (height n) (height r))))"
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    34
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    35
instance ..
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    36
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    37
end
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    38
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    39
text{* Balanced: *}
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    40
fun bal :: "'a tree234 \<Rightarrow> bool" where
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    41
"bal Leaf = True" |
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    42
"bal (Node2 l _ r) = (bal l & bal r & height l = height r)" |
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    43
"bal (Node3 l _ m _ r) = (bal l & bal m & bal r & height l = height m & height m = height r)" |
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    44
"bal (Node4 l _ m _ n _ r) = (bal l & bal m & bal n & bal r & height l = height m & height m = height n & height n = height r)"
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    45
44c9198f210c no CRLF
nipkow
parents: 61515
diff changeset
    46
end