1 (* Author: Tobias Nipkow *) |
1 (* Author: Tobias Nipkow *) |
2 |
2 |
3 section {* Tree Implementation of Sets *} |
3 section \<open>Unbalanced Tree Implementation of Set\<close> |
4 |
4 |
5 theory Tree_Set |
5 theory Tree_Set |
6 imports |
6 imports |
7 "~~/src/HOL/Library/Tree" |
7 "~~/src/HOL/Library/Tree" |
8 Cmp |
8 Cmp |
9 Set_by_Ordered |
9 Set_by_Ordered |
10 begin |
10 begin |
11 |
11 |
12 fun isin :: "'a::cmp tree \<Rightarrow> 'a \<Rightarrow> bool" where |
12 fun isin :: "'a::linorder tree \<Rightarrow> 'a \<Rightarrow> bool" where |
13 "isin Leaf x = False" | |
13 "isin Leaf x = False" | |
14 "isin (Node l a r) x = |
14 "isin (Node l a r) x = |
15 (case cmp x a of |
15 (case cmp x a of |
16 LT \<Rightarrow> isin l x | |
16 LT \<Rightarrow> isin l x | |
17 EQ \<Rightarrow> True | |
17 EQ \<Rightarrow> True | |
18 GT \<Rightarrow> isin r x)" |
18 GT \<Rightarrow> isin r x)" |
19 |
19 |
20 hide_const (open) insert |
20 hide_const (open) insert |
21 |
21 |
22 fun insert :: "'a::cmp \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where |
22 fun insert :: "'a::linorder \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where |
23 "insert x Leaf = Node Leaf x Leaf" | |
23 "insert x Leaf = Node Leaf x Leaf" | |
24 "insert x (Node l a r) = |
24 "insert x (Node l a r) = |
25 (case cmp x a of |
25 (case cmp x a of |
26 LT \<Rightarrow> Node (insert x l) a r | |
26 LT \<Rightarrow> Node (insert x l) a r | |
27 EQ \<Rightarrow> Node l a r | |
27 EQ \<Rightarrow> Node l a r | |
29 |
29 |
30 fun del_min :: "'a tree \<Rightarrow> 'a * 'a tree" where |
30 fun del_min :: "'a tree \<Rightarrow> 'a * 'a tree" where |
31 "del_min (Node l a r) = |
31 "del_min (Node l a r) = |
32 (if l = Leaf then (a,r) else let (x,l') = del_min l in (x, Node l' a r))" |
32 (if l = Leaf then (a,r) else let (x,l') = del_min l in (x, Node l' a r))" |
33 |
33 |
34 fun delete :: "'a::cmp \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where |
34 fun delete :: "'a::linorder \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where |
35 "delete x Leaf = Leaf" | |
35 "delete x Leaf = Leaf" | |
36 "delete x (Node l a r) = |
36 "delete x (Node l a r) = |
37 (case cmp x a of |
37 (case cmp x a of |
38 LT \<Rightarrow> Node (delete x l) a r | |
38 LT \<Rightarrow> Node (delete x l) a r | |
39 GT \<Rightarrow> Node l a (delete x r) | |
39 GT \<Rightarrow> Node l a (delete x r) | |