author | wenzelm |
Fri, 19 Aug 1994 15:37:05 +0200 | |
changeset 554 | c7d9018cc9e6 |
parent 0 | a5a9c433f639 |
child 1474 | 3f7d67927fe2 |
permissions | -rw-r--r-- |
(* Title: CTT/bool ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1991 University of Cambridge The two-element type (booleans and conditionals) *) Bool = CTT + consts Bool :: "t" true,false :: "i" cond :: "[i,i,i]=>i" rules Bool_def "Bool == T+T" true_def "true == inl(tt)" false_def "false == inr(tt)" cond_def "cond(a,b,c) == when(a, %u.b, %u.c)" end