src/HOL/Data_Structures/Tree23_Map.thy
changeset 63412 def97df48390
parent 63411 e051eea34990
child 63636 6f38b7abb648
equal deleted inserted replaced
63410:9789ccc2a477 63412:def97df48390
     6 imports
     6 imports
     7   Tree23_Set
     7   Tree23_Set
     8   Map_by_Ordered
     8   Map_by_Ordered
     9 begin
     9 begin
    10 
    10 
    11 fun lookup :: "('a::cmp * 'b) tree23 \<Rightarrow> 'a \<Rightarrow> 'b option" where
    11 fun lookup :: "('a::linorder * 'b) tree23 \<Rightarrow> 'a \<Rightarrow> 'b option" where
    12 "lookup Leaf x = None" |
    12 "lookup Leaf x = None" |
    13 "lookup (Node2 l (a,b) r) x = (case cmp x a of
    13 "lookup (Node2 l (a,b) r) x = (case cmp x a of
    14   LT \<Rightarrow> lookup l x |
    14   LT \<Rightarrow> lookup l x |
    15   GT \<Rightarrow> lookup r x |
    15   GT \<Rightarrow> lookup r x |
    16   EQ \<Rightarrow> Some b)" |
    16   EQ \<Rightarrow> Some b)" |
    20   GT \<Rightarrow> (case cmp x a2 of
    20   GT \<Rightarrow> (case cmp x a2 of
    21           LT \<Rightarrow> lookup m x |
    21           LT \<Rightarrow> lookup m x |
    22           EQ \<Rightarrow> Some b2 |
    22           EQ \<Rightarrow> Some b2 |
    23           GT \<Rightarrow> lookup r x))"
    23           GT \<Rightarrow> lookup r x))"
    24 
    24 
    25 fun upd :: "'a::cmp \<Rightarrow> 'b \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) up\<^sub>i" where
    25 fun upd :: "'a::linorder \<Rightarrow> 'b \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) up\<^sub>i" where
    26 "upd x y Leaf = Up\<^sub>i Leaf (x,y) Leaf" |
    26 "upd x y Leaf = Up\<^sub>i Leaf (x,y) Leaf" |
    27 "upd x y (Node2 l ab r) = (case cmp x (fst ab) of
    27 "upd x y (Node2 l ab r) = (case cmp x (fst ab) of
    28    LT \<Rightarrow> (case upd x y l of
    28    LT \<Rightarrow> (case upd x y l of
    29            T\<^sub>i l' => T\<^sub>i (Node2 l' ab r)
    29            T\<^sub>i l' => T\<^sub>i (Node2 l' ab r)
    30          | Up\<^sub>i l1 ab' l2 => T\<^sub>i (Node3 l1 ab' l2 ab r)) |
    30          | Up\<^sub>i l1 ab' l2 => T\<^sub>i (Node3 l1 ab' l2 ab r)) |
    44            EQ \<Rightarrow> T\<^sub>i (Node3 l ab1 m (x,y) r) |
    44            EQ \<Rightarrow> T\<^sub>i (Node3 l ab1 m (x,y) r) |
    45            GT \<Rightarrow> (case upd x y r of
    45            GT \<Rightarrow> (case upd x y r of
    46                    T\<^sub>i r' => T\<^sub>i (Node3 l ab1 m ab2 r')
    46                    T\<^sub>i r' => T\<^sub>i (Node3 l ab1 m ab2 r')
    47                  | Up\<^sub>i r1 ab' r2 => Up\<^sub>i (Node2 l ab1 m) ab2 (Node2 r1 ab' r2))))"
    47                  | Up\<^sub>i r1 ab' r2 => Up\<^sub>i (Node2 l ab1 m) ab2 (Node2 r1 ab' r2))))"
    48 
    48 
    49 definition update :: "'a::cmp \<Rightarrow> 'b \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) tree23" where
    49 definition update :: "'a::linorder \<Rightarrow> 'b \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) tree23" where
    50 "update a b t = tree\<^sub>i(upd a b t)"
    50 "update a b t = tree\<^sub>i(upd a b t)"
    51 
    51 
    52 fun del :: "'a::cmp \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) up\<^sub>d" where
    52 fun del :: "'a::linorder \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) up\<^sub>d" where
    53 "del x Leaf = T\<^sub>d Leaf" |
    53 "del x Leaf = T\<^sub>d Leaf" |
    54 "del x (Node2 Leaf ab1 Leaf) = (if x=fst ab1 then Up\<^sub>d Leaf else T\<^sub>d(Node2 Leaf ab1 Leaf))" |
    54 "del x (Node2 Leaf ab1 Leaf) = (if x=fst ab1 then Up\<^sub>d Leaf else T\<^sub>d(Node2 Leaf ab1 Leaf))" |
    55 "del x (Node3 Leaf ab1 Leaf ab2 Leaf) = T\<^sub>d(if x=fst ab1 then Node2 Leaf ab2 Leaf
    55 "del x (Node3 Leaf ab1 Leaf ab2 Leaf) = T\<^sub>d(if x=fst ab1 then Node2 Leaf ab2 Leaf
    56   else if x=fst ab2 then Node2 Leaf ab1 Leaf else Node3 Leaf ab1 Leaf ab2 Leaf)" |
    56   else if x=fst ab2 then Node2 Leaf ab1 Leaf else Node3 Leaf ab1 Leaf ab2 Leaf)" |
    57 "del x (Node2 l ab1 r) = (case cmp x (fst ab1) of
    57 "del x (Node2 l ab1 r) = (case cmp x (fst ab1) of
    64   GT \<Rightarrow> (case cmp x (fst ab2) of
    64   GT \<Rightarrow> (case cmp x (fst ab2) of
    65            LT \<Rightarrow> node32 l ab1 (del x m) ab2 r |
    65            LT \<Rightarrow> node32 l ab1 (del x m) ab2 r |
    66            EQ \<Rightarrow> let (ab2',r') = del_min r in node33 l ab1 m ab2' r' |
    66            EQ \<Rightarrow> let (ab2',r') = del_min r in node33 l ab1 m ab2' r' |
    67            GT \<Rightarrow> node33 l ab1 m ab2 (del x r)))"
    67            GT \<Rightarrow> node33 l ab1 m ab2 (del x r)))"
    68 
    68 
    69 definition delete :: "'a::cmp \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) tree23" where
    69 definition delete :: "'a::linorder \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) tree23" where
    70 "delete x t = tree\<^sub>d(del x t)"
    70 "delete x t = tree\<^sub>d(del x t)"
    71 
    71 
    72 
    72 
    73 subsection \<open>Functional Correctness\<close>
    73 subsection \<open>Functional Correctness\<close>
    74 
    74