author | hoelzl |
Fri, 16 Sep 2016 13:56:51 +0200 | |
changeset 63886 | 685fb01256af |
parent 63411 | e051eea34990 |
child 67929 | 30486b96274d |
permissions | -rw-r--r-- |
61224 | 1 |
(* Author: Tobias Nipkow *) |
2 |
||
3 |
section \<open>Function \textit{isin} for Tree2\<close> |
|
4 |
||
5 |
theory Isin2 |
|
6 |
imports |
|
7 |
Tree2 |
|
61692
cb595e12451d
removed lemmas that were only needed for old version of isin.
nipkow
parents:
61229
diff
changeset
|
8 |
Cmp |
61224 | 9 |
Set_by_Ordered |
10 |
begin |
|
11 |
||
63411
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents:
62390
diff
changeset
|
12 |
fun isin :: "('a::linorder,'b) tree \<Rightarrow> 'a \<Rightarrow> bool" where |
61224 | 13 |
"isin Leaf x = False" | |
61692
cb595e12451d
removed lemmas that were only needed for old version of isin.
nipkow
parents:
61229
diff
changeset
|
14 |
"isin (Node _ l a r) x = |
cb595e12451d
removed lemmas that were only needed for old version of isin.
nipkow
parents:
61229
diff
changeset
|
15 |
(case cmp x a of |
cb595e12451d
removed lemmas that were only needed for old version of isin.
nipkow
parents:
61229
diff
changeset
|
16 |
LT \<Rightarrow> isin l x | |
cb595e12451d
removed lemmas that were only needed for old version of isin.
nipkow
parents:
61229
diff
changeset
|
17 |
EQ \<Rightarrow> True | |
cb595e12451d
removed lemmas that were only needed for old version of isin.
nipkow
parents:
61229
diff
changeset
|
18 |
GT \<Rightarrow> isin r x)" |
61224 | 19 |
|
20 |
lemma "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> elems(inorder t))" |
|
61229 | 21 |
by (induction t) (auto simp: elems_simps1) |
61224 | 22 |
|
23 |
lemma isin_set: "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> elems(inorder t))" |
|
61229 | 24 |
by (induction t) (auto simp: elems_simps2) |
61224 | 25 |
|
62390 | 26 |
end |