src/HOL/Library/Tree.thy
author nipkow
Tue, 29 May 2018 14:05:59 +0200
changeset 68312 e9b5f25f6712
parent 68109 cebf36c14226
child 68998 818898556504
permissions -rw-r--r--
canonical names
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
57250
cddaf5b93728 new theory of binary trees
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
64887
266fb24c80bd tuned/minimized
nipkow
parents: 64771
diff changeset
     2
(* Todo: minimal ipl of balanced trees *)
57250
cddaf5b93728 new theory of binary trees
nipkow
parents:
diff changeset
     3
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59928
diff changeset
     4
section \<open>Binary Tree\<close>
57250
cddaf5b93728 new theory of binary trees
nipkow
parents:
diff changeset
     5
cddaf5b93728 new theory of binary trees
nipkow
parents:
diff changeset
     6
theory Tree
cddaf5b93728 new theory of binary trees
nipkow
parents:
diff changeset
     7
imports Main
cddaf5b93728 new theory of binary trees
nipkow
parents:
diff changeset
     8
begin
cddaf5b93728 new theory of binary trees
nipkow
parents:
diff changeset
     9
58424
cbbba613b6ab added nice standard syntax
nipkow
parents: 58310
diff changeset
    10
datatype 'a tree =
64887
266fb24c80bd tuned/minimized
nipkow
parents: 64771
diff changeset
    11
  Leaf ("\<langle>\<rangle>") |
266fb24c80bd tuned/minimized
nipkow
parents: 64771
diff changeset
    12
  Node "'a tree" (root_val: 'a) "'a tree" ("(1\<langle>_,/ _,/ _\<rangle>)")
57569
e20a999f7161 register tree with datatype_compat ot support QuickCheck
hoelzl
parents: 57530
diff changeset
    13
datatype_compat tree
57250
cddaf5b93728 new theory of binary trees
nipkow
parents:
diff changeset
    14
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59928
diff changeset
    15
text\<open>Can be seen as counting the number of leaves rather than nodes:\<close>
58438
566117a31cc0 added function size1
nipkow
parents: 58424
diff changeset
    16
566117a31cc0 added function size1
nipkow
parents: 58424
diff changeset
    17
definition size1 :: "'a tree \<Rightarrow> nat" where
566117a31cc0 added function size1
nipkow
parents: 58424
diff changeset
    18
"size1 t = size t + 1"
566117a31cc0 added function size1
nipkow
parents: 58424
diff changeset
    19
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    20
fun subtrees :: "'a tree \<Rightarrow> 'a tree set" where
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    21
"subtrees \<langle>\<rangle> = {\<langle>\<rangle>}" |
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    22
"subtrees (\<langle>l, a, r\<rangle>) = insert \<langle>l, a, r\<rangle> (subtrees l \<union> subtrees r)"
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    23
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    24
fun mirror :: "'a tree \<Rightarrow> 'a tree" where
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    25
"mirror \<langle>\<rangle> = Leaf" |
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    26
"mirror \<langle>l,x,r\<rangle> = \<langle>mirror r, x, mirror l\<rangle>"
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    27
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    28
class height = fixes height :: "'a \<Rightarrow> nat"
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    29
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    30
instantiation tree :: (type)height
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    31
begin
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    32
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    33
fun height_tree :: "'a tree => nat" where
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    34
"height Leaf = 0" |
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    35
"height (Node t1 a t2) = max (height t1) (height t2) + 1"
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    36
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    37
instance ..
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    38
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    39
end
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    40
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    41
fun min_height :: "'a tree \<Rightarrow> nat" where
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    42
"min_height Leaf = 0" |
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    43
"min_height (Node l _ r) = min (min_height l) (min_height r) + 1"
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    44
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    45
fun complete :: "'a tree \<Rightarrow> bool" where
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    46
"complete Leaf = True" |
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    47
"complete (Node l x r) = (complete l \<and> complete r \<and> height l = height r)"
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    48
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    49
definition balanced :: "'a tree \<Rightarrow> bool" where
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    50
"balanced t = (height t - min_height t \<le> 1)"
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    51
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    52
text \<open>Weight balanced:\<close>
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    53
fun wbalanced :: "'a tree \<Rightarrow> bool" where
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    54
"wbalanced Leaf = True" |
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    55
"wbalanced (Node l x r) = (abs(int(size l) - int(size r)) \<le> 1 \<and> wbalanced l \<and> wbalanced r)"
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    56
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    57
text \<open>Internal path length:\<close>
64887
266fb24c80bd tuned/minimized
nipkow
parents: 64771
diff changeset
    58
fun ipl :: "'a tree \<Rightarrow> nat" where
266fb24c80bd tuned/minimized
nipkow
parents: 64771
diff changeset
    59
"ipl Leaf = 0 " |
266fb24c80bd tuned/minimized
nipkow
parents: 64771
diff changeset
    60
"ipl (Node l _ r) = ipl l + size l + ipl r + size r"
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    61
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    62
fun preorder :: "'a tree \<Rightarrow> 'a list" where
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    63
"preorder \<langle>\<rangle> = []" |
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    64
"preorder \<langle>l, x, r\<rangle> = x # preorder l @ preorder r"
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    65
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    66
fun inorder :: "'a tree \<Rightarrow> 'a list" where
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    67
"inorder \<langle>\<rangle> = []" |
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    68
"inorder \<langle>l, x, r\<rangle> = inorder l @ [x] @ inorder r"
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    69
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    70
text\<open>A linear version avoiding append:\<close>
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    71
fun inorder2 :: "'a tree \<Rightarrow> 'a list \<Rightarrow> 'a list" where
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    72
"inorder2 \<langle>\<rangle> xs = xs" |
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    73
"inorder2 \<langle>l, x, r\<rangle> xs = inorder2 l (x # inorder2 r xs)"
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    74
64925
5eda89787621 added postorder
nipkow
parents: 64924
diff changeset
    75
fun postorder :: "'a tree \<Rightarrow> 'a list" where
5eda89787621 added postorder
nipkow
parents: 64924
diff changeset
    76
"postorder \<langle>\<rangle> = []" |
5eda89787621 added postorder
nipkow
parents: 64924
diff changeset
    77
"postorder \<langle>l, x, r\<rangle> = postorder l @ postorder r @ [x]"
5eda89787621 added postorder
nipkow
parents: 64924
diff changeset
    78
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    79
text\<open>Binary Search Tree:\<close>
66606
f23f044148d3 introduced bst_wrt
nipkow
parents: 65340
diff changeset
    80
fun bst_wrt :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a tree \<Rightarrow> bool" where
f23f044148d3 introduced bst_wrt
nipkow
parents: 65340
diff changeset
    81
"bst_wrt P \<langle>\<rangle> \<longleftrightarrow> True" |
f23f044148d3 introduced bst_wrt
nipkow
parents: 65340
diff changeset
    82
"bst_wrt P \<langle>l, a, r\<rangle> \<longleftrightarrow>
f23f044148d3 introduced bst_wrt
nipkow
parents: 65340
diff changeset
    83
 bst_wrt P l \<and> bst_wrt P r \<and> (\<forall>x\<in>set_tree l. P x a) \<and> (\<forall>x\<in>set_tree r. P a x)"
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    84
66606
f23f044148d3 introduced bst_wrt
nipkow
parents: 65340
diff changeset
    85
abbreviation bst :: "('a::linorder) tree \<Rightarrow> bool" where
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66659
diff changeset
    86
"bst \<equiv> bst_wrt (<)"
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    87
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    88
fun (in linorder) heap :: "'a tree \<Rightarrow> bool" where
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    89
"heap Leaf = True" |
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    90
"heap (Node l m r) =
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    91
  (heap l \<and> heap r \<and> (\<forall>x \<in> set_tree l \<union> set_tree r. m \<le> x))"
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    92
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
    93
65339
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
    94
subsection \<open>@{const map_tree}\<close>
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
    95
65340
nipkow
parents: 65339
diff changeset
    96
lemma eq_map_tree_Leaf[simp]: "map_tree f t = Leaf \<longleftrightarrow> t = Leaf"
65339
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
    97
by (rule tree.map_disc_iff)
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
    98
65340
nipkow
parents: 65339
diff changeset
    99
lemma eq_Leaf_map_tree[simp]: "Leaf = map_tree f t \<longleftrightarrow> t = Leaf"
65339
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   100
by (cases t) auto
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   101
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   102
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   103
subsection \<open>@{const size}\<close>
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   104
58438
566117a31cc0 added function size1
nipkow
parents: 58424
diff changeset
   105
lemma size1_simps[simp]:
566117a31cc0 added function size1
nipkow
parents: 58424
diff changeset
   106
  "size1 \<langle>\<rangle> = 1"
566117a31cc0 added function size1
nipkow
parents: 58424
diff changeset
   107
  "size1 \<langle>l, x, r\<rangle> = size1 l + size1 r"
566117a31cc0 added function size1
nipkow
parents: 58424
diff changeset
   108
by (simp_all add: size1_def)
566117a31cc0 added function size1
nipkow
parents: 58424
diff changeset
   109
62650
7e6bb43e7217 added tree lemmas
nipkow
parents: 62202
diff changeset
   110
lemma size1_ge0[simp]: "0 < size1 t"
7e6bb43e7217 added tree lemmas
nipkow
parents: 62202
diff changeset
   111
by (simp add: size1_def)
7e6bb43e7217 added tree lemmas
nipkow
parents: 62202
diff changeset
   112
65340
nipkow
parents: 65339
diff changeset
   113
lemma eq_size_0[simp]: "size t = 0 \<longleftrightarrow> t = Leaf"
65339
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   114
by(cases t) auto
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   115
65340
nipkow
parents: 65339
diff changeset
   116
lemma eq_0_size[simp]: "0 = size t \<longleftrightarrow> t = Leaf"
60505
9e6584184315 added funs and lemmas
nipkow
parents: 59928
diff changeset
   117
by(cases t) auto
9e6584184315 added funs and lemmas
nipkow
parents: 59928
diff changeset
   118
58424
cbbba613b6ab added nice standard syntax
nipkow
parents: 58310
diff changeset
   119
lemma neq_Leaf_iff: "(t \<noteq> \<langle>\<rangle>) = (\<exists>l a r. t = \<langle>l, a, r\<rangle>)"
cbbba613b6ab added nice standard syntax
nipkow
parents: 58310
diff changeset
   120
by (cases t) auto
57530
439f881c8744 added lemma
nipkow
parents: 57450
diff changeset
   121
59776
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   122
lemma size_map_tree[simp]: "size (map_tree f t) = size t"
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   123
by (induction t) auto
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   124
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   125
lemma size1_map_tree[simp]: "size1 (map_tree f t) = size1 t"
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   126
by (simp add: size1_def)
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   127
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   128
65339
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   129
subsection \<open>@{const set_tree}\<close>
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   130
65340
nipkow
parents: 65339
diff changeset
   131
lemma eq_set_tree_empty[simp]: "set_tree t = {} \<longleftrightarrow> t = Leaf"
65339
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   132
by (cases t) auto
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   133
65340
nipkow
parents: 65339
diff changeset
   134
lemma eq_empty_set_tree[simp]: "{} = set_tree t \<longleftrightarrow> t = Leaf"
65339
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   135
by (cases t) auto
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   136
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   137
lemma finite_set_tree[simp]: "finite(set_tree t)"
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   138
by(induction t) auto
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   139
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   140
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   141
subsection \<open>@{const subtrees}\<close>
60808
fd26519b1a6a depth -> height; removed del_rightmost (too specifi)
nipkow
parents: 60507
diff changeset
   142
65340
nipkow
parents: 65339
diff changeset
   143
lemma neq_subtrees_empty[simp]: "subtrees t \<noteq> {}"
nipkow
parents: 65339
diff changeset
   144
by (cases t)(auto)
nipkow
parents: 65339
diff changeset
   145
nipkow
parents: 65339
diff changeset
   146
lemma neq_empty_subtrees[simp]: "{} \<noteq> subtrees t"
nipkow
parents: 65339
diff changeset
   147
by (cases t)(auto)
nipkow
parents: 65339
diff changeset
   148
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   149
lemma set_treeE: "a \<in> set_tree t \<Longrightarrow> \<exists>l r. \<langle>l, a, r\<rangle> \<in> subtrees t"
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   150
by (induction t)(auto)
59776
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   151
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   152
lemma Node_notin_subtrees_if[simp]: "a \<notin> set_tree t \<Longrightarrow> Node l a r \<notin> subtrees t"
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   153
by (induction t) auto
59776
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   154
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   155
lemma in_set_tree_if: "\<langle>l, a, r\<rangle> \<in> subtrees t \<Longrightarrow> a \<in> set_tree t"
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   156
by (metis Node_notin_subtrees_if)
60808
fd26519b1a6a depth -> height; removed del_rightmost (too specifi)
nipkow
parents: 60507
diff changeset
   157
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   158
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   159
subsection \<open>@{const height} and @{const min_height}\<close>
60808
fd26519b1a6a depth -> height; removed del_rightmost (too specifi)
nipkow
parents: 60507
diff changeset
   160
65340
nipkow
parents: 65339
diff changeset
   161
lemma eq_height_0[simp]: "height t = 0 \<longleftrightarrow> t = Leaf"
65339
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   162
by(cases t) auto
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   163
65340
nipkow
parents: 65339
diff changeset
   164
lemma eq_0_height[simp]: "0 = height t \<longleftrightarrow> t = Leaf"
63665
15f48ce7ec23 added lemma
nipkow
parents: 63598
diff changeset
   165
by(cases t) auto
15f48ce7ec23 added lemma
nipkow
parents: 63598
diff changeset
   166
60808
fd26519b1a6a depth -> height; removed del_rightmost (too specifi)
nipkow
parents: 60507
diff changeset
   167
lemma height_map_tree[simp]: "height (map_tree f t) = height t"
59776
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   168
by (induction t) auto
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   169
64414
f8be2208e99c added lemma
nipkow
parents: 63861
diff changeset
   170
lemma height_le_size_tree: "height t \<le> size (t::'a tree)"
f8be2208e99c added lemma
nipkow
parents: 63861
diff changeset
   171
by (induction t) auto
f8be2208e99c added lemma
nipkow
parents: 63861
diff changeset
   172
64533
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   173
lemma size1_height: "size1 t \<le> 2 ^ height (t::'a tree)"
62202
e5bc7cbb0bcc added lemma
nipkow
parents: 62160
diff changeset
   174
proof(induction t)
e5bc7cbb0bcc added lemma
nipkow
parents: 62160
diff changeset
   175
  case (Node l a r)
e5bc7cbb0bcc added lemma
nipkow
parents: 62160
diff changeset
   176
  show ?case
e5bc7cbb0bcc added lemma
nipkow
parents: 62160
diff changeset
   177
  proof (cases "height l \<le> height r")
e5bc7cbb0bcc added lemma
nipkow
parents: 62160
diff changeset
   178
    case True
64533
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   179
    have "size1(Node l a r) = size1 l + size1 r" by simp
64918
nipkow
parents: 64887
diff changeset
   180
    also have "\<dots> \<le> 2 ^ height l + 2 ^ height r" using Node.IH by arith
nipkow
parents: 64887
diff changeset
   181
    also have "\<dots> \<le> 2 ^ height r + 2 ^ height r" using True by simp
64922
nipkow
parents: 64921
diff changeset
   182
    also have "\<dots> = 2 ^ height (Node l a r)"
64918
nipkow
parents: 64887
diff changeset
   183
      using True by (auto simp: max_def mult_2)
nipkow
parents: 64887
diff changeset
   184
    finally show ?thesis .
62202
e5bc7cbb0bcc added lemma
nipkow
parents: 62160
diff changeset
   185
  next
e5bc7cbb0bcc added lemma
nipkow
parents: 62160
diff changeset
   186
    case False
64533
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   187
    have "size1(Node l a r) = size1 l + size1 r" by simp
64918
nipkow
parents: 64887
diff changeset
   188
    also have "\<dots> \<le> 2 ^ height l + 2 ^ height r" using Node.IH by arith
nipkow
parents: 64887
diff changeset
   189
    also have "\<dots> \<le> 2 ^ height l + 2 ^ height l" using False by simp
62202
e5bc7cbb0bcc added lemma
nipkow
parents: 62160
diff changeset
   190
    finally show ?thesis using False by (auto simp: max_def mult_2)
e5bc7cbb0bcc added lemma
nipkow
parents: 62160
diff changeset
   191
  qed
e5bc7cbb0bcc added lemma
nipkow
parents: 62160
diff changeset
   192
qed simp
e5bc7cbb0bcc added lemma
nipkow
parents: 62160
diff changeset
   193
63755
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   194
corollary size_height: "size t \<le> 2 ^ height (t::'a tree) - 1"
64533
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   195
using size1_height[of t, unfolded size1_def] by(arith)
63755
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   196
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   197
lemma height_subtrees: "s \<in> subtrees t \<Longrightarrow> height s \<le> height t"
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   198
by (induction t) auto
57687
cca7e8788481 added more functions and lemmas
nipkow
parents: 57569
diff changeset
   199
63598
025d6e52d86f added min_height
nipkow
parents: 63413
diff changeset
   200
64540
f1f4ba6d02c9 spelling
nipkow
parents: 64533
diff changeset
   201
lemma min_height_le_height: "min_height t \<le> height t"
63598
025d6e52d86f added min_height
nipkow
parents: 63413
diff changeset
   202
by(induction t) auto
025d6e52d86f added min_height
nipkow
parents: 63413
diff changeset
   203
025d6e52d86f added min_height
nipkow
parents: 63413
diff changeset
   204
lemma min_height_map_tree[simp]: "min_height (map_tree f t) = min_height t"
025d6e52d86f added min_height
nipkow
parents: 63413
diff changeset
   205
by (induction t) auto
025d6e52d86f added min_height
nipkow
parents: 63413
diff changeset
   206
64533
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   207
lemma min_height_size1: "2 ^ min_height t \<le> size1 t"
63598
025d6e52d86f added min_height
nipkow
parents: 63413
diff changeset
   208
proof(induction t)
025d6e52d86f added min_height
nipkow
parents: 63413
diff changeset
   209
  case (Node l a r)
025d6e52d86f added min_height
nipkow
parents: 63413
diff changeset
   210
  have "(2::nat) ^ min_height (Node l a r) \<le> 2 ^ min_height l + 2 ^ min_height r"
025d6e52d86f added min_height
nipkow
parents: 63413
diff changeset
   211
    by (simp add: min_def)
64533
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   212
  also have "\<dots> \<le> size1(Node l a r)" using Node.IH by simp
63598
025d6e52d86f added min_height
nipkow
parents: 63413
diff changeset
   213
  finally show ?case .
025d6e52d86f added min_height
nipkow
parents: 63413
diff changeset
   214
qed simp
025d6e52d86f added min_height
nipkow
parents: 63413
diff changeset
   215
025d6e52d86f added min_height
nipkow
parents: 63413
diff changeset
   216
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   217
subsection \<open>@{const complete}\<close>
63036
1ba3aacfa4d3 added "balanced" predicate
nipkow
parents: 62650
diff changeset
   218
63755
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   219
lemma complete_iff_height: "complete t \<longleftrightarrow> (min_height t = height t)"
63598
025d6e52d86f added min_height
nipkow
parents: 63413
diff changeset
   220
apply(induction t)
025d6e52d86f added min_height
nipkow
parents: 63413
diff changeset
   221
 apply simp
025d6e52d86f added min_height
nipkow
parents: 63413
diff changeset
   222
apply (simp add: min_def max_def)
64540
f1f4ba6d02c9 spelling
nipkow
parents: 64533
diff changeset
   223
by (metis le_antisym le_trans min_height_le_height)
63598
025d6e52d86f added min_height
nipkow
parents: 63413
diff changeset
   224
63770
a67397b13eb5 added lemmas
nipkow
parents: 63765
diff changeset
   225
lemma size1_if_complete: "complete t \<Longrightarrow> size1 t = 2 ^ height t"
63036
1ba3aacfa4d3 added "balanced" predicate
nipkow
parents: 62650
diff changeset
   226
by (induction t) auto
1ba3aacfa4d3 added "balanced" predicate
nipkow
parents: 62650
diff changeset
   227
63755
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   228
lemma size_if_complete: "complete t \<Longrightarrow> size t = 2 ^ height t - 1"
63770
a67397b13eb5 added lemmas
nipkow
parents: 63765
diff changeset
   229
using size1_if_complete[simplified size1_def] by fastforce
a67397b13eb5 added lemmas
nipkow
parents: 63765
diff changeset
   230
64533
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   231
lemma complete_if_size1_height: "size1 t = 2 ^ height t \<Longrightarrow> complete t"
63770
a67397b13eb5 added lemmas
nipkow
parents: 63765
diff changeset
   232
proof (induct "height t" arbitrary: t)
65340
nipkow
parents: 65339
diff changeset
   233
  case 0 thus ?case by (simp)
63770
a67397b13eb5 added lemmas
nipkow
parents: 63765
diff changeset
   234
next
a67397b13eb5 added lemmas
nipkow
parents: 63765
diff changeset
   235
  case (Suc h)
a67397b13eb5 added lemmas
nipkow
parents: 63765
diff changeset
   236
  hence "t \<noteq> Leaf" by auto
a67397b13eb5 added lemmas
nipkow
parents: 63765
diff changeset
   237
  then obtain l a r where [simp]: "t = Node l a r"
a67397b13eb5 added lemmas
nipkow
parents: 63765
diff changeset
   238
    by (auto simp: neq_Leaf_iff)
a67397b13eb5 added lemmas
nipkow
parents: 63765
diff changeset
   239
  have 1: "height l \<le> h" and 2: "height r \<le> h" using Suc(2) by(auto)
64533
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   240
  have 3: "\<not> height l < h"
63770
a67397b13eb5 added lemmas
nipkow
parents: 63765
diff changeset
   241
  proof
a67397b13eb5 added lemmas
nipkow
parents: 63765
diff changeset
   242
    assume 0: "height l < h"
64533
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   243
    have "size1 t = size1 l + size1 r" by simp
64918
nipkow
parents: 64887
diff changeset
   244
    also have "\<dots> \<le> 2 ^ height l + 2 ^ height r"
nipkow
parents: 64887
diff changeset
   245
      using size1_height[of l] size1_height[of r] by arith
nipkow
parents: 64887
diff changeset
   246
    also have " \<dots> < 2 ^ h + 2 ^ height r" using 0 by (simp)
nipkow
parents: 64887
diff changeset
   247
    also have " \<dots> \<le> 2 ^ h + 2 ^ h" using 2 by (simp)
nipkow
parents: 64887
diff changeset
   248
    also have "\<dots> = 2 ^ (Suc h)" by (simp)
64533
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   249
    also have "\<dots> = size1 t" using Suc(2,3) by simp
64918
nipkow
parents: 64887
diff changeset
   250
    finally have "size1 t < size1 t" .
nipkow
parents: 64887
diff changeset
   251
    thus False by (simp)
63770
a67397b13eb5 added lemmas
nipkow
parents: 63765
diff changeset
   252
  qed
64918
nipkow
parents: 64887
diff changeset
   253
  have 4: "\<not> height r < h"
63770
a67397b13eb5 added lemmas
nipkow
parents: 63765
diff changeset
   254
  proof
a67397b13eb5 added lemmas
nipkow
parents: 63765
diff changeset
   255
    assume 0: "height r < h"
64533
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   256
    have "size1 t = size1 l + size1 r" by simp
64918
nipkow
parents: 64887
diff changeset
   257
    also have "\<dots> \<le> 2 ^ height l + 2 ^ height r"
nipkow
parents: 64887
diff changeset
   258
      using size1_height[of l] size1_height[of r] by arith
nipkow
parents: 64887
diff changeset
   259
    also have " \<dots> < 2 ^ height l + 2 ^ h" using 0 by (simp)
nipkow
parents: 64887
diff changeset
   260
    also have " \<dots> \<le> 2 ^ h + 2 ^ h" using 1 by (simp)
nipkow
parents: 64887
diff changeset
   261
    also have "\<dots> = 2 ^ (Suc h)" by (simp)
64533
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   262
    also have "\<dots> = size1 t" using Suc(2,3) by simp
64918
nipkow
parents: 64887
diff changeset
   263
    finally have "size1 t < size1 t" .
nipkow
parents: 64887
diff changeset
   264
    thus False by (simp)
63770
a67397b13eb5 added lemmas
nipkow
parents: 63765
diff changeset
   265
  qed
a67397b13eb5 added lemmas
nipkow
parents: 63765
diff changeset
   266
  from 1 2 3 4 have *: "height l = h" "height r = h" by linarith+
64533
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   267
  hence "size1 l = 2 ^ height l" "size1 r = 2 ^ height r"
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   268
    using Suc(3) size1_height[of l] size1_height[of r] by (auto)
63770
a67397b13eb5 added lemmas
nipkow
parents: 63765
diff changeset
   269
  with * Suc(1) show ?case by simp
a67397b13eb5 added lemmas
nipkow
parents: 63765
diff changeset
   270
qed
a67397b13eb5 added lemmas
nipkow
parents: 63765
diff changeset
   271
64533
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   272
text\<open>The following proof involves \<open>\<ge>\<close>/\<open>>\<close> chains rather than the standard
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   273
\<open>\<le>\<close>/\<open><\<close> chains. To chain the elements together the transitivity rules \<open>xtrans\<close>
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   274
are used.\<close>
63755
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   275
64533
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   276
lemma complete_if_size1_min_height: "size1 t = 2 ^ min_height t \<Longrightarrow> complete t"
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   277
proof (induct "min_height t" arbitrary: t)
65340
nipkow
parents: 65339
diff changeset
   278
  case 0 thus ?case by (simp add: size1_def)
63755
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   279
next
64533
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   280
  case (Suc h)
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   281
  hence "t \<noteq> Leaf" by auto
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   282
  then obtain l a r where [simp]: "t = Node l a r"
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   283
    by (auto simp: neq_Leaf_iff)
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   284
  have 1: "h \<le> min_height l" and 2: "h \<le> min_height r" using Suc(2) by(auto)
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   285
  have 3: "\<not> h < min_height l"
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   286
  proof
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   287
    assume 0: "h < min_height l"
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   288
    have "size1 t = size1 l + size1 r" by simp
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   289
    also note min_height_size1[of l]
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   290
    also(xtrans) note min_height_size1[of r]
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   291
    also(xtrans) have "(2::nat) ^ min_height l > 2 ^ h"
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   292
        using 0 by (simp add: diff_less_mono)
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   293
    also(xtrans) have "(2::nat) ^ min_height r \<ge> 2 ^ h" using 2 by simp
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   294
    also(xtrans) have "(2::nat) ^ h + 2 ^ h = 2 ^ (Suc h)" by (simp)
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   295
    also have "\<dots> = size1 t" using Suc(2,3) by simp
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   296
    finally show False by (simp add: diff_le_mono)
63755
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   297
  qed
64533
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   298
  have 4: "\<not> h < min_height r"
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   299
  proof
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   300
    assume 0: "h < min_height r"
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   301
    have "size1 t = size1 l + size1 r" by simp
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   302
    also note min_height_size1[of l]
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   303
    also(xtrans) note min_height_size1[of r]
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   304
    also(xtrans) have "(2::nat) ^ min_height r > 2 ^ h"
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   305
        using 0 by (simp add: diff_less_mono)
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   306
    also(xtrans) have "(2::nat) ^ min_height l \<ge> 2 ^ h" using 1 by simp
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   307
    also(xtrans) have "(2::nat) ^ h + 2 ^ h = 2 ^ (Suc h)" by (simp)
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   308
    also have "\<dots> = size1 t" using Suc(2,3) by simp
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   309
    finally show False by (simp add: diff_le_mono)
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   310
  qed
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   311
  from 1 2 3 4 have *: "min_height l = h" "min_height r = h" by linarith+
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   312
  hence "size1 l = 2 ^ min_height l" "size1 r = 2 ^ min_height r"
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   313
    using Suc(3) min_height_size1[of l] min_height_size1[of r] by (auto)
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   314
  with * Suc(1) show ?case
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   315
    by (simp add: complete_iff_height)
63755
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   316
qed
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   317
64533
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   318
lemma complete_iff_size1: "complete t \<longleftrightarrow> size1 t = 2 ^ height t"
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   319
using complete_if_size1_height size1_if_complete by blast
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   320
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   321
text\<open>Better bounds for incomplete trees:\<close>
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   322
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   323
lemma size1_height_if_incomplete:
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   324
  "\<not> complete t \<Longrightarrow> size1 t < 2 ^ height t"
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   325
by (meson antisym_conv complete_iff_size1 not_le size1_height)
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   326
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   327
lemma min_height_size1_if_incomplete:
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   328
  "\<not> complete t \<Longrightarrow> 2 ^ min_height t < size1 t"
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   329
by (metis complete_if_size1_min_height le_less min_height_size1)
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   330
63755
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   331
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   332
subsection \<open>@{const balanced}\<close>
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   333
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   334
lemma balanced_subtreeL: "balanced (Node l x r) \<Longrightarrow> balanced l"
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   335
by(simp add: balanced_def)
63755
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   336
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   337
lemma balanced_subtreeR: "balanced (Node l x r) \<Longrightarrow> balanced r"
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   338
by(simp add: balanced_def)
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   339
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   340
lemma balanced_subtrees: "\<lbrakk> balanced t; s \<in> subtrees t \<rbrakk> \<Longrightarrow> balanced s"
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   341
using [[simp_depth_limit=1]]
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   342
by(induction t arbitrary: s)
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   343
  (auto simp add: balanced_subtreeL balanced_subtreeR)
63755
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   344
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   345
text\<open>Balanced trees have optimal height:\<close>
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   346
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   347
lemma balanced_optimal:
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   348
fixes t :: "'a tree" and t' :: "'b tree"
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   349
assumes "balanced t" "size t \<le> size t'" shows "height t \<le> height t'"
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   350
proof (cases "complete t")
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   351
  case True
64924
nipkow
parents: 64923
diff changeset
   352
  have "(2::nat) ^ height t \<le> 2 ^ height t'"
63755
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   353
  proof -
64924
nipkow
parents: 64923
diff changeset
   354
    have "2 ^ height t = size1 t"
nipkow
parents: 64923
diff changeset
   355
      using True by (simp add: complete_iff_height size1_if_complete)
nipkow
parents: 64923
diff changeset
   356
    also have "\<dots> \<le> size1 t'" using assms(2) by(simp add: size1_def)
nipkow
parents: 64923
diff changeset
   357
    also have "\<dots> \<le> 2 ^ height t'" by (rule size1_height)
63755
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   358
    finally show ?thesis .
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   359
  qed
64924
nipkow
parents: 64923
diff changeset
   360
  thus ?thesis by (simp)
63755
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   361
next
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   362
  case False
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   363
  have "(2::nat) ^ min_height t < 2 ^ height t'"
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   364
  proof -
64533
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   365
    have "(2::nat) ^ min_height t < size1 t"
172f3a047f4a more lemmas, tuned proofs
nipkow
parents: 64414
diff changeset
   366
      by(rule min_height_size1_if_incomplete[OF False])
64918
nipkow
parents: 64887
diff changeset
   367
    also have "\<dots> \<le> size1 t'" using assms(2) by (simp add: size1_def)
nipkow
parents: 64887
diff changeset
   368
    also have "\<dots> \<le> 2 ^ height t'"  by(rule size1_height)
nipkow
parents: 64887
diff changeset
   369
    finally have "(2::nat) ^ min_height t < (2::nat) ^ height t'" .
64924
nipkow
parents: 64923
diff changeset
   370
    thus ?thesis .
63755
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   371
  qed
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   372
  hence *: "min_height t < height t'" by simp
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   373
  have "min_height t + 1 = height t"
64540
f1f4ba6d02c9 spelling
nipkow
parents: 64533
diff changeset
   374
    using min_height_le_height[of t] assms(1) False
63829
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63770
diff changeset
   375
    by (simp add: complete_iff_height balanced_def)
63755
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   376
  with * show ?thesis by arith
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63665
diff changeset
   377
qed
63036
1ba3aacfa4d3 added "balanced" predicate
nipkow
parents: 62650
diff changeset
   378
1ba3aacfa4d3 added "balanced" predicate
nipkow
parents: 62650
diff changeset
   379
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   380
subsection \<open>@{const wbalanced}\<close>
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   381
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   382
lemma wbalanced_subtrees: "\<lbrakk> wbalanced t; s \<in> subtrees t \<rbrakk> \<Longrightarrow> wbalanced s"
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   383
using [[simp_depth_limit=1]] by(induction t arbitrary: s) auto
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   384
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   385
64887
266fb24c80bd tuned/minimized
nipkow
parents: 64771
diff changeset
   386
subsection \<open>@{const ipl}\<close>
63413
9fe2d9dc095e added path_len
nipkow
parents: 63036
diff changeset
   387
9fe2d9dc095e added path_len
nipkow
parents: 63036
diff changeset
   388
text \<open>The internal path length of a tree:\<close>
9fe2d9dc095e added path_len
nipkow
parents: 63036
diff changeset
   389
64923
7c340dcbc323 int version slicker
nipkow
parents: 64922
diff changeset
   390
lemma ipl_if_complete_int:
7c340dcbc323 int version slicker
nipkow
parents: 64922
diff changeset
   391
  "complete t \<Longrightarrow> int(ipl t) = (int(height t) - 2) * 2^(height t) + 2"
7c340dcbc323 int version slicker
nipkow
parents: 64922
diff changeset
   392
apply(induction t)
7c340dcbc323 int version slicker
nipkow
parents: 64922
diff changeset
   393
 apply simp
7c340dcbc323 int version slicker
nipkow
parents: 64922
diff changeset
   394
apply simp
7c340dcbc323 int version slicker
nipkow
parents: 64922
diff changeset
   395
apply (simp add: algebra_simps size_if_complete of_nat_diff)
7c340dcbc323 int version slicker
nipkow
parents: 64922
diff changeset
   396
done
63413
9fe2d9dc095e added path_len
nipkow
parents: 63036
diff changeset
   397
9fe2d9dc095e added path_len
nipkow
parents: 63036
diff changeset
   398
59776
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   399
subsection "List of entries"
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   400
65340
nipkow
parents: 65339
diff changeset
   401
lemma eq_inorder_Nil[simp]: "inorder t = [] \<longleftrightarrow> t = Leaf"
65339
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   402
by (cases t) auto
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   403
65340
nipkow
parents: 65339
diff changeset
   404
lemma eq_Nil_inorder[simp]: "[] = inorder t \<longleftrightarrow> t = Leaf"
65339
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   405
by (cases t) auto
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   406
57449
f81da03b9ebd Library/Tree: use datatype_new, bst is an inductive predicate
hoelzl
parents: 57250
diff changeset
   407
lemma set_inorder[simp]: "set (inorder t) = set_tree t"
58424
cbbba613b6ab added nice standard syntax
nipkow
parents: 58310
diff changeset
   408
by (induction t) auto
57250
cddaf5b93728 new theory of binary trees
nipkow
parents:
diff changeset
   409
59776
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   410
lemma set_preorder[simp]: "set (preorder t) = set_tree t"
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   411
by (induction t) auto
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   412
64925
5eda89787621 added postorder
nipkow
parents: 64924
diff changeset
   413
lemma set_postorder[simp]: "set (postorder t) = set_tree t"
5eda89787621 added postorder
nipkow
parents: 64924
diff changeset
   414
by (induction t) auto
5eda89787621 added postorder
nipkow
parents: 64924
diff changeset
   415
59776
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   416
lemma length_preorder[simp]: "length (preorder t) = size t"
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   417
by (induction t) auto
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   418
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   419
lemma length_inorder[simp]: "length (inorder t) = size t"
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   420
by (induction t) auto
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   421
64925
5eda89787621 added postorder
nipkow
parents: 64924
diff changeset
   422
lemma length_postorder[simp]: "length (postorder t) = size t"
5eda89787621 added postorder
nipkow
parents: 64924
diff changeset
   423
by (induction t) auto
5eda89787621 added postorder
nipkow
parents: 64924
diff changeset
   424
59776
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   425
lemma preorder_map: "preorder (map_tree f t) = map f (preorder t)"
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   426
by (induction t) auto
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   427
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   428
lemma inorder_map: "inorder (map_tree f t) = map f (inorder t)"
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   429
by (induction t) auto
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   430
64925
5eda89787621 added postorder
nipkow
parents: 64924
diff changeset
   431
lemma postorder_map: "postorder (map_tree f t) = map f (postorder t)"
5eda89787621 added postorder
nipkow
parents: 64924
diff changeset
   432
by (induction t) auto
5eda89787621 added postorder
nipkow
parents: 64924
diff changeset
   433
63765
e60020520b15 added inorder2
nipkow
parents: 63755
diff changeset
   434
lemma inorder2_inorder: "inorder2 t xs = inorder t @ xs"
e60020520b15 added inorder2
nipkow
parents: 63755
diff changeset
   435
by (induction t arbitrary: xs) auto
e60020520b15 added inorder2
nipkow
parents: 63755
diff changeset
   436
57687
cca7e8788481 added more functions and lemmas
nipkow
parents: 57569
diff changeset
   437
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   438
subsection \<open>Binary Search Tree\<close>
59561
1a84beaa239b added new tree material
nipkow
parents: 58881
diff changeset
   439
66606
f23f044148d3 introduced bst_wrt
nipkow
parents: 65340
diff changeset
   440
lemma bst_wrt_mono: "(\<And>x y. P x y \<Longrightarrow> Q x y) \<Longrightarrow> bst_wrt P t \<Longrightarrow> bst_wrt Q t"
59928
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents: 59776
diff changeset
   441
by (induction t) (auto)
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents: 59776
diff changeset
   442
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66659
diff changeset
   443
lemma bst_wrt_le_if_bst: "bst t \<Longrightarrow> bst_wrt (\<le>) t"
66606
f23f044148d3 introduced bst_wrt
nipkow
parents: 65340
diff changeset
   444
using bst_wrt_mono less_imp_le by blast
f23f044148d3 introduced bst_wrt
nipkow
parents: 65340
diff changeset
   445
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66659
diff changeset
   446
lemma bst_wrt_le_iff_sorted: "bst_wrt (\<le>) t \<longleftrightarrow> sorted (inorder t)"
59561
1a84beaa239b added new tree material
nipkow
parents: 58881
diff changeset
   447
apply (induction t)
1a84beaa239b added new tree material
nipkow
parents: 58881
diff changeset
   448
 apply(simp)
68109
cebf36c14226 new def of sorted and sorted_wrt
nipkow
parents: 67399
diff changeset
   449
by (fastforce simp: sorted_append intro: less_imp_le less_trans)
59561
1a84beaa239b added new tree material
nipkow
parents: 58881
diff changeset
   450
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66659
diff changeset
   451
lemma bst_iff_sorted_wrt_less: "bst t \<longleftrightarrow> sorted_wrt (<) (inorder t)"
59928
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents: 59776
diff changeset
   452
apply (induction t)
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents: 59776
diff changeset
   453
 apply simp
68109
cebf36c14226 new def of sorted and sorted_wrt
nipkow
parents: 67399
diff changeset
   454
apply (fastforce simp: sorted_wrt_append)
59928
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents: 59776
diff changeset
   455
done
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents: 59776
diff changeset
   456
59776
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   457
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   458
subsection \<open>@{const heap}\<close>
60505
9e6584184315 added funs and lemmas
nipkow
parents: 59928
diff changeset
   459
9e6584184315 added funs and lemmas
nipkow
parents: 59928
diff changeset
   460
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63829
diff changeset
   461
subsection \<open>@{const mirror}\<close>
59561
1a84beaa239b added new tree material
nipkow
parents: 58881
diff changeset
   462
1a84beaa239b added new tree material
nipkow
parents: 58881
diff changeset
   463
lemma mirror_Leaf[simp]: "mirror t = \<langle>\<rangle> \<longleftrightarrow> t = \<langle>\<rangle>"
1a84beaa239b added new tree material
nipkow
parents: 58881
diff changeset
   464
by (induction t) simp_all
1a84beaa239b added new tree material
nipkow
parents: 58881
diff changeset
   465
65339
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   466
lemma Leaf_mirror[simp]: "\<langle>\<rangle> = mirror t \<longleftrightarrow> t = \<langle>\<rangle>"
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   467
using mirror_Leaf by fastforce
c4531ddafe72 more lemmas
nipkow
parents: 64925
diff changeset
   468
59561
1a84beaa239b added new tree material
nipkow
parents: 58881
diff changeset
   469
lemma size_mirror[simp]: "size(mirror t) = size t"
1a84beaa239b added new tree material
nipkow
parents: 58881
diff changeset
   470
by (induction t) simp_all
1a84beaa239b added new tree material
nipkow
parents: 58881
diff changeset
   471
1a84beaa239b added new tree material
nipkow
parents: 58881
diff changeset
   472
lemma size1_mirror[simp]: "size1(mirror t) = size1 t"
1a84beaa239b added new tree material
nipkow
parents: 58881
diff changeset
   473
by (simp add: size1_def)
1a84beaa239b added new tree material
nipkow
parents: 58881
diff changeset
   474
60808
fd26519b1a6a depth -> height; removed del_rightmost (too specifi)
nipkow
parents: 60507
diff changeset
   475
lemma height_mirror[simp]: "height(mirror t) = height t"
59776
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   476
by (induction t) simp_all
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   477
66659
d5bf4bdb4fb7 added lemmas
nipkow
parents: 66606
diff changeset
   478
lemma min_height_mirror [simp]: "min_height (mirror t) = min_height t"
d5bf4bdb4fb7 added lemmas
nipkow
parents: 66606
diff changeset
   479
by (induction t) simp_all  
d5bf4bdb4fb7 added lemmas
nipkow
parents: 66606
diff changeset
   480
d5bf4bdb4fb7 added lemmas
nipkow
parents: 66606
diff changeset
   481
lemma ipl_mirror [simp]: "ipl (mirror t) = ipl t"
d5bf4bdb4fb7 added lemmas
nipkow
parents: 66606
diff changeset
   482
by (induction t) simp_all
d5bf4bdb4fb7 added lemmas
nipkow
parents: 66606
diff changeset
   483
59776
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   484
lemma inorder_mirror: "inorder(mirror t) = rev(inorder t)"
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   485
by (induction t) simp_all
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   486
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   487
lemma map_mirror: "map_tree f (mirror t) = mirror (map_tree f t)"
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   488
by (induction t) simp_all
f54af3307334 added funs and lemmas
nipkow
parents: 59561
diff changeset
   489
59561
1a84beaa239b added new tree material
nipkow
parents: 58881
diff changeset
   490
lemma mirror_mirror[simp]: "mirror(mirror t) = t"
1a84beaa239b added new tree material
nipkow
parents: 58881
diff changeset
   491
by (induction t) simp_all
1a84beaa239b added new tree material
nipkow
parents: 58881
diff changeset
   492
57250
cddaf5b93728 new theory of binary trees
nipkow
parents:
diff changeset
   493
end