src/ZF/Bool.thy
author paulson
Tue, 02 Jul 2002 13:28:08 +0200
changeset 13269 3ba9be497c33
parent 13239 f599984ec4c2
child 13328 703de709a64b
permissions -rw-r--r--
Tidying and introduction of various new theorems
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      ZF/bool.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
Booleans in Zermelo-Fraenkel Set Theory 
837
778f01546669 Re-indented declarations; declared the number 2
lcp
parents: 799
diff changeset
     7
778f01546669 Re-indented declarations; declared the number 2
lcp
parents: 799
diff changeset
     8
2 is equal to bool, but serves a different purpose
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
13239
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    11
theory Bool = pair:
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    12
2539
ddd22ceee8cc turned some consts into syntax;
wenzelm
parents: 2469
diff changeset
    13
syntax
ddd22ceee8cc turned some consts into syntax;
wenzelm
parents: 2469
diff changeset
    14
    "1"         :: i             ("1")
ddd22ceee8cc turned some consts into syntax;
wenzelm
parents: 2469
diff changeset
    15
    "2"         :: i             ("2")
ddd22ceee8cc turned some consts into syntax;
wenzelm
parents: 2469
diff changeset
    16
14
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 0
diff changeset
    17
translations
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 0
diff changeset
    18
   "1"  == "succ(0)"
837
778f01546669 Re-indented declarations; declared the number 2
lcp
parents: 799
diff changeset
    19
   "2"  == "succ(1)"
14
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 0
diff changeset
    20
13239
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    21
constdefs
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    22
  bool        :: i
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    23
    "bool == {0,1}"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    24
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    25
  cond        :: "[i,i,i]=>i"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    26
    "cond(b,c,d) == if(b=1,c,d)"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    27
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    28
  not         :: "i=>i"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    29
    "not(b) == cond(b,0,1)"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    30
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    31
  "and"       :: "[i,i]=>i"      (infixl "and" 70)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    32
    "a and b == cond(a,b,0)"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    33
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    34
  or          :: "[i,i]=>i"      (infixl "or" 65)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    35
    "a or b == cond(a,1,b)"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    36
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    37
  xor         :: "[i,i]=>i"      (infixl "xor" 65)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    38
    "a xor b == cond(a,not(b),b)"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    39
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    40
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    41
lemmas bool_defs = bool_def cond_def
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    42
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    43
lemma singleton_0: "{0} = 1"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    44
by (simp add: succ_def)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    45
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    46
(* Introduction rules *)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    47
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    48
lemma bool_1I [simp,TC]: "1 : bool"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    49
by (simp add: bool_defs )
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    50
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    51
lemma bool_0I [simp,TC]: "0 : bool"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    52
by (simp add: bool_defs)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    53
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    54
lemma one_not_0: "1~=0"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    55
by (simp add: bool_defs )
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    56
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    57
(** 1=0 ==> R **)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    58
lemmas one_neq_0 = one_not_0 [THEN notE, standard]
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    59
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    60
lemma boolE:
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    61
    "[| c: bool;  c=1 ==> P;  c=0 ==> P |] ==> P"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    62
by (simp add: bool_defs, blast)  
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    63
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    64
(** cond **)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    65
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    66
(*1 means true*)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    67
lemma cond_1 [simp]: "cond(1,c,d) = c"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    68
by (simp add: bool_defs )
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    69
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    70
(*0 means false*)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    71
lemma cond_0 [simp]: "cond(0,c,d) = d"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    72
by (simp add: bool_defs )
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    73
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    74
lemma cond_type [TC]: "[| b: bool;  c: A(1);  d: A(0) |] ==> cond(b,c,d): A(b)"
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13239
diff changeset
    75
by (simp add: bool_defs, blast)
13239
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    76
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    77
(*For Simp_tac and Blast_tac*)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    78
lemma cond_simple_type: "[| b: bool;  c: A;  d: A |] ==> cond(b,c,d): A"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    79
by (simp add: bool_defs )
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    80
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    81
lemma def_cond_1: "[| !!b. j(b)==cond(b,c,d) |] ==> j(1) = c"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    82
by simp
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    83
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    84
lemma def_cond_0: "[| !!b. j(b)==cond(b,c,d) |] ==> j(0) = d"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    85
by simp
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    86
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    87
lemmas not_1 = not_def [THEN def_cond_1, standard, simp]
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    88
lemmas not_0 = not_def [THEN def_cond_0, standard, simp]
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    89
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    90
lemmas and_1 = and_def [THEN def_cond_1, standard, simp]
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    91
lemmas and_0 = and_def [THEN def_cond_0, standard, simp]
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    92
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    93
lemmas or_1 = or_def [THEN def_cond_1, standard, simp]
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    94
lemmas or_0 = or_def [THEN def_cond_0, standard, simp]
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    95
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    96
lemmas xor_1 = xor_def [THEN def_cond_1, standard, simp]
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    97
lemmas xor_0 = xor_def [THEN def_cond_0, standard, simp]
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    98
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
    99
lemma not_type [TC]: "a:bool ==> not(a) : bool"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   100
by (simp add: not_def)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   101
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   102
lemma and_type [TC]: "[| a:bool;  b:bool |] ==> a and b : bool"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   103
by (simp add: and_def)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   104
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   105
lemma or_type [TC]: "[| a:bool;  b:bool |] ==> a or b : bool"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   106
by (simp add: or_def)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   107
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   108
lemma xor_type [TC]: "[| a:bool;  b:bool |] ==> a xor b : bool"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   109
by (simp add: xor_def)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   110
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   111
lemmas bool_typechecks = bool_1I bool_0I cond_type not_type and_type
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   112
                         or_type xor_type
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   113
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   114
(*** Laws for 'not' ***)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   115
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   116
lemma not_not [simp]: "a:bool ==> not(not(a)) = a"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   117
by (elim boolE, auto)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   118
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   119
lemma not_and [simp]: "a:bool ==> not(a and b) = not(a) or not(b)"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   120
by (elim boolE, auto)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   121
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   122
lemma not_or [simp]: "a:bool ==> not(a or b) = not(a) and not(b)"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   123
by (elim boolE, auto)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   124
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   125
(*** Laws about 'and' ***)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   126
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   127
lemma and_absorb [simp]: "a: bool ==> a and a = a"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   128
by (elim boolE, auto)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   129
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   130
lemma and_commute: "[| a: bool; b:bool |] ==> a and b = b and a"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   131
by (elim boolE, auto)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   132
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   133
lemma and_assoc: "a: bool ==> (a and b) and c  =  a and (b and c)"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   134
by (elim boolE, auto)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   135
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   136
lemma and_or_distrib: "[| a: bool; b:bool; c:bool |] ==>  
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   137
       (a or b) and c  =  (a and c) or (b and c)"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   138
by (elim boolE, auto)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   139
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13239
diff changeset
   140
(** binary 'or' **)
13239
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   141
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   142
lemma or_absorb [simp]: "a: bool ==> a or a = a"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   143
by (elim boolE, auto)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   144
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   145
lemma or_commute: "[| a: bool; b:bool |] ==> a or b = b or a"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   146
by (elim boolE, auto)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   147
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   148
lemma or_assoc: "a: bool ==> (a or b) or c  =  a or (b or c)"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   149
by (elim boolE, auto)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   150
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   151
lemma or_and_distrib: "[| a: bool; b: bool; c: bool |] ==>  
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   152
           (a and b) or c  =  (a or c) and (b or c)"
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   153
by (elim boolE, auto)
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   154
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13239
diff changeset
   155
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13239
diff changeset
   156
constdefs bool_of_o :: "o=>i"
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13239
diff changeset
   157
   "bool_of_o(P) == (if P then 1 else 0)"
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13239
diff changeset
   158
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13239
diff changeset
   159
lemma [simp]: "bool_of_o(True) = 1"
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13239
diff changeset
   160
by (simp add: bool_of_o_def) 
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13239
diff changeset
   161
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13239
diff changeset
   162
lemma [simp]: "bool_of_o(False) = 0"
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13239
diff changeset
   163
by (simp add: bool_of_o_def) 
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13239
diff changeset
   164
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13239
diff changeset
   165
lemma [simp,TC]: "bool_of_o(P) \<in> bool"
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13239
diff changeset
   166
by (simp add: bool_of_o_def) 
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13239
diff changeset
   167
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13239
diff changeset
   168
lemma [simp]: "(bool_of_o(P) = 1) <-> P"
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13239
diff changeset
   169
by (simp add: bool_of_o_def) 
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13239
diff changeset
   170
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13239
diff changeset
   171
lemma [simp]: "(bool_of_o(P) = 0) <-> ~P"
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13239
diff changeset
   172
by (simp add: bool_of_o_def) 
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13239
diff changeset
   173
13239
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   174
ML
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   175
{*
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   176
val bool_def = thm "bool_def";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   177
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   178
val bool_defs = thms "bool_defs";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   179
val singleton_0 = thm "singleton_0";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   180
val bool_1I = thm "bool_1I";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   181
val bool_0I = thm "bool_0I";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   182
val one_not_0 = thm "one_not_0";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   183
val one_neq_0 = thm "one_neq_0";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   184
val boolE = thm "boolE";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   185
val cond_1 = thm "cond_1";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   186
val cond_0 = thm "cond_0";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   187
val cond_type = thm "cond_type";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   188
val cond_simple_type = thm "cond_simple_type";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   189
val def_cond_1 = thm "def_cond_1";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   190
val def_cond_0 = thm "def_cond_0";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   191
val not_1 = thm "not_1";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   192
val not_0 = thm "not_0";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   193
val and_1 = thm "and_1";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   194
val and_0 = thm "and_0";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   195
val or_1 = thm "or_1";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   196
val or_0 = thm "or_0";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   197
val xor_1 = thm "xor_1";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   198
val xor_0 = thm "xor_0";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   199
val not_type = thm "not_type";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   200
val and_type = thm "and_type";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   201
val or_type = thm "or_type";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   202
val xor_type = thm "xor_type";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   203
val bool_typechecks = thms "bool_typechecks";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   204
val not_not = thm "not_not";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   205
val not_and = thm "not_and";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   206
val not_or = thm "not_or";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   207
val and_absorb = thm "and_absorb";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   208
val and_commute = thm "and_commute";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   209
val and_assoc = thm "and_assoc";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   210
val and_or_distrib = thm "and_or_distrib";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   211
val or_absorb = thm "or_absorb";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   212
val or_commute = thm "or_commute";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   213
val or_assoc = thm "or_assoc";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   214
val or_and_distrib = thm "or_and_distrib";
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   215
*}
f599984ec4c2 converted Bool, Trancl, Rel to Isar format
paulson
parents: 13220
diff changeset
   216
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   217
end