src/HOL/ex/BT.thy
author nipkow
Wed, 18 Aug 2004 11:09:40 +0200
changeset 15140 322485b816ac
parent 11024 23bf8d787b04
child 16417 9bc16273c2d4
permissions -rw-r--r--
import -> imports
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
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
     6
Binary trees (based on the ZF version).
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
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    11
theory BT = Main:
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
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    18
  n_nodes :: "'a bt => nat"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    19
  n_leaves :: "'a bt => nat"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    20
  reflect :: "'a bt => 'a bt"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    21
  bt_map :: "('a => 'b) => ('a bt => 'b bt)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    22
  preorder :: "'a bt => 'a list"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    23
  inorder :: "'a bt => 'a list"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    24
  postorder :: "'a bt => 'a list"
1167
cbd32a0f2f41 New theory and proofs including preorder, inorder, ..., initially
lcp
parents:
diff changeset
    25
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 1896
diff changeset
    26
primrec
1896
df4e40b9ff6d Simplified primrec - definitions.
berghofe
parents: 1476
diff changeset
    27
  "n_nodes (Lf) = 0"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    28
  "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
    29
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 1896
diff changeset
    30
primrec
1896
df4e40b9ff6d Simplified primrec - definitions.
berghofe
parents: 1476
diff changeset
    31
  "n_leaves (Lf) = Suc 0"
df4e40b9ff6d Simplified primrec - definitions.
berghofe
parents: 1476
diff changeset
    32
  "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
    33
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 1896
diff changeset
    34
primrec
1896
df4e40b9ff6d Simplified primrec - definitions.
berghofe
parents: 1476
diff changeset
    35
  "reflect (Lf) = Lf"
df4e40b9ff6d Simplified primrec - definitions.
berghofe
parents: 1476
diff changeset
    36
  "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
    37
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 1896
diff changeset
    38
primrec
1896
df4e40b9ff6d Simplified primrec - definitions.
berghofe
parents: 1476
diff changeset
    39
  "bt_map f Lf = Lf"
df4e40b9ff6d Simplified primrec - definitions.
berghofe
parents: 1476
diff changeset
    40
  "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
    41
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 1896
diff changeset
    42
primrec
1896
df4e40b9ff6d Simplified primrec - definitions.
berghofe
parents: 1476
diff changeset
    43
  "preorder (Lf) = []"
df4e40b9ff6d Simplified primrec - definitions.
berghofe
parents: 1476
diff changeset
    44
  "preorder (Br a t1 t2) = [a] @ (preorder t1) @ (preorder t2)"
1167
cbd32a0f2f41 New theory and proofs including preorder, inorder, ..., initially
lcp
parents:
diff changeset
    45
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 1896
diff changeset
    46
primrec
1896
df4e40b9ff6d Simplified primrec - definitions.
berghofe
parents: 1476
diff changeset
    47
  "inorder (Lf) = []"
df4e40b9ff6d Simplified primrec - definitions.
berghofe
parents: 1476
diff changeset
    48
  "inorder (Br a t1 t2) = (inorder t1) @ [a] @ (inorder t2)"
1167
cbd32a0f2f41 New theory and proofs including preorder, inorder, ..., initially
lcp
parents:
diff changeset
    49
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 1896
diff changeset
    50
primrec
1896
df4e40b9ff6d Simplified primrec - definitions.
berghofe
parents: 1476
diff changeset
    51
  "postorder (Lf) = []"
df4e40b9ff6d Simplified primrec - definitions.
berghofe
parents: 1476
diff changeset
    52
  "postorder (Br a t1 t2) = (postorder t1) @ (postorder t2) @ [a]"
1167
cbd32a0f2f41 New theory and proofs including preorder, inorder, ..., initially
lcp
parents:
diff changeset
    53
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    54
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    55
text {* \medskip BT simplification *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    56
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    57
lemma n_leaves_reflect: "n_leaves (reflect t) = n_leaves t"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    58
  apply (induct t)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    59
   apply auto
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    60
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    61
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    62
lemma n_nodes_reflect: "n_nodes (reflect t) = n_nodes t"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    63
  apply (induct t)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    64
   apply auto
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    65
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    66
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    67
text {*
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    68
  The famous relationship between the numbers of leaves and nodes.
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    69
*}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    70
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    71
lemma n_leaves_nodes: "n_leaves t = Suc (n_nodes t)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    72
  apply (induct t)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    73
   apply auto
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    74
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    75
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    76
lemma reflect_reflect_ident: "reflect (reflect t) = t"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    77
  apply (induct t)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    78
   apply auto
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    79
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    80
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    81
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
    82
  apply (induct t)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    83
   apply simp_all
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
    84
  done
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 inorder_bt_map: "inorder (bt_map f t) = map f (inorder 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 simp_all
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 preorder_reflect: "preorder (reflect t) = rev (postorder 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 simp_all
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 inorder_reflect: "inorder (reflect t) = rev (inorder 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
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   101
lemma postorder_reflect: "postorder (reflect t) = rev (preorder t)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   102
  apply (induct t)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   103
   apply simp_all
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   104
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8339
diff changeset
   105
1167
cbd32a0f2f41 New theory and proofs including preorder, inorder, ..., initially
lcp
parents:
diff changeset
   106
end