src/ZF/Coind/Types.thy
author paulson
Mon, 28 Dec 1998 16:54:01 +0100
changeset 6046 2c8a8be36c94
parent 3840 e0baea4d485a
child 6068 2d8f3e1f1151
permissions -rw-r--r--
converted to use new primrec section
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      ZF/Coind/Types.thy
916
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
     2
    ID:         $Id$
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     3
    Author:     Jacob Frost, Cambridge University Computer Laboratory
916
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
     4
    Copyright   1995  University of Cambridge
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
     5
*)
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
     6
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
     7
Types = Language +
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
     8
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
     9
consts
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    10
  Ty :: i                       (* Datatype of types *)
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    11
  TyConst :: i          (* Abstract type of type constants *)
934
2e0203309287 Replaced rules by defs. Also got rid of tyconstU by
lcp
parents: 916
diff changeset
    12
datatype <= "univ(TyConst)"
2e0203309287 Replaced rules by defs. Also got rid of tyconstU by
lcp
parents: 916
diff changeset
    13
  "Ty" = t_const("tc:TyConst")
2e0203309287 Replaced rules by defs. Also got rid of tyconstU by
lcp
parents: 916
diff changeset
    14
       | t_fun("t1:Ty","t2:Ty")
916
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    15
  
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    16
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    17
(* Definition of type environments and associated operators *)
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    18
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    19
consts
1401
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1155
diff changeset
    20
  TyEnv :: i
934
2e0203309287 Replaced rules by defs. Also got rid of tyconstU by
lcp
parents: 916
diff changeset
    21
datatype <= "univ(Ty Un ExVar)"
2e0203309287 Replaced rules by defs. Also got rid of tyconstU by
lcp
parents: 916
diff changeset
    22
  "TyEnv" = te_emp
2e0203309287 Replaced rules by defs. Also got rid of tyconstU by
lcp
parents: 916
diff changeset
    23
          | te_owr("te:TyEnv","x:ExVar","t:Ty") 
916
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    24
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    25
consts
1401
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1155
diff changeset
    26
  te_dom :: i => i
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1155
diff changeset
    27
  te_app :: [i,i] => i
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    28
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    29
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    30
primrec (*domain of the type environment*)
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    31
  "te_dom (te_emp) = 0"
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    32
  "te_dom(te_owr(te,x,v)) = te_dom(te) Un {x}"
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    33
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    34
primrec (*lookup up identifiers in the type environment*)
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    35
  "te_app (te_emp,x) = 0"
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    36
  "te_app (te_owr(te,y,t),x) = if(x=y, t, te_app(te,x))"
916
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    37
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    38
end
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    39
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    40
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    41
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    42
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    43