src/ZF/ex/Term.thy
author lcp
Fri, 12 Aug 1994 12:28:46 +0200
changeset 515 abcc438e7c27
child 740 281881c08397
permissions -rw-r--r--
installation of new inductive/datatype sections
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/Term.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
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
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    12
  term_rec  :: "[i, [i,i,i]=>i] => i"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    13
  term_map  :: "[i=>i, i] => i"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    14
  term_size :: "i=>i"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    15
  reflect   :: "i=>i"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    16
  preorder  :: "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
  term      :: "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
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    21
  "term(A)" = Apply ("a: A", "l: list(term(A))")
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    22
  monos	      "[list_mono]"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    23
  type_intrs  "[list_univ RS subsetD]"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    24
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    25
rules
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    26
  term_rec_def
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    27
   "term_rec(t,d) == \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    28
\   Vrec(t, %t g. term_case(%x zs. d(x, zs, map(%z.g`z, zs)), t))"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    29
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    30
  term_map_def	"term_map(f,t) == term_rec(t, %x zs rs. Apply(f(x), rs))"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    31
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    32
  term_size_def	"term_size(t) == term_rec(t, %x zs rs. succ(list_add(rs)))"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    33
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    34
  reflect_def	"reflect(t) == term_rec(t, %x zs rs. Apply(x, rev(rs)))"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    35
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    36
  preorder_def	"preorder(t) == term_rec(t, %x zs rs. Cons(x, flat(rs)))"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    37
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    38
end