src/ZF/ex/TF.thy
author lcp
Tue, 29 Nov 1994 00:31:31 +0100
changeset 753 ec86863e87c8
parent 581 465075fd257b
child 1155 928a16e02f9f
permissions -rw-r--r--
replaced "rules" by "defs"
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     1
(*  Title: 	ZF/ex/TF.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     2
    ID:         $Id$
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     3
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     5
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     6
Trees & forests, a mutually recursive type definition.
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     7
*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     8
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     9
TF = List +
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    10
consts
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    11
  TF_rec      :: "[i, [i,i,i]=>i, i, [i,i,i,i]=>i] => i"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    12
  TF_map      :: "[i=>i, i] => i"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    13
  TF_size     :: "i=>i"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    14
  TF_preorder :: "i=>i"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    15
  list_of_TF  :: "i=>i"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    16
  TF_of_list  :: "i=>i"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    17
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    18
  tree, forest, tree_forest    :: "i=>i"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    19
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    20
datatype
581
465075fd257b removal of needless quotes
lcp
parents: 515
diff changeset
    21
  "tree(A)"   = Tcons ("a: A",  "f: forest(A)")
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    22
and
581
465075fd257b removal of needless quotes
lcp
parents: 515
diff changeset
    23
  "forest(A)" = Fnil  |  Fcons ("t: tree(A)",  "f: forest(A)")
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    24
753
ec86863e87c8 replaced "rules" by "defs"
lcp
parents: 581
diff changeset
    25
defs
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    26
  TF_rec_def
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    27
    "TF_rec(z,b,c,d) == Vrec(z,  			\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    28
\      %z r. tree_forest_case(%x f. b(x, f, r`f), 	\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    29
\                             c, 			\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    30
\		              %t f. d(t, f, r`t, r`f), z))"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    31
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    32
  list_of_TF_def
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    33
    "list_of_TF(z) == TF_rec(z, %x f r. [Tcons(x,f)], [], \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    34
\		             %t f r1 r2. Cons(t, r2))"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    35
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    36
  TF_of_list_def
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    37
    "TF_of_list(f) == list_rec(f, Fnil,  %t f r. Fcons(t,r))"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    38
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    39
  TF_map_def
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    40
    "TF_map(h,z) == TF_rec(z, %x f r.Tcons(h(x),r), Fnil, \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    41
\                           %t f r1 r2. Fcons(r1,r2))"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    42
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    43
  TF_size_def
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    44
    "TF_size(z) == TF_rec(z, %x f r.succ(r), 0, %t f r1 r2. r1#+r2)"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    45
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    46
  TF_preorder_def
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    47
    "TF_preorder(z) == TF_rec(z, %x f r.Cons(x,r), Nil, %t f r1 r2. r1@r2)"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    48
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    49
end