| author | wenzelm | 
| Wed, 09 Dec 2015 16:36:26 +0100 | |
| changeset 61814 | 1ca1142e1711 | 
| parent 61692 | cb595e12451d | 
| child 62390 | 842917225d56 | 
| 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: 
61229diff
changeset | 8 | Cmp | 
| 61224 | 9 | Set_by_Ordered | 
| 10 | begin | |
| 11 | ||
| 61692 
cb595e12451d
removed lemmas that were only needed for old version of isin.
 nipkow parents: 
61229diff
changeset | 12 | fun isin :: "('a::cmp,'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: 
61229diff
changeset | 14 | "isin (Node _ l a r) x = | 
| 
cb595e12451d
removed lemmas that were only needed for old version of isin.
 nipkow parents: 
61229diff
changeset | 15 | (case cmp x a of | 
| 
cb595e12451d
removed lemmas that were only needed for old version of isin.
 nipkow parents: 
61229diff
changeset | 16 | LT \<Rightarrow> isin l x | | 
| 
cb595e12451d
removed lemmas that were only needed for old version of isin.
 nipkow parents: 
61229diff
changeset | 17 | EQ \<Rightarrow> True | | 
| 
cb595e12451d
removed lemmas that were only needed for old version of isin.
 nipkow parents: 
61229diff
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 | |
| 26 | end |