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