# HG changeset patch # User nipkow # Date 1406306513 -7200 # Node ID cca7e8788481f2a69bac417a8facebbef92f76b9 # Parent 5b16e2370ccb10e20d2769a0250e3353d775ac07 added more functions and lemmas diff -r 5b16e2370ccb -r cca7e8788481 src/HOL/Library/Tree.thy --- a/src/HOL/Library/Tree.thy Fri Jul 25 17:13:30 2014 +0200 +++ b/src/HOL/Library/Tree.thy Fri Jul 25 18:41:53 2014 +0200 @@ -18,6 +18,12 @@ lemma set_tree_Node2: "set_tree(Node l x r) = insert x (set_tree l \ set_tree r)" by auto +lemma finite_set_tree[simp]: "finite(set_tree t)" +by(induction t) auto + + +subsection "The set of subtrees" + fun subtrees :: "'a tree \ 'a tree set" where "subtrees Leaf = {Leaf}" | "subtrees (Node l a r) = insert (Node l a r) (subtrees l \ subtrees r)" @@ -31,6 +37,9 @@ lemma in_set_tree_if: "Node l a r \ subtrees t \ a \ set_tree t" by (metis Node_notin_subtrees_if) + +subsection "Inorder list of entries" + fun inorder :: "'a tree \ 'a list" where "inorder Leaf = []" | "inorder (Node l x r) = inorder l @ [x] @ inorder r" @@ -38,6 +47,7 @@ lemma set_inorder[simp]: "set (inorder t) = set_tree t" by (induction t) auto + subsection {* Binary Search Tree predicate *} fun (in linorder) bst :: "'a tree \ bool" where @@ -47,4 +57,53 @@ 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) + +subsection "Deletion of the rightmost entry" + +fun del_rightmost :: "'a tree \ 'a tree * 'a" where +"del_rightmost (Node l a Leaf) = (l,a)" | +"del_rightmost (Node l a r) = (let (r',x) = del_rightmost r in (Node l a r', x))" + +lemma del_rightmost_set_tree_if_bst: + "\ del_rightmost t = (t',x); bst t; t \ Leaf \ + \ x \ set_tree t \ set_tree t' = set_tree t - {x}" +apply(induction t arbitrary: t' rule: del_rightmost.induct) + apply (fastforce simp: ball_Un split: prod.splits)+ +done + +lemma del_rightmost_set_tree: + "\ del_rightmost t = (t',x); t \ Leaf \ \ set_tree t = insert x (set_tree t')" +apply(induction t arbitrary: t' rule: del_rightmost.induct) +by (auto split: prod.splits) auto + +lemma del_rightmost_bst: + "\ del_rightmost t = (t',x); bst t; t \ Leaf \ \ bst t'" +proof(induction t arbitrary: t' rule: del_rightmost.induct) + case (2 l a rl b rr) + let ?r = "Node rl b rr" + from "2.prems"(1) obtain r' where 1: "del_rightmost ?r = (r',x)" and [simp]: "t' = Node l a r'" + by(simp split: prod.splits) + from "2.prems"(2) 1 del_rightmost_set_tree[OF 1] show ?case by(auto)(simp add: "2.IH") +qed auto + + +lemma del_rightmost_greater: "\ del_rightmost t = (t',x); bst t; t \ Leaf \ + \ \a\set_tree t'. a < x" +proof(induction t arbitrary: t' rule: del_rightmost.induct) + case (2 l a rl b rr) + from "2.prems"(1) obtain r' + where dm: "del_rightmost (Node rl b rr) = (r',x)" and [simp]: "t' = Node l a r'" + by(simp split: prod.splits) + show ?case using "2.prems"(2) "2.IH"[OF dm] del_rightmost_set_tree_if_bst[OF dm] + by (fastforce simp add: ball_Un) +qed simp_all + +(* should be moved but metis not available in much of Main *) +lemma Max_insert1: "\ finite A; \a\A. a \ x \ \ Max(insert x A) = x" +by (metis Max_in Max_insert Max_singleton antisym max_def) + +lemma del_rightmost_Max: + "\ del_rightmost t = (t',x); bst t; t \ Leaf \ \ x = Max(set_tree t)" +by (metis Max_insert1 del_rightmost_greater del_rightmost_set_tree finite_set_tree less_le_not_le) + end