src/HOLCF/Holcfb.thy
author paulson
Fri, 08 Dec 1995 10:39:52 +0100
changeset 1394 a1d2735f5ade
parent 1274 ea0668a1c0ba
child 1479 21eb5e156d91
permissions -rw-r--r--
New function read_cterms is a combination of read_def_cterm and read_cterm. It reads and simultaneously type-checks a list of strings. cterm_of no longer catches exception TYPE; instead it must be caught later on and a message printed using Sign.exn_type_msg. More informative messages can be printed this way.

(*  Title: 	HOLCF/holcfb.thy
    ID:         $Id$
    Author: 	Franz Regensburger
    Copyright   1993  Technische Universitaet Muenchen

Basic definitions for the embedding of LCF into HOL.

*)

Holcfb = Nat + 

consts
	theleast     :: "(nat=>bool)=>nat"
defs

theleast_def    "theleast P == (@z.(P z & (!n.P n --> z<=n)))"

(* start 8bit 1 *)
(* end 8bit 1 *)

end