--- a/src/HOL/Data_Structures/Tree_Set.thy	Wed Nov 04 15:07:23 2015 +0100
+++ b/src/HOL/Data_Structures/Tree_Set.thy	Thu Nov 05 08:27:14 2015 +0100
@@ -5,31 +5,34 @@
 theory Tree_Set
 imports
   "~~/src/HOL/Library/Tree"
+  Cmp
   Set_by_Ordered
 begin
 
-fun isin :: "'a::linorder tree \<Rightarrow> 'a \<Rightarrow> bool" where
+fun isin :: "'a::cmp tree \<Rightarrow> 'a \<Rightarrow> bool" where
 "isin Leaf x = False" |
-"isin (Node l a r) x = (x < a \<and> isin l x \<or> x=a \<or> isin r x)"
+"isin (Node l a r) x =
+  (case cmp x a of LT \<Rightarrow> isin l x | EQ \<Rightarrow> True | GT \<Rightarrow> isin r x)"
 
 hide_const (open) insert
 
-fun insert :: "'a::linorder \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
+fun insert :: "'a::cmp \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
 "insert x Leaf = Node Leaf x Leaf" |
-"insert x (Node l a r) =
-   (if x < a then Node (insert x l) a r else
-    if x = a then Node l a r
-    else Node l a (insert x r))"
+"insert x (Node l a r) = (case cmp x a of
+      LT \<Rightarrow> Node (insert x l) a r |
+      EQ \<Rightarrow> Node l a r |
+      GT \<Rightarrow> Node l a (insert x r))"
 
 fun del_min :: "'a tree \<Rightarrow> 'a * 'a tree" where
 "del_min (Node Leaf a r) = (a, r)" |
 "del_min (Node l a r) = (let (x,l') = del_min l in (x, Node l' a r))"
 
-fun delete :: "'a::linorder \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
+fun delete :: "'a::cmp \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
 "delete x Leaf = Leaf" |
-"delete x (Node l a r) = (if x < a then Node (delete x l) a r else
-  if x > a then Node l a (delete x r) else
-  if r = Leaf then l else let (a',r') = del_min r in Node l a' r')"
+"delete x (Node l a r) = (case cmp x a of
+  LT \<Rightarrow>  Node (delete x l) a r |
+  GT \<Rightarrow>  Node l a (delete x r) |
+  EQ \<Rightarrow> if r = Leaf then l else let (a',r') = del_min r in Node l a' r')"
 
 
 subsection "Functional Correctness Proofs"
@@ -56,7 +59,6 @@
   "sorted(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
 by(induction t) (auto simp: del_list_simps del_minD split: prod.splits)
 
-
 interpretation Set_by_Ordered
 where empty = Leaf and isin = isin and insert = insert and delete = delete
 and inorder = inorder and wf = "\<lambda>_. True"