| author | wenzelm | 
| Wed, 27 Oct 1999 18:12:40 +0200 | |
| changeset 7956 | edaca60a54cd | 
| parent 5184 | 9b8547a9496a | 
| child 8339 | bcadeb9c7095 | 
| permissions | -rw-r--r-- | 
| 1476 | 1 | (* Title: HOL/ex/BT.thy | 
| 1167 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 2 | ID: $Id$ | 
| 1476 | 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 | |
| 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 6 | Binary trees (based on the ZF version) | 
| 
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 | |
| 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 9 | BT = List + | 
| 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 10 | |
| 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 11 | datatype 'a bt = Lf  |  Br 'a ('a bt) ('a bt)
 | 
| 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 12 | |
| 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 13 | consts | 
| 1476 | 14 | n_nodes :: 'a bt => nat | 
| 15 | n_leaves :: 'a bt => nat | |
| 16 | reflect :: 'a bt => 'a bt | |
| 1376 | 17 |     bt_map      :: ('a=>'b) => ('a bt => 'b bt)
 | 
| 18 | preorder :: 'a bt => 'a list | |
| 19 | inorder :: 'a bt => 'a list | |
| 20 | postorder :: 'a bt => 'a list | |
| 1167 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 21 | |
| 5184 | 22 | primrec | 
| 1896 | 23 | "n_nodes (Lf) = 0" | 
| 24 | "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 | 25 | |
| 5184 | 26 | primrec | 
| 1896 | 27 | "n_leaves (Lf) = Suc 0" | 
| 28 | "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 | 29 | |
| 5184 | 30 | primrec | 
| 1896 | 31 | "reflect (Lf) = Lf" | 
| 32 | "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 | 33 | |
| 5184 | 34 | primrec | 
| 1896 | 35 | "bt_map f Lf = Lf" | 
| 36 | "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 | 37 | |
| 5184 | 38 | primrec | 
| 1896 | 39 | "preorder (Lf) = []" | 
| 40 | "preorder (Br a t1 t2) = [a] @ (preorder t1) @ (preorder t2)" | |
| 1167 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 41 | |
| 5184 | 42 | primrec | 
| 1896 | 43 | "inorder (Lf) = []" | 
| 44 | "inorder (Br a t1 t2) = (inorder t1) @ [a] @ (inorder t2)" | |
| 1167 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 45 | |
| 5184 | 46 | primrec | 
| 1896 | 47 | "postorder (Lf) = []" | 
| 48 | "postorder (Br a t1 t2) = (postorder t1) @ (postorder t2) @ [a]" | |
| 1167 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 49 | |
| 
cbd32a0f2f41
New theory and proofs including preorder, inorder, ..., initially
 lcp parents: diff
changeset | 50 | end | 
| 1896 | 51 |