src/ZF/ex/Term.thy
author paulson
Mon, 21 May 2001 14:36:24 +0200
changeset 11316 b4e71bd751e4
parent 6159 833b76d0e6dc
child 11354 9b80fe19407f
permissions -rw-r--r--
X-symbols for set theory
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      ZF/ex/Term.thy
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     2
    ID:         $Id$
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
515
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
Terms over an alphabet.
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     7
Illustrates the list functor (essentially the same type as in Trees & Forests)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     8
*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     9
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    10
Term = List +
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    11
consts
1401
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1155
diff changeset
    12
  term      :: i=>i
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    13
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    14
datatype
11316
b4e71bd751e4 X-symbols for set theory
paulson
parents: 6159
diff changeset
    15
  "term(A)" = Apply ("a \\<in> A", "l \\<in> list(term(A))")
6117
f9aad8ccd590 tidying of datatype and inductive definitions
paulson
parents: 6112
diff changeset
    16
  monos        list_mono
740
281881c08397 added comments on alternative uses of type_intrs/elims
lcp
parents: 515
diff changeset
    17
281881c08397 added comments on alternative uses of type_intrs/elims
lcp
parents: 515
diff changeset
    18
  type_elims  "[make_elim (list_univ RS subsetD)]"
281881c08397 added comments on alternative uses of type_intrs/elims
lcp
parents: 515
diff changeset
    19
(*Or could have
281881c08397 added comments on alternative uses of type_intrs/elims
lcp
parents: 515
diff changeset
    20
    type_intrs  "[list_univ RS subsetD]"
281881c08397 added comments on alternative uses of type_intrs/elims
lcp
parents: 515
diff changeset
    21
*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    22
6159
833b76d0e6dc constdefs
paulson
parents: 6117
diff changeset
    23
constdefs
6112
5e4871c5136b datatype package improvements
paulson
parents: 3840
diff changeset
    24
  term_rec  :: [i, [i,i,i]=>i] => i
1155
928a16e02f9f removed \...\ inside strings
clasohm
parents: 753
diff changeset
    25
   "term_rec(t,d) == 
6159
833b76d0e6dc constdefs
paulson
parents: 6117
diff changeset
    26
    Vrec(t, %t g. term_case(%x zs. d(x, zs, map(%z. g`z, zs)), t))"
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    27
6159
833b76d0e6dc constdefs
paulson
parents: 6117
diff changeset
    28
  term_map  :: [i=>i, i] => i
833b76d0e6dc constdefs
paulson
parents: 6117
diff changeset
    29
    "term_map(f,t) == term_rec(t, %x zs rs. Apply(f(x), rs))"
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    30
6159
833b76d0e6dc constdefs
paulson
parents: 6117
diff changeset
    31
  term_size :: i=>i
833b76d0e6dc constdefs
paulson
parents: 6117
diff changeset
    32
   "term_size(t) == term_rec(t, %x zs rs. succ(list_add(rs)))"
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    33
6159
833b76d0e6dc constdefs
paulson
parents: 6117
diff changeset
    34
  reflect   :: i=>i
833b76d0e6dc constdefs
paulson
parents: 6117
diff changeset
    35
    "reflect(t) == term_rec(t, %x zs rs. Apply(x, rev(rs)))"
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    36
6159
833b76d0e6dc constdefs
paulson
parents: 6117
diff changeset
    37
  preorder  :: i=>i
833b76d0e6dc constdefs
paulson
parents: 6117
diff changeset
    38
    "preorder(t) == term_rec(t, %x zs rs. Cons(x, flat(rs)))"
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    39
6159
833b76d0e6dc constdefs
paulson
parents: 6117
diff changeset
    40
  postorder :: i=>i
833b76d0e6dc constdefs
paulson
parents: 6117
diff changeset
    41
    "postorder(t) == term_rec(t, %x zs rs. flat(rs) @ [x])"
6112
5e4871c5136b datatype package improvements
paulson
parents: 3840
diff changeset
    42
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    43
end