src/HOL/Data_Structures/Lookup2.thy
author wenzelm
Wed, 09 Dec 2015 16:36:26 +0100
changeset 61814 1ca1142e1711
parent 61790 0494964bb226
child 62390 842917225d56
permissions -rw-r--r--
clarified type Token.src: plain token list, with usual implicit value assignment; clarified type Token.name_value, notably for head of Token.src; clarified Attrib/Method check_src vs. parser;

(* Author: Tobias Nipkow *)

section \<open>Function \textit{lookup} for Tree2\<close>

theory Lookup2
imports
  Tree2
  Cmp
  Map_by_Ordered
begin

fun lookup :: "('a::cmp * 'b, 'c) tree \<Rightarrow> 'a \<Rightarrow> 'b option" where
"lookup Leaf x = None" |
"lookup (Node _ l (a,b) r) x =
  (case cmp x a of LT \<Rightarrow> lookup l x | GT \<Rightarrow> lookup r x | EQ \<Rightarrow> Some b)"

lemma lookup_map_of:
  "sorted1(inorder t) \<Longrightarrow> lookup t x = map_of (inorder t) x"
by(induction t) (auto simp: map_of_simps split: option.split)

end