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::linorder * 'b) tree23 \<Rightarrow> 'a \<Rightarrow> 'b option" where |
11 fun lookup :: "('a::cmp * '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 = |
13 "lookup (Node2 l (a,b) r) x = (case cmp x a of |
14 (if x < a then lookup l x else |
14 LT \<Rightarrow> lookup l x | |
15 if a < x then lookup r x else Some b)" | |
15 GT \<Rightarrow> lookup r x | |
16 "lookup (Node3 l (a1,b1) m (a2,b2) r) x = |
16 EQ \<Rightarrow> Some b)" | |
17 (if x < a1 then lookup l x else |
17 "lookup (Node3 l (a1,b1) m (a2,b2) r) x = (case cmp x a1 of |
18 if x = a1 then Some b1 else |
18 LT \<Rightarrow> lookup l x | |
19 if x < a2 then lookup m x else |
19 EQ \<Rightarrow> Some b1 | |
20 if x = a2 then Some b2 |
20 GT \<Rightarrow> (case cmp x a2 of |
21 else lookup r x)" |
21 LT \<Rightarrow> lookup m x | |
|
22 EQ \<Rightarrow> Some b2 | |
|
23 GT \<Rightarrow> lookup r x))" |
22 |
24 |
23 fun upd :: "'a::linorder \<Rightarrow> 'b \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) up\<^sub>i" where |
25 fun upd :: "'a::cmp \<Rightarrow> 'b \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) up\<^sub>i" where |
24 "upd x y Leaf = Up\<^sub>i Leaf (x,y) Leaf" | |
26 "upd x y Leaf = Up\<^sub>i Leaf (x,y) Leaf" | |
25 "upd x y (Node2 l ab r) = |
27 "upd x y (Node2 l ab r) = (case cmp x (fst ab) of |
26 (if x < fst ab then |
28 LT \<Rightarrow> (case upd x y l of |
27 (case upd x y l of |
|
28 T\<^sub>i l' => T\<^sub>i (Node2 l' ab r) |
29 T\<^sub>i l' => T\<^sub>i (Node2 l' ab r) |
29 | 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)) | |
30 else if x = fst ab then T\<^sub>i (Node2 l (x,y) r) |
31 EQ \<Rightarrow> T\<^sub>i (Node2 l (x,y) r) | |
31 else |
32 GT \<Rightarrow> (case upd x y r of |
32 (case upd x y r of |
|
33 T\<^sub>i r' => T\<^sub>i (Node2 l ab r') |
33 T\<^sub>i r' => T\<^sub>i (Node2 l ab r') |
34 | Up\<^sub>i r1 ab' r2 => T\<^sub>i (Node3 l ab r1 ab' r2)))" | |
34 | Up\<^sub>i r1 ab' r2 => T\<^sub>i (Node3 l ab r1 ab' r2)))" | |
35 "upd x y (Node3 l ab1 m ab2 r) = |
35 "upd x y (Node3 l ab1 m ab2 r) = (case cmp x (fst ab1) of |
36 (if x < fst ab1 then |
36 LT \<Rightarrow> (case upd x y l of |
37 (case upd x y l of |
|
38 T\<^sub>i l' => T\<^sub>i (Node3 l' ab1 m ab2 r) |
37 T\<^sub>i l' => T\<^sub>i (Node3 l' ab1 m ab2 r) |
39 | Up\<^sub>i l1 ab' l2 => Up\<^sub>i (Node2 l1 ab' l2) ab1 (Node2 m ab2 r)) |
38 | Up\<^sub>i l1 ab' l2 => Up\<^sub>i (Node2 l1 ab' l2) ab1 (Node2 m ab2 r)) | |
40 else if x = fst ab1 then T\<^sub>i (Node3 l (x,y) m ab2 r) |
39 EQ \<Rightarrow> T\<^sub>i (Node3 l (x,y) m ab2 r) | |
41 else if x < fst ab2 then |
40 GT \<Rightarrow> (case cmp x (fst ab2) of |
42 (case upd x y m of |
41 LT \<Rightarrow> (case upd x y m of |
43 T\<^sub>i m' => T\<^sub>i (Node3 l ab1 m' ab2 r) |
42 T\<^sub>i m' => T\<^sub>i (Node3 l ab1 m' ab2 r) |
44 | Up\<^sub>i m1 ab' m2 => Up\<^sub>i (Node2 l ab1 m1) ab' (Node2 m2 ab2 r)) |
43 | Up\<^sub>i m1 ab' m2 => Up\<^sub>i (Node2 l ab1 m1) ab' (Node2 m2 ab2 r)) | |
45 else if x = fst ab2 then T\<^sub>i (Node3 l ab1 m (x,y) r) |
44 EQ \<Rightarrow> T\<^sub>i (Node3 l ab1 m (x,y) r) | |
46 else |
45 GT \<Rightarrow> (case upd x y r of |
47 (case upd x y r of |
46 T\<^sub>i r' => T\<^sub>i (Node3 l ab1 m ab2 r') |
48 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))))" |
49 | Up\<^sub>i r1 ab' r2 => Up\<^sub>i (Node2 l ab1 m) ab2 (Node2 r1 ab' r2)))" |
|
50 |
48 |
51 definition update :: "'a::linorder \<Rightarrow> 'b \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) tree23" where |
49 definition update :: "'a::cmp \<Rightarrow> 'b \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) tree23" where |
52 "update a b t = tree\<^sub>i(upd a b t)" |
50 "update a b t = tree\<^sub>i(upd a b t)" |
53 |
51 |
54 fun del :: "'a::linorder \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) up\<^sub>d" |
52 fun del :: "'a::cmp \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) up\<^sub>d" where |
55 where |
|
56 "del x Leaf = T\<^sub>d Leaf" | |
53 "del x Leaf = T\<^sub>d Leaf" | |
57 "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))" | |
58 "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 |
59 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)" | |
60 "del x (Node2 l ab1 r) = (if x<fst ab1 then node21 (del x l) ab1 r else |
57 "del x (Node2 l ab1 r) = (case cmp x (fst ab1) of |
61 if x > fst ab1 then node22 l ab1 (del x r) else |
58 LT \<Rightarrow> node21 (del x l) ab1 r | |
62 let (ab1',t) = del_min r in node22 l ab1' t)" | |
59 GT \<Rightarrow> node22 l ab1 (del x r) | |
63 "del x (Node3 l ab1 m ab2 r) = (if x<fst ab1 then node31 (del x l) ab1 m ab2 r else |
60 EQ \<Rightarrow> let (ab1',t) = del_min r in node22 l ab1' t)" | |
64 if x = fst ab1 then let (ab1',m') = del_min m in node32 l ab1' m' ab2 r else |
61 "del x (Node3 l ab1 m ab2 r) = (case cmp x (fst ab1) of |
65 if x < fst ab2 then node32 l ab1 (del x m) ab2 r else |
62 LT \<Rightarrow> node31 (del x l) ab1 m ab2 r | |
66 if x = fst ab2 then let (ab2',r') = del_min r in node33 l ab1 m ab2' r' |
63 EQ \<Rightarrow> let (ab1',m') = del_min m in node32 l ab1' m' ab2 r | |
67 else node33 l ab1 m ab2 (del x r))" |
64 GT \<Rightarrow> (case cmp x (fst ab2) of |
|
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' | |
|
67 GT \<Rightarrow> node33 l ab1 m ab2 (del x r)))" |
68 |
68 |
69 definition delete :: "'a::linorder \<Rightarrow> ('a*'b) tree23 \<Rightarrow> ('a*'b) tree23" where |
69 definition delete :: "'a::cmp \<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 |