src/HOL/Data_Structures/Lookup2.thy
changeset 61693 f6b9f528c89c
parent 61232 c46faf9762f7
child 61790 0494964bb226
--- a/src/HOL/Data_Structures/Lookup2.thy	Tue Nov 17 11:44:10 2015 +0100
+++ b/src/HOL/Data_Structures/Lookup2.thy	Tue Nov 17 12:01:19 2015 +0100
@@ -5,17 +5,16 @@
 theory Lookup2
 imports
   Tree2
+  Cmp
   Map_by_Ordered
 begin
 
-fun lookup :: "('a::linorder * 'b, 'c) tree \<Rightarrow> 'a \<Rightarrow> 'b option" where
+fun lookup :: "('a::cmp * 'b, 'c) tree \<Rightarrow> 'a \<Rightarrow> 'b option" where
 "lookup Leaf x = None" |
 "lookup (Node _ l (a,b) r) x =
-  (if x < a then lookup l x else
-   if x > a then lookup r x else Some b)"
+  (case cmp x a of LT \<Rightarrow> lookup l x | GT \<Rightarrow> lookup r x | EQ \<Rightarrow> Some b)"
 
-lemma lookup_eq:
-  "sorted1(inorder t) \<Longrightarrow> lookup t x = map_of (inorder t) x"
+lemma lookup_eq: "sorted1(inorder t) \<Longrightarrow> lookup t x = map_of (inorder t) x"
 by(induction t) (auto simp: map_of_simps split: option.split)
 
 end
\ No newline at end of file