src/HOL/ex/Term.thy
author lcp
Tue, 25 Apr 1995 11:14:03 +0200
changeset 1072 0140ff702b23
parent 969 b051e2fc2e34
child 1151 c820b3cc3df0
permissions -rw-r--r--
updated version
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     1
(*  Title: 	HOL/ex/Term
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     3
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     5
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     6
Terms over a given alphabet -- function applications; illustrates list functor
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     7
  (essentially the same type as in Trees & Forests)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     8
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     9
There is no constructor APP because it is simply cons ($) 
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    10
*)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    11
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    12
Term = SList +
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    13
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    14
types   'a term
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    15
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    16
arities term :: (term)term
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    17
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    18
consts
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    19
  term		:: "'a item set => 'a item set"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    20
  Rep_term	:: "'a term => 'a item"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    21
  Abs_term	:: "'a item => 'a term"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    22
  Rep_Tlist	:: "'a term list => 'a item"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    23
  Abs_Tlist	:: "'a item => 'a term list"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    24
  App		:: "['a, ('a term)list] => 'a term"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    25
  Term_rec	:: "['a item, ['a item , 'a item, 'b list]=>'b] => 'b"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    26
  term_rec	:: "['a term, ['a ,'a term list, 'b list]=>'b] => 'b"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    27
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    28
inductive "term(A)"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    29
  intrs
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    30
    APP_I "[| M: A;  N : list(term(A)) |] ==> M$N : term(A)"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    31
  monos   "[list_mono]"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    32
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    33
defs
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    34
  (*defining abstraction/representation functions for term list...*)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    35
  Rep_Tlist_def	"Rep_Tlist == Rep_map(Rep_term)"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    36
  Abs_Tlist_def	"Abs_Tlist == Abs_map(Abs_term)"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    37
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    38
  (*defining the abstract constants*)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    39
  App_def 	"App a ts == Abs_term(Leaf(a) $ Rep_Tlist(ts))"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    40
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    41
  (*list recursion*)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    42
  Term_rec_def	
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    43
   "Term_rec M d == wfrec (trancl pred_sexp) M \
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    44
\           (Split(%x y g. d x y (Abs_map g y)))"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    45
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    46
  term_rec_def
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    47
   "term_rec t d == \
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    48
\   Term_rec (Rep_term t) (%x y r. d (Inv Leaf x) (Abs_Tlist(y)) r)"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    49
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    50
rules
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    51
    (*faking a type definition for term...*)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    52
  Rep_term 		"Rep_term(n): term(range(Leaf))"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    53
  Rep_term_inverse 	"Abs_term(Rep_term(t)) = t"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    54
  Abs_term_inverse 	"M: term(range(Leaf)) ==> Rep_term(Abs_term(M)) = M"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    55
end