src/HOL/Data_Structures/RBT.thy
author wenzelm
Sat, 28 Nov 2020 15:15:53 +0100
changeset 72755 8dffbe01a3e1
parent 71830 7a997ead54b0
permissions -rw-r--r--
support for Scala compile-time positions;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
61224
759b5299a9f2 added red black trees
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
759b5299a9f2 added red black trees
nipkow
parents:
diff changeset
     2
61469
cd82b1023932 added 2-3 trees (simpler and more complete than the version in ex/Tree23)
nipkow
parents: 61231
diff changeset
     3
section \<open>Red-Black Trees\<close>
61224
759b5299a9f2 added red black trees
nipkow
parents:
diff changeset
     4
759b5299a9f2 added red black trees
nipkow
parents:
diff changeset
     5
theory RBT
759b5299a9f2 added red black trees
nipkow
parents:
diff changeset
     6
imports Tree2
759b5299a9f2 added red black trees
nipkow
parents:
diff changeset
     7
begin
759b5299a9f2 added red black trees
nipkow
parents:
diff changeset
     8
759b5299a9f2 added red black trees
nipkow
parents:
diff changeset
     9
datatype color = Red | Black
759b5299a9f2 added red black trees
nipkow
parents:
diff changeset
    10
70755
3fb16bed5d6c replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents: 70571
diff changeset
    11
type_synonym 'a rbt = "('a*color)tree"
61224
759b5299a9f2 added red black trees
nipkow
parents:
diff changeset
    12
70755
3fb16bed5d6c replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents: 70571
diff changeset
    13
abbreviation R where "R l a r \<equiv> Node l (a, Red) r"
3fb16bed5d6c replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents: 70571
diff changeset
    14
abbreviation B where "B l a r \<equiv> Node l (a, Black) r"
61224
759b5299a9f2 added red black trees
nipkow
parents:
diff changeset
    15
64960
8be78855ee7a split balance into two, clearer etc
nipkow
parents: 64952
diff changeset
    16
fun baliL :: "'a rbt \<Rightarrow> 'a \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt" where
70571
e72daea2aab6 tuned names
nipkow
parents: 68413
diff changeset
    17
"baliL (R (R t1 a t2) b t3) c t4 = R (B t1 a t2) b (B t3 c t4)" |
e72daea2aab6 tuned names
nipkow
parents: 68413
diff changeset
    18
"baliL (R t1 a (R t2 b t3)) c t4 = R (B t1 a t2) b (B t3 c t4)" |
64960
8be78855ee7a split balance into two, clearer etc
nipkow
parents: 64952
diff changeset
    19
"baliL t1 a t2 = B t1 a t2"
8be78855ee7a split balance into two, clearer etc
nipkow
parents: 64952
diff changeset
    20
8be78855ee7a split balance into two, clearer etc
nipkow
parents: 64952
diff changeset
    21
fun baliR :: "'a rbt \<Rightarrow> 'a \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt" where
70571
e72daea2aab6 tuned names
nipkow
parents: 68413
diff changeset
    22
"baliR t1 a (R t2 b (R t3 c t4)) = R (B t1 a t2) b (B t3 c t4)" |
e72daea2aab6 tuned names
nipkow
parents: 68413
diff changeset
    23
"baliR t1 a (R (R t2 b t3) c t4) = R (B t1 a t2) b (B t3 c t4)" |
64960
8be78855ee7a split balance into two, clearer etc
nipkow
parents: 64952
diff changeset
    24
"baliR t1 a t2 = B t1 a t2"
61224
759b5299a9f2 added red black trees
nipkow
parents:
diff changeset
    25
61749
7f530d7e552d paint root black after insert and delete
nipkow
parents: 61678
diff changeset
    26
fun paint :: "color \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt" where
7f530d7e552d paint root black after insert and delete
nipkow
parents: 61678
diff changeset
    27
"paint c Leaf = Leaf" |
70755
3fb16bed5d6c replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents: 70571
diff changeset
    28
"paint c (Node l (a,_) r) = Node l (a,c) r"
61224
759b5299a9f2 added red black trees
nipkow
parents:
diff changeset
    29
64960
8be78855ee7a split balance into two, clearer etc
nipkow
parents: 64952
diff changeset
    30
fun baldL :: "'a rbt \<Rightarrow> 'a \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt" where
70571
e72daea2aab6 tuned names
nipkow
parents: 68413
diff changeset
    31
"baldL (R t1 a t2) b t3 = R (B t1 a t2) b t3" |
71349
nipkow
parents: 71348
diff changeset
    32
"baldL t1 a (B t2 b t3) = baliR t1 a (R t2 b t3)" |
nipkow
parents: 71348
diff changeset
    33
"baldL t1 a (R (B t2 b t3) c t4) = R (B t1 a t2) b (baliR t3 c (paint Red t4))" |
70571
e72daea2aab6 tuned names
nipkow
parents: 68413
diff changeset
    34
"baldL t1 a t2 = R t1 a t2"
61224
759b5299a9f2 added red black trees
nipkow
parents:
diff changeset
    35
64960
8be78855ee7a split balance into two, clearer etc
nipkow
parents: 64952
diff changeset
    36
fun baldR :: "'a rbt \<Rightarrow> 'a \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt" where
70571
e72daea2aab6 tuned names
nipkow
parents: 68413
diff changeset
    37
"baldR t1 a (R t2 b t3) = R t1 a (B t2 b t3)" |
e72daea2aab6 tuned names
nipkow
parents: 68413
diff changeset
    38
"baldR (B t1 a t2) b t3 = baliL (R t1 a t2) b t3" |
e72daea2aab6 tuned names
nipkow
parents: 68413
diff changeset
    39
"baldR (R t1 a (B t2 b t3)) c t4 = R (baliL (paint Red t1) a t2) b (B t3 c t4)" |
e72daea2aab6 tuned names
nipkow
parents: 68413
diff changeset
    40
"baldR t1 a t2 = R t1 a t2"
61224
759b5299a9f2 added red black trees
nipkow
parents:
diff changeset
    41
71830
7a997ead54b0 "app" -> "join" for RBTs
nipkow
parents: 71463
diff changeset
    42
fun join :: "'a rbt \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt" where
7a997ead54b0 "app" -> "join" for RBTs
nipkow
parents: 71463
diff changeset
    43
"join Leaf t = t" |
7a997ead54b0 "app" -> "join" for RBTs
nipkow
parents: 71463
diff changeset
    44
"join t Leaf = t" |
7a997ead54b0 "app" -> "join" for RBTs
nipkow
parents: 71463
diff changeset
    45
"join (R t1 a t2) (R t3 c t4) =
7a997ead54b0 "app" -> "join" for RBTs
nipkow
parents: 71463
diff changeset
    46
  (case join t2 t3 of
61678
b594e9277be3 tuned white space
nipkow
parents: 61534
diff changeset
    47
     R u2 b u3 \<Rightarrow> (R (R t1 a u2) b (R u3 c t4)) |
b594e9277be3 tuned white space
nipkow
parents: 61534
diff changeset
    48
     t23 \<Rightarrow> R t1 a (R t23 c t4))" |
71830
7a997ead54b0 "app" -> "join" for RBTs
nipkow
parents: 71463
diff changeset
    49
"join (B t1 a t2) (B t3 c t4) =
7a997ead54b0 "app" -> "join" for RBTs
nipkow
parents: 71463
diff changeset
    50
  (case join t2 t3 of
71348
nipkow
parents: 70755
diff changeset
    51
     R u2 b u3 \<Rightarrow> R (B t1 a u2) b (B u3 c t4) |
64960
8be78855ee7a split balance into two, clearer etc
nipkow
parents: 64952
diff changeset
    52
     t23 \<Rightarrow> baldL t1 a (B t23 c t4))" |
71830
7a997ead54b0 "app" -> "join" for RBTs
nipkow
parents: 71463
diff changeset
    53
"join t1 (R t2 a t3) = R (join t1 t2) a t3" |
7a997ead54b0 "app" -> "join" for RBTs
nipkow
parents: 71463
diff changeset
    54
"join (R t1 a t2) t3 = R t1 a (join t2 t3)" 
61224
759b5299a9f2 added red black trees
nipkow
parents:
diff changeset
    55
759b5299a9f2 added red black trees
nipkow
parents:
diff changeset
    56
end