src/HOL/Data_Structures/Isin2.thy
author wenzelm
Wed, 19 May 2021 11:54:58 +0200
changeset 73740 c46ff0efa1ce
parent 70755 3fb16bed5d6c
permissions -rw-r--r--
more direct use of latex tools: avoid diversion into "isabelle latex -o pdf" and its confusion of ISABELLE_PDFLATEX vs. ISABELLE_LUALATEX; clarified ISABELLE_MAKEINDEX options;

(* Author: Tobias Nipkow *)

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

theory Isin2
imports
  Tree2
  Cmp
  Set_Specs
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_inorder: "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> set(inorder t))"
by (induction t rule: tree2_induct) (auto simp: isin_simps)

lemma isin_set_tree: "bst t \<Longrightarrow> isin t x \<longleftrightarrow> x \<in> set_tree t"
by(induction t rule: tree2_induct) auto

end