src/HOL/Data_Structures/Tree2.thy
changeset 72586 e3ba2578ad9d
parent 72080 2030eacf3a72
--- a/src/HOL/Data_Structures/Tree2.thy	Thu Oct 22 15:18:08 2020 +0200
+++ b/src/HOL/Data_Structures/Tree2.thy	Thu Nov 12 12:44:17 2020 +0100
@@ -21,11 +21,11 @@
 
 fun set_tree :: "('a*'b) tree \<Rightarrow> 'a set" where
 "set_tree Leaf = {}" |
-"set_tree (Node l (a,_) r) = Set.insert a (set_tree l \<union> set_tree r)"
+"set_tree (Node l (a,_) r) = {a} \<union> set_tree l \<union> set_tree r"
 
 fun bst :: "('a::linorder*'b) tree \<Rightarrow> bool" where
 "bst Leaf = True" |
-"bst (Node l (a, _) r) = (bst l \<and> bst r \<and> (\<forall>x \<in> set_tree l. x < a) \<and> (\<forall>x \<in> set_tree r. a < x))"
+"bst (Node l (a, _) r) = ((\<forall>x \<in> set_tree l. x < a) \<and> (\<forall>x \<in> set_tree r. a < x) \<and> bst l \<and> bst r)"
 
 lemma finite_set_tree[simp]: "finite(set_tree t)"
 by(induction t) auto