src/HOL/Data_Structures/Tree_Set.thy
changeset 67929 30486b96274d
parent 66453 cc19f7ca2ed6
child 67964 08cc5ab18c84
--- a/src/HOL/Data_Structures/Tree_Set.thy	Thu Mar 22 17:18:33 2018 +0100
+++ b/src/HOL/Data_Structures/Tree_Set.thy	Fri Mar 23 11:37:02 2018 +0100
@@ -42,12 +42,8 @@
 
 subsection "Functional Correctness Proofs"
 
-lemma "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> elems (inorder t))"
-by (induction t) (auto simp: elems_simps1)
-
-lemma isin_set: "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> elems (inorder t))"
-by (induction t) (auto simp: elems_simps2)
-
+lemma isin_set: "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> set (inorder t))"
+by (induction t) (auto simp: isin_simps)
 
 lemma inorder_insert:
   "sorted(inorder t) \<Longrightarrow> inorder(insert x t) = ins_list x (inorder t)"