# HG changeset patch # User nipkow # Date 1424557333 -3600 # Node ID 1a84beaa239bcae9de097a55260a7013594dc465 # Parent efd44495ecf8555936184c63722f8b31968ee6b5 added new tree material diff -r efd44495ecf8 -r 1a84beaa239b src/HOL/Library/Tree.thy --- a/src/HOL/Library/Tree.thy Fri Feb 20 16:59:25 2015 +0000 +++ b/src/HOL/Library/Tree.thy Sat Feb 21 23:22:13 2015 +0100 @@ -66,6 +66,36 @@ lemma (in linorder) bst_imp_sorted: "bst t \ sorted (inorder t)" by (induction t) (auto simp: sorted_append sorted_Cons intro: less_imp_le less_trans) +text{* In case there are 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)" + +lemma (in linorder) bst_eq_imp_sorted: "bst_eq t \ sorted (inorder t)" +apply (induction t) + apply(simp) +by (fastforce simp: sorted_append sorted_Cons intro: less_imp_le less_trans) + +subsection "Function @{text mirror}" + +fun mirror :: "'a tree \ 'a tree" where +"mirror \\ = Leaf" | +"mirror \l,x,r\ = \mirror r, x, mirror l\" + +lemma mirror_Leaf[simp]: "mirror t = \\ \ t = \\" +by (induction t) simp_all + +lemma size_mirror[simp]: "size(mirror t) = size t" +by (induction t) simp_all + +lemma size1_mirror[simp]: "size1(mirror t) = size1 t" +by (simp add: size1_def) + +lemma mirror_mirror[simp]: "mirror(mirror t) = t" +by (induction t) simp_all + subsection "Deletion of the rightmost entry"