src/ZF/Bool.ML
author nipkow
Mon, 06 Aug 2001 13:43:24 +0200
changeset 11464 ddea204de5bc
parent 9907 473a6604da94
permissions -rw-r--r--
turned translation for 1::nat into def. introduced 1' and replaced most occurrences of 1 by 1'.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 760
diff changeset
     1
(*  Title:      ZF/bool
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 760
diff changeset
     3
    Author:     Martin D Coen, 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
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5137
diff changeset
     6
Booleans in Zermelo-Fraenkel Set Theory 
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 6153
diff changeset
     9
bind_thms ("bool_defs", [bool_def,cond_def]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4091
diff changeset
    11
Goalw [succ_def] "{0} = 1";
2493
bdeb5024353a Removal of sum_cs and eq_cs
paulson
parents: 2469
diff changeset
    12
by (rtac refl 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    13
qed "singleton_0";
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    14
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
(* Introduction rules *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    16
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4091
diff changeset
    17
Goalw bool_defs "1 : bool";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    18
by (rtac (consI1 RS consI2) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 129
diff changeset
    19
qed "bool_1I";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    20
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4091
diff changeset
    21
Goalw bool_defs "0 : bool";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    22
by (rtac consI1 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 129
diff changeset
    23
qed "bool_0I";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    24
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5268
diff changeset
    25
Addsimps [bool_1I, bool_0I];
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6053
diff changeset
    26
AddTCs   [bool_1I, bool_0I];
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5268
diff changeset
    27
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4091
diff changeset
    28
Goalw bool_defs "1~=0";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    29
by (rtac succ_not_0 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 129
diff changeset
    30
qed "one_not_0";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    32
(** 1=0 ==> R **)
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 6153
diff changeset
    33
bind_thm ("one_neq_0", one_not_0 RS notE);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    34
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5137
diff changeset
    35
val major::prems = Goalw bool_defs
119
0e58da397b1d boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents: 37
diff changeset
    36
    "[| c: bool;  c=1 ==> P;  c=0 ==> P |] ==> P";
129
dc50a4b96d7b expandshort and other trivial changes
lcp
parents: 119
diff changeset
    37
by (rtac (major RS consE) 1);
119
0e58da397b1d boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents: 37
diff changeset
    38
by (REPEAT (eresolve_tac (singletonE::prems) 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 129
diff changeset
    39
qed "boolE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    40
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    41
(** cond **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    42
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    43
(*1 means true*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4091
diff changeset
    44
Goalw bool_defs "cond(1,c,d) = c";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    45
by (rtac (refl RS if_P) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 129
diff changeset
    46
qed "cond_1";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    47
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    48
(*0 means false*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4091
diff changeset
    49
Goalw bool_defs "cond(0,c,d) = d";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    50
by (rtac (succ_not_0 RS not_sym RS if_not_P) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 129
diff changeset
    51
qed "cond_0";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    52
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    53
Addsimps [cond_1, cond_0];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    54
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2493
diff changeset
    55
fun bool_tac i = fast_tac (claset() addSEs [boolE] addss (simpset())) i;
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    56
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    57
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5268
diff changeset
    58
Goal "[| b: bool;  c: A(1);  d: A(0) |] ==> cond(b,c,d): A(b)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    59
by (bool_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 129
diff changeset
    60
qed "cond_type";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6053
diff changeset
    61
AddTCs [cond_type];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    62
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5268
diff changeset
    63
(*For Simp_tac and Blast_tac*)
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5268
diff changeset
    64
Goal "[| b: bool;  c: A;  d: A |] ==> cond(b,c,d): A";
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5268
diff changeset
    65
by (bool_tac 1);
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5268
diff changeset
    66
qed "cond_simple_type";
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5268
diff changeset
    67
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5137
diff changeset
    68
val [rew] = Goal "[| !!b. j(b)==cond(b,c,d) |] ==> j(1) = c";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    69
by (rewtac rew);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    70
by (rtac cond_1 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 129
diff changeset
    71
qed "def_cond_1";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    72
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5137
diff changeset
    73
val [rew] = Goal "[| !!b. j(b)==cond(b,c,d) |] ==> j(0) = d";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    74
by (rewtac rew);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    75
by (rtac cond_0 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 129
diff changeset
    76
qed "def_cond_0";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    77
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    78
fun conds def = [standard (def RS def_cond_1), standard (def RS def_cond_0)];
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    79
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 6153
diff changeset
    80
val [not_1, not_0] = conds not_def;
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 6153
diff changeset
    81
val [and_1, and_0] = conds and_def;
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 6153
diff changeset
    82
val [or_1, or_0]   = conds or_def;
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 6153
diff changeset
    83
val [xor_1, xor_0] = conds xor_def;
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 6153
diff changeset
    84
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 6153
diff changeset
    85
bind_thm ("not_1", not_1);
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 6153
diff changeset
    86
bind_thm ("not_0", not_0);
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 6153
diff changeset
    87
bind_thm ("and_1", and_1);
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 6153
diff changeset
    88
bind_thm ("and_0", and_0);
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 6153
diff changeset
    89
bind_thm ("or_1", or_1);
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 6153
diff changeset
    90
bind_thm ("or_0", or_0);
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 6153
diff changeset
    91
bind_thm ("xor_1", xor_1);
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 6153
diff changeset
    92
bind_thm ("xor_0", xor_0);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    93
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    94
Addsimps [not_1,not_0,and_1,and_0,or_1,or_0,xor_1,xor_0];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    95
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6053
diff changeset
    96
Goalw [not_def] "a:bool ==> not(a) : bool";
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6053
diff changeset
    97
by (Asm_simp_tac 1);
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6053
diff changeset
    98
qed "not_type";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    99
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6053
diff changeset
   100
Goalw [and_def] "[| a:bool;  b:bool |] ==> a and b : bool";
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6053
diff changeset
   101
by (Asm_simp_tac 1);
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6053
diff changeset
   102
qed "and_type";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   103
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6053
diff changeset
   104
Goalw [or_def] "[| a:bool;  b:bool |] ==> a or b : bool";
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6053
diff changeset
   105
by (Asm_simp_tac 1);
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6053
diff changeset
   106
qed "or_type";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   107
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6053
diff changeset
   108
AddTCs [not_type, and_type, or_type];
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5268
diff changeset
   109
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6053
diff changeset
   110
Goalw [xor_def] "[| a:bool;  b:bool |] ==> a xor b : bool";
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6053
diff changeset
   111
by (Asm_simp_tac 1);
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6053
diff changeset
   112
qed "xor_type";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   113
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6053
diff changeset
   114
AddTCs [xor_type];
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5268
diff changeset
   115
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 6153
diff changeset
   116
bind_thms ("bool_typechecks",
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 6153
diff changeset
   117
  [bool_1I, bool_0I, cond_type, not_type, and_type, or_type, xor_type]);
119
0e58da397b1d boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents: 37
diff changeset
   118
0e58da397b1d boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents: 37
diff changeset
   119
(*** Laws for 'not' ***)
0e58da397b1d boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents: 37
diff changeset
   120
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   121
Goal "a:bool ==> not(not(a)) = a";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   122
by (bool_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 129
diff changeset
   123
qed "not_not";
119
0e58da397b1d boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents: 37
diff changeset
   124
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   125
Goal "a:bool ==> not(a and b) = not(a) or not(b)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   126
by (bool_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 129
diff changeset
   127
qed "not_and";
119
0e58da397b1d boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents: 37
diff changeset
   128
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   129
Goal "a:bool ==> not(a or b) = not(a) and not(b)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   130
by (bool_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 129
diff changeset
   131
qed "not_or";
119
0e58da397b1d boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents: 37
diff changeset
   132
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   133
Addsimps [not_not, not_and, not_or];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   134
119
0e58da397b1d boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents: 37
diff changeset
   135
(*** Laws about 'and' ***)
0e58da397b1d boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents: 37
diff changeset
   136
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   137
Goal "a: bool ==> a and a = a";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   138
by (bool_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 129
diff changeset
   139
qed "and_absorb";
119
0e58da397b1d boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents: 37
diff changeset
   140
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   141
Addsimps [and_absorb];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   142
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   143
Goal "[| a: bool; b:bool |] ==> a and b = b and a";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   144
by (bool_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 129
diff changeset
   145
qed "and_commute";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   146
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   147
Goal "a: bool ==> (a and b) and c  =  a and (b and c)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   148
by (bool_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 129
diff changeset
   149
qed "and_assoc";
119
0e58da397b1d boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents: 37
diff changeset
   150
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5137
diff changeset
   151
Goal "[| a: bool; b:bool; c:bool |] ==> \
119
0e58da397b1d boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents: 37
diff changeset
   152
\      (a or b) and c  =  (a and c) or (b and c)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   153
by (bool_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 129
diff changeset
   154
qed "and_or_distrib";
119
0e58da397b1d boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents: 37
diff changeset
   155
0e58da397b1d boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents: 37
diff changeset
   156
(** binary orion **)
0e58da397b1d boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents: 37
diff changeset
   157
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   158
Goal "a: bool ==> a or a = a";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   159
by (bool_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 129
diff changeset
   160
qed "or_absorb";
119
0e58da397b1d boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents: 37
diff changeset
   161
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   162
Addsimps [or_absorb];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   163
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   164
Goal "[| a: bool; b:bool |] ==> a or b = b or a";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   165
by (bool_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 129
diff changeset
   166
qed "or_commute";
119
0e58da397b1d boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents: 37
diff changeset
   167
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   168
Goal "a: bool ==> (a or b) or c  =  a or (b or c)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   169
by (bool_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 129
diff changeset
   170
qed "or_assoc";
119
0e58da397b1d boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents: 37
diff changeset
   171
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5137
diff changeset
   172
Goal "[| a: bool; b: bool; c: bool |] ==> \
119
0e58da397b1d boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents: 37
diff changeset
   173
\          (a and b) or c  =  (a or c) and (b or c)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   174
by (bool_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 129
diff changeset
   175
qed "or_and_distrib";
119
0e58da397b1d boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents: 37
diff changeset
   176