# HG changeset patch # User nipkow # Date 1411549745 -7200 # Node ID cbbba613b6abf372328feabcee602b211417e660 # Parent e4d540c0dd57509c14c980ecb319d89a8085af7a added nice standard syntax diff -r e4d540c0dd57 -r cbbba613b6ab src/HOL/Library/Tree.thy --- a/src/HOL/Library/Tree.thy Mon Sep 22 21:45:59 2014 +0200 +++ b/src/HOL/Library/Tree.thy Wed Sep 24 11:09:05 2014 +0200 @@ -6,17 +6,16 @@ imports Main begin -datatype 'a tree = Leaf | Node (left: "'a tree") (val: 'a) (right: "'a tree") +datatype 'a tree = + Leaf ("\\") | + Node (left: "'a tree") (val: 'a) (right: "'a tree") ("\_, _, _\") where "left Leaf = Leaf" | "right Leaf = Leaf" datatype_compat tree -lemma neq_Leaf_iff: "(t \ Leaf) = (\l a r. t = Node l a r)" - by (cases t) auto - -lemma set_tree_Node2: "set_tree(Node l x r) = insert x (set_tree l \ set_tree r)" -by auto +lemma neq_Leaf_iff: "(t \ \\) = (\l a r. t = \l, a, r\)" +by (cases t) auto lemma finite_set_tree[simp]: "finite(set_tree t)" by(induction t) auto @@ -25,44 +24,44 @@ 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)" + "subtrees \\ = {\\}" | + "subtrees (\l, a, r\) = insert \l, a, r\ (subtrees l \ subtrees r)" -lemma set_treeE: "a \ set_tree t \ \l r. Node l a r \ subtrees t" - by (induction t)(auto) +lemma set_treeE: "a \ set_tree t \ \l r. \l, a, r\ \ subtrees t" +by (induction t)(auto) lemma Node_notin_subtrees_if[simp]: "a \ set_tree t \ Node l a r \ subtrees t" - by (induction t) auto +by (induction t) auto -lemma in_set_tree_if: "Node l a r \ subtrees t \ a \ set_tree t" - by (metis Node_notin_subtrees_if) +lemma in_set_tree_if: "\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" +"inorder \\ = []" | +"inorder \l, x, r\ = inorder l @ [x] @ inorder r" lemma set_inorder[simp]: "set (inorder t) = set_tree t" - by (induction t) auto +by (induction t) auto subsection {* Binary Search Tree predicate *} fun (in linorder) bst :: "'a tree \ bool" where - "bst Leaf \ True" | - "bst (Node l a r) \ bst l \ bst r \ (\x\set_tree l. x < a) \ (\x\set_tree r. a < x)" +"bst \\ \ True" | +"bst \l, a, r\ \ bst l \ bst r \ (\x\set_tree l. x < a) \ (\x\set_tree r. a < x)" 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) +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))" +"del_rightmost \l, a, \\\ = (l,a)" | +"del_rightmost \l, a, r\ = (let (r',x) = del_rightmost r in (\l, a, r'\, x))" lemma del_rightmost_set_tree_if_bst: "\ del_rightmost t = (t',x); bst t; t \ Leaf \ @@ -72,12 +71,12 @@ done lemma del_rightmost_set_tree: - "\ del_rightmost t = (t',x); t \ Leaf \ \ set_tree t = insert x (set_tree t')" + "\ del_rightmost t = (t',x); t \ \\ \ \ 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'" + "\ del_rightmost t = (t',x); bst t; t \ \\ \ \ bst t'" proof(induction t arbitrary: t' rule: del_rightmost.induct) case (2 l a rl b rr) let ?r = "Node rl b rr" @@ -87,7 +86,7 @@ qed auto -lemma del_rightmost_greater: "\ del_rightmost t = (t',x); bst t; t \ Leaf \ +lemma del_rightmost_greater: "\ del_rightmost t = (t',x); bst t; t \ \\ \ \ \a\set_tree t'. a < x" proof(induction t arbitrary: t' rule: del_rightmost.induct) case (2 l a rl b rr) @@ -103,7 +102,7 @@ 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)" + "\ del_rightmost t = (t',x); bst t; t \ \\ \ \ x = Max(set_tree t)" by (metis Max_insert1 del_rightmost_greater del_rightmost_set_tree finite_set_tree less_le_not_le) end