| author | wenzelm | 
| Tue, 01 Oct 2024 21:30:59 +0200 | |
| changeset 81092 | c92efbf32bfe | 
| parent 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 61640 | 1 | (* Author: Tobias Nipkow *) | 
| 2 | ||
| 67406 | 3 | section \<open>2-3-4 Trees\<close> | 
| 61640 | 4 | |
| 5 | theory Tree234 | |
| 6 | imports Main | |
| 7 | begin | |
| 8 | ||
| 9 | class height = | |
| 10 | fixes height :: "'a \<Rightarrow> nat" | |
| 11 | ||
| 12 | datatype 'a tree234 = | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
67406diff
changeset | 13 | Leaf (\<open>\<langle>\<rangle>\<close>) | | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
67406diff
changeset | 14 | Node2 "'a tree234" 'a "'a tree234" (\<open>\<langle>_, _, _\<rangle>\<close>) | | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
67406diff
changeset | 15 | Node3 "'a tree234" 'a "'a tree234" 'a "'a tree234" (\<open>\<langle>_, _, _, _, _\<rangle>\<close>) | | 
| 61640 | 16 | Node4 "'a tree234" 'a "'a tree234" 'a "'a tree234" 'a "'a tree234" | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
67406diff
changeset | 17 | (\<open>\<langle>_, _, _, _, _, _, _\<rangle>\<close>) | 
| 61640 | 18 | |
| 19 | fun inorder :: "'a tree234 \<Rightarrow> 'a list" where | |
| 20 | "inorder Leaf = []" | | |
| 21 | "inorder(Node2 l a r) = inorder l @ a # inorder r" | | |
| 22 | "inorder(Node3 l a m b r) = inorder l @ a # inorder m @ b # inorder r" | | |
| 23 | "inorder(Node4 l a m b n c r) = inorder l @ a # inorder m @ b # inorder n @ c # inorder r" | |
| 24 | ||
| 25 | ||
| 26 | instantiation tree234 :: (type)height | |
| 27 | begin | |
| 28 | ||
| 29 | fun height_tree234 :: "'a tree234 \<Rightarrow> nat" where | |
| 30 | "height Leaf = 0" | | |
| 31 | "height (Node2 l _ r) = Suc(max (height l) (height r))" | | |
| 32 | "height (Node3 l _ m _ r) = Suc(max (height l) (max (height m) (height r)))" | | |
| 33 | "height (Node4 l _ m _ n _ r) = Suc(max (height l) (max (height m) (max (height n) (height r))))" | |
| 34 | ||
| 35 | instance .. | |
| 36 | ||
| 37 | end | |
| 38 | ||
| 67406 | 39 | text\<open>Balanced:\<close> | 
| 61640 | 40 | fun bal :: "'a tree234 \<Rightarrow> bool" where | 
| 41 | "bal Leaf = True" | | |
| 42 | "bal (Node2 l _ r) = (bal l & bal r & height l = height r)" | | |
| 43 | "bal (Node3 l _ m _ r) = (bal l & bal m & bal r & height l = height m & height m = height r)" | | |
| 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)" | |
| 45 | ||
| 46 | end |