src/CTT/Bool.thy
author wenzelm
Fri, 16 Sep 2005 23:01:29 +0200
changeset 17441 5b5feca0344a
parent 3837 d7f033c74b38
child 19761 5cd82054c2c6
permissions -rw-r--r--
converted to Isar theory format;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1474
3f7d67927fe2 expanded tabs
clasohm
parents: 0
diff changeset
     1
(*  Title:      CTT/bool
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1474
3f7d67927fe2 expanded tabs
clasohm
parents: 0
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1991  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
17441
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
     7
header {* The two-element type (booleans and conditionals) *}
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
     8
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
     9
theory Bool
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    10
imports CTT
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    11
uses "~~/src/Provers/typedsimp.ML" "rew.ML"
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    12
begin
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
17441
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    14
consts
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    15
  Bool        :: "t"
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    16
  true        :: "i"
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    17
  false       :: "i"
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    18
  cond        :: "[i,i,i]=>i"
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    19
defs
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    20
  Bool_def:   "Bool == T+T"
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    21
  true_def:   "true == inl(tt)"
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    22
  false_def:  "false == inr(tt)"
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    23
  cond_def:   "cond(a,b,c) == when(a, %u. b, %u. c)"
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    24
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    25
ML {* use_legacy_bindings (the_context ()) *}
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    26
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    27
end