Journal of Formalized Mathematics
Volume 11, 1999
University of Bialystok
Copyright (c) 1999 Association of Mizar Users

The abstract of the Mizar article:

Predicate Calculus for Boolean Valued Functions. Part XI

by
Shunichi Kobayashi

Received November 15, 1999

MML identifier: BVFUNC19
[ Mizar article, MML identifier index ]


environ

 vocabulary FUNCT_2, MARGREL1, PARTIT1, EQREL_1, BVFUNC_2, ZF_LANG;
 notation XBOOLE_0, SUBSET_1, FRAENKEL, MARGREL1, VALUAT_1, EQREL_1, BVFUNC_1,
      BVFUNC_2;
 constructors BVFUNC_2, BVFUNC_1;
 clusters MARGREL1, XBOOLE_0;


begin

reserve Y for non empty set,
        a for Element of Funcs(Y,BOOLEAN),
        G for Subset of PARTITIONS(Y),
        A,B for a_partition of Y;

canceled 5;

theorem :: BVFUNC19:6
G is independent implies
All('not' Ex(a,A,G),B,G) '<' All(All('not' a,B,G),A,G);

Back to top