| author | nipkow |
| Sun, 08 Apr 2018 09:46:33 +0200 | |
| changeset 67964 | 08cc5ab18c84 |
| parent 67929 | 30486b96274d |
| child 67965 | aaa31cd0caef |
| permissions | -rw-r--r-- |
(* Author: Tobias Nipkow *) section \<open>Function \textit{isin} for Tree2\<close> theory Isin2 imports Tree2 Cmp Set_Interfaces begin fun isin :: "('a::linorder,'b) tree \<Rightarrow> 'a \<Rightarrow> bool" where "isin Leaf x = False" | "isin (Node _ l a r) x = (case cmp x a of LT \<Rightarrow> isin l x | EQ \<Rightarrow> True | GT \<Rightarrow> isin r x)" lemma isin_set: "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> set(inorder t))" by (induction t) (auto simp: isin_simps) end