src/HOL/ex/BT.thy
author boehmes
Mon, 17 Aug 2009 10:59:12 +0200
changeset 32381 11542bebe4d4
parent 23236 91d71bde1112
child 39246 9e58f0499f57
permissions -rw-r--r--
made Mirabelle a component
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1476
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1376
diff changeset
     1
(*  Title:      HOL/ex/BT.thy
1167
cbd32a0f2f41 New theory and proofs including preorder, inorder, ..., initially
lcp
parents:
diff changeset
     2
    ID:         $Id$
1476
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1376
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
1167
cbd32a0f2f41 New theory and proofs including preorder, inorder, ..., initially
lcp
parents:
diff changeset
     4
    Copyright   1995  University of Cambridge
cbd32a0f2f41 New theory and proofs including preorder, inorder, ..., initially
lcp
parents:
diff changeset
     5
20711
8b52cdaee86c fixed the definition of "depth"
paulson
parents: 19526
diff changeset
     6
Binary trees
1167
cbd32a0f2f41 New theory and proofs including preorder, inorder, ..., initially
lcp
parents:
diff changeset
     7
*)
cbd32a0f2f41 New theory and proofs including preorder, inorder, ..., initially
lcp
parents:
diff changeset
     8
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
     9
header {* Binary trees *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    10
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 11024
diff changeset
    11
theory BT imports Main begin
1167
cbd32a0f2f41 New theory and proofs including preorder, inorder, ..., initially
lcp
parents:
diff changeset
    12
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    13
datatype 'a bt =
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    14
    Lf
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    15
  | Br 'a  "'a bt"  "'a bt"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    16
1167
cbd32a0f2f41 New theory and proofs including preorder, inorder, ..., initially
lcp
parents:
diff changeset
    17
consts
19478
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
    18
  n_nodes   :: "'a bt => nat"
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
    19
  n_leaves  :: "'a bt => nat"
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
    20
  depth     :: "'a bt => nat"
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
    21
  reflect   :: "'a bt => 'a bt"
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
    22
  bt_map    :: "('a => 'b) => ('a bt => 'b bt)"
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
    23
  preorder  :: "'a bt => 'a list"
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
    24
  inorder   :: "'a bt => 'a list"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    25
  postorder :: "'a bt => 'a list"
23236
91d71bde1112 reverted appnd to append
haftmann
parents: 23023
diff changeset
    26
  append     :: "'a bt => 'a bt => 'a bt"
1167
cbd32a0f2f41 New theory and proofs including preorder, inorder, ..., initially
lcp
parents:
diff changeset
    27
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 1896
diff changeset
    28
primrec
19478
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
    29
  "n_nodes Lf = 0"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    30
  "n_nodes (Br a t1 t2) = Suc (n_nodes t1 + n_nodes t2)"
1167
cbd32a0f2f41 New theory and proofs including preorder, inorder, ..., initially
lcp
parents:
diff changeset
    31
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 1896
diff changeset
    32
primrec
19478
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
    33
  "n_leaves Lf = Suc 0"
1896
df4e40b9ff6d Simplified primrec - definitions.
berghofe
parents: 1476
diff changeset
    34
  "n_leaves (Br a t1 t2) = n_leaves t1 + n_leaves t2"
1167
cbd32a0f2f41 New theory and proofs including preorder, inorder, ..., initially
lcp
parents:
diff changeset
    35
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 1896
diff changeset
    36
primrec
19478
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
    37
  "depth Lf = 0"
20711
8b52cdaee86c fixed the definition of "depth"
paulson
parents: 19526
diff changeset
    38
  "depth (Br a t1 t2) = Suc (max (depth t1) (depth t2))"
19478
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
    39
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
    40
primrec
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
    41
  "reflect Lf = Lf"
1896
df4e40b9ff6d Simplified primrec - definitions.
berghofe
parents: 1476
diff changeset
    42
  "reflect (Br a t1 t2) = Br a (reflect t2) (reflect t1)"
1167
cbd32a0f2f41 New theory and proofs including preorder, inorder, ..., initially
lcp
parents:
diff changeset
    43
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 1896
diff changeset
    44
primrec
1896
df4e40b9ff6d Simplified primrec - definitions.
berghofe
parents: 1476
diff changeset
    45
  "bt_map f Lf = Lf"
df4e40b9ff6d Simplified primrec - definitions.
berghofe
parents: 1476
diff changeset
    46
  "bt_map f (Br a t1 t2) = Br (f a) (bt_map f t1) (bt_map f t2)"
1167
cbd32a0f2f41 New theory and proofs including preorder, inorder, ..., initially
lcp
parents:
diff changeset
    47
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 1896
diff changeset
    48
primrec
19478
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
    49
  "preorder Lf = []"
1896
df4e40b9ff6d Simplified primrec - definitions.
berghofe
parents: 1476
diff changeset
    50
  "preorder (Br a t1 t2) = [a] @ (preorder t1) @ (preorder t2)"
1167
cbd32a0f2f41 New theory and proofs including preorder, inorder, ..., initially
lcp
parents:
diff changeset
    51
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 1896
diff changeset
    52
primrec
19478
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
    53
  "inorder Lf = []"
1896
df4e40b9ff6d Simplified primrec - definitions.
berghofe
parents: 1476
diff changeset
    54
  "inorder (Br a t1 t2) = (inorder t1) @ [a] @ (inorder t2)"
1167
cbd32a0f2f41 New theory and proofs including preorder, inorder, ..., initially
lcp
parents:
diff changeset
    55
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 1896
diff changeset
    56
primrec
19478
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
    57
  "postorder Lf = []"
1896
df4e40b9ff6d Simplified primrec - definitions.
berghofe
parents: 1476
diff changeset
    58
  "postorder (Br a t1 t2) = (postorder t1) @ (postorder t2) @ [a]"
1167
cbd32a0f2f41 New theory and proofs including preorder, inorder, ..., initially
lcp
parents:
diff changeset
    59
19478
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
    60
primrec
23236
91d71bde1112 reverted appnd to append
haftmann
parents: 23023
diff changeset
    61
  "append Lf t = t"
91d71bde1112 reverted appnd to append
haftmann
parents: 23023
diff changeset
    62
  "append (Br a t1 t2) t = Br a (append t1 t) (append t2 t)"
19478
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
    63
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    64
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    65
text {* \medskip BT simplification *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    66
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    67
lemma n_leaves_reflect: "n_leaves (reflect t) = n_leaves t"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    68
  apply (induct t)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    69
   apply auto
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    70
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    71
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    72
lemma n_nodes_reflect: "n_nodes (reflect t) = n_nodes t"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    73
  apply (induct t)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    74
   apply auto
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    75
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    76
19478
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
    77
lemma depth_reflect: "depth (reflect t) = depth t"
20711
8b52cdaee86c fixed the definition of "depth"
paulson
parents: 19526
diff changeset
    78
  apply (induct t) 
8b52cdaee86c fixed the definition of "depth"
paulson
parents: 19526
diff changeset
    79
   apply auto
19478
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
    80
  done
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
    81
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    82
text {*
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    83
  The famous relationship between the numbers of leaves and nodes.
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    84
*}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    85
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    86
lemma n_leaves_nodes: "n_leaves t = Suc (n_nodes t)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    87
  apply (induct t)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    88
   apply auto
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    89
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    90
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    91
lemma reflect_reflect_ident: "reflect (reflect t) = t"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    92
  apply (induct t)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    93
   apply auto
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    94
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    95
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    96
lemma bt_map_reflect: "bt_map f (reflect t) = reflect (bt_map f t)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    97
  apply (induct t)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    98
   apply simp_all
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    99
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   100
19526
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   101
lemma preorder_bt_map: "preorder (bt_map f t) = map f (preorder t)"
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   102
  apply (induct t)
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   103
   apply simp_all
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   104
  done
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   105
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   106
lemma inorder_bt_map: "inorder (bt_map f t) = map f (inorder t)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   107
  apply (induct t)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   108
   apply simp_all
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   109
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   110
19526
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   111
lemma postorder_bt_map: "postorder (bt_map f t) = map f (postorder t)"
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   112
  apply (induct t)
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   113
   apply simp_all
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   114
  done
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   115
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   116
lemma depth_bt_map [simp]: "depth (bt_map f t) = depth t"
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   117
  apply (induct t)
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   118
   apply simp_all
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   119
  done
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   120
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   121
lemma n_leaves_bt_map [simp]: "n_leaves (bt_map f t) = n_leaves t"
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   122
  apply (induct t)
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   123
   apply (simp_all add: left_distrib)
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   124
  done
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   125
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   126
lemma preorder_reflect: "preorder (reflect t) = rev (postorder t)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   127
  apply (induct t)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   128
   apply simp_all
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   129
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   130
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   131
lemma inorder_reflect: "inorder (reflect t) = rev (inorder t)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   132
  apply (induct t)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   133
   apply simp_all
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   134
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   135
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   136
lemma postorder_reflect: "postorder (reflect t) = rev (preorder t)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   137
  apply (induct t)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   138
   apply simp_all
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   139
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   140
19478
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
   141
text {*
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
   142
 Analogues of the standard properties of the append function for lists.
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
   143
*}
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
   144
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
   145
lemma append_assoc [simp]:
23236
91d71bde1112 reverted appnd to append
haftmann
parents: 23023
diff changeset
   146
     "append (append t1 t2) t3 = append t1 (append t2 t3)"
19478
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
   147
  apply (induct t1)
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
   148
   apply simp_all
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
   149
  done
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
   150
23236
91d71bde1112 reverted appnd to append
haftmann
parents: 23023
diff changeset
   151
lemma append_Lf2 [simp]: "append t Lf = t"
19478
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
   152
  apply (induct t)
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
   153
   apply simp_all
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
   154
  done
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
   155
23236
91d71bde1112 reverted appnd to append
haftmann
parents: 23023
diff changeset
   156
lemma depth_append [simp]: "depth (append t1 t2) = depth t1 + depth t2"
19478
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
   157
  apply (induct t1)
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
   158
   apply (simp_all add: max_add_distrib_left)
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
   159
  done
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
   160
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
   161
lemma n_leaves_append [simp]:
23236
91d71bde1112 reverted appnd to append
haftmann
parents: 23023
diff changeset
   162
     "n_leaves (append t1 t2) = n_leaves t1 * n_leaves t2"
19478
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
   163
  apply (induct t1)
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
   164
   apply (simp_all add: left_distrib)
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
   165
  done
25778eacbe21 some new functions
paulson
parents: 16417
diff changeset
   166
19526
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   167
lemma bt_map_append:
23236
91d71bde1112 reverted appnd to append
haftmann
parents: 23023
diff changeset
   168
     "bt_map f (append t1 t2) = append (bt_map f t1) (bt_map f t2)"
19526
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   169
  apply (induct t1)
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   170
   apply simp_all
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   171
  done
90fb9e092e66 a few more examples
paulson
parents: 19478
diff changeset
   172
1167
cbd32a0f2f41 New theory and proofs including preorder, inorder, ..., initially
lcp
parents:
diff changeset
   173
end