| author | wenzelm | 
| Tue, 28 Oct 1997 17:36:16 +0100 | |
| changeset 4022 | 0770a19c48d3 | 
| parent 3840 | e0baea4d485a | 
| child 6046 | 2c8a8be36c94 | 
| permissions | -rw-r--r-- | 
| 1478 | 1 | (* Title: Redex.thy | 
| 1048 | 2 | ID: $Id$ | 
| 1478 | 3 | Author: Ole Rasmussen | 
| 1048 | 4 | Copyright 1995 University of Cambridge | 
| 5 | Logic Image: ZF | |
| 6 | *) | |
| 7 | ||
| 2874 
b1e7e2179597
Datatype declarations now require theory Datatype--NOT in quotes
 paulson parents: 
1478diff
changeset | 8 | Redex = Datatype + | 
| 1048 | 9 | consts | 
| 1401 | 10 | redexes :: i | 
| 1048 | 11 | |
| 12 | datatype | |
| 1478 | 13 |   "redexes" = Var ("n: nat")            
 | 
| 1048 | 14 |             | Fun ("t: redexes")
 | 
| 15 |             | App ("b:bool" ,"f:redexes" , "a:redexes")
 | |
| 16 | type_intrs "[bool_into_univ]" | |
| 17 | ||
| 18 | ||
| 19 | consts | |
| 1478 | 20 | redex_rec :: [i, i=>i, [i,i]=>i, [i,i,i,i,i]=>i]=>i | 
| 1048 | 21 | |
| 22 | defs | |
| 23 | redex_rec_def | |
| 1155 | 24 | "redex_rec(p,b,c,d) == | 
| 3840 | 25 | Vrec(p, %p g. redexes_case(b, %x. c(x,g`x), | 
| 26 | %n x y. d(n, x, y, g`x, g`y), p))" | |
| 1048 | 27 | end | 
| 28 |