# HG changeset patch # User nipkow # Date 1504634632 -7200 # Node ID a8edca8c4a680ded543a45a69f50cba7d8bbbbf7 # Parent 261dcd52c5a0b4dd79e6b33ec1d2b1f2c0fc0ad1# Parent f23f044148d308da3a597a047c49a679c6651772 merged diff -r 261dcd52c5a0 -r a8edca8c4a68 src/HOL/Library/Tree.thy --- a/src/HOL/Library/Tree.thy Tue Sep 05 16:45:23 2017 +0200 +++ b/src/HOL/Library/Tree.thy Tue Sep 05 20:03:52 2017 +0200 @@ -77,15 +77,13 @@ "postorder \l, x, r\ = postorder l @ postorder r @ [x]" text\Binary Search Tree:\ -fun (in linorder) bst :: "'a tree \ bool" where -"bst \\ \ True" | -"bst \l, a, r\ \ bst l \ bst r \ (\x\set_tree l. x < a) \ (\x\set_tree r. a < x)" +fun bst_wrt :: "('a \ 'a \ bool) \ 'a tree \ bool" where +"bst_wrt P \\ \ True" | +"bst_wrt P \l, a, r\ \ + bst_wrt P l \ bst_wrt P r \ (\x\set_tree l. P x a) \ (\x\set_tree r. P a x)" -text\Binary Search Tree with duplicates:\ -fun (in linorder) bst_eq :: "'a tree \ bool" where -"bst_eq \\ \ True" | -"bst_eq \l,a,r\ \ - bst_eq l \ bst_eq r \ (\x\set_tree l. x \ a) \ (\x\set_tree r. a \ x)" +abbreviation bst :: "('a::linorder) tree \ bool" where +"bst \ bst_wrt (op <)" fun (in linorder) heap :: "'a tree \ bool" where "heap Leaf = True" | @@ -439,24 +437,21 @@ subsection \Binary Search Tree\ -lemma (in linorder) bst_eq_if_bst: "bst t \ bst_eq t" +lemma bst_wrt_mono: "(\x y. P x y \ Q x y) \ bst_wrt P t \ bst_wrt Q t" by (induction t) (auto) -lemma (in linorder) bst_eq_imp_sorted: "bst_eq t \ sorted (inorder t)" +lemma bst_wrt_le_if_bst: "bst t \ bst_wrt (op \) t" +using bst_wrt_mono less_imp_le by blast + +lemma bst_wrt_le_iff_sorted: "bst_wrt (op \) t \ sorted (inorder t)" apply (induction t) apply(simp) by (fastforce simp: sorted_append sorted_Cons intro: less_imp_le less_trans) -lemma (in linorder) distinct_preorder_if_bst: "bst t \ distinct (preorder t)" +lemma bst_iff_sorted_wrt_less: "bst t \ sorted_wrt (op <) (inorder t)" apply (induction t) apply simp -apply(fastforce elim: order.asym) -done - -lemma (in linorder) distinct_inorder_if_bst: "bst t \ distinct (inorder t)" -apply (induction t) - apply simp -apply(fastforce elim: order.asym) +apply (fastforce simp: sorted_wrt_Cons sorted_wrt_append) done