author | lcp |
Thu, 04 Aug 1994 11:45:59 +0200 | |
changeset 507 | a00301e9e64b |
parent 467 | 92868dab2939 |
child 753 | ec86863e87c8 |
permissions | -rw-r--r-- |
(* Title: ZF/OrderType.thy ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Order types. The order type of a well-ordering is the least ordinal isomorphic to it. *) OrderType = OrderArith + Ordinal + consts ordermap :: "[i,i]=>i" ordertype :: "[i,i]=>i" rules ordermap_def "ordermap(A,r) == lam x:A. wfrec[A](r, x, %x f. f `` pred(A,x,r))" ordertype_def "ordertype(A,r) == ordermap(A,r)``A" end