:: HEYTING1 semantic presentation

theorem Th1: :: HEYTING1:1
for A, B, C being non empty set
for f being Function of A,B st ( for x being Element of A holds f . x in C ) holds
f is Function of A,C
proof end;

definition
let A be non empty set ;
let B be Element of Fin A, C be Element of Fin A;
redefine pred c= as c2 c= c3 means :: HEYTING1:def 1
for a being Element of A st a in B holds
a in C;
compatibility
( B c= C iff for a being Element of A st a in B holds
a in C )
proof end;
end;

:: deftheorem Def1 defines c= HEYTING1:def 1 :
for A being non empty set
for B, C being Element of Fin A holds
( B c= C iff for a being Element of A st a in B holds
a in C );

definition
let A be set ;
assume E21: not A is empty ;
func [c1] -> non empty set equals :Def2: :: HEYTING1:def 2
A;
correctness
coherence
A is non empty set
;
by ;
end;

:: deftheorem Def2 defines [ HEYTING1:def 2 :
for A being set st not A is empty holds
[A] = A;

theorem Th2: :: HEYTING1:2
canceled;

theorem Th3: :: HEYTING1:3
for A being set
for B being Element of Fin (DISJOINT_PAIRS A) st B = {} holds
mi B = {}
proof end;

E40: now
let A be set ;
let a be Element of DISJOINT_PAIRS A;
reconsider B = {.a.} as Element of Fin (DISJOINT_PAIRS A) ;
now
let c be Element of DISJOINT_PAIRS A;
let b be Element of DISJOINT_PAIRS A;
assume that
E21: ( c in B & b in B ) and
c c= b ;
( c = a & b = a ) by , TARSKI:def 1;
hence c = b ;
end;
hence {a} is Element of Normal_forms_on A by NORMFORM:53;
end;

registration
let A be set ;
cluster non empty Element of Normal_forms_on a1;
existence
not for b1 being Element of Normal_forms_on A holds b1 is empty
proof end;
end;

definition
let A be set ;
let a be Element of DISJOINT_PAIRS A;
redefine func { as {c2} -> Element of Normal_forms_on a1;
coherence
{a} is Element of Normal_forms_on A
by ;
end;

definition
let A be set ;
let u be Element of (NormForm A);
func @ c2 -> Element of Normal_forms_on a1 equals :: HEYTING1:def 3
u;
coherence
u is Element of Normal_forms_on A
by NORMFORM:def 14;
end;

:: deftheorem Def3 defines @ HEYTING1:def 3 :
for A being set
for u being Element of (NormForm A) holds @ u = u;

Lemma41: for A being set
for u, v being Element of (NormForm A) holds mi ((@ u) ^ (@ v)) = the L_meet of (NormForm A) . u,v
by NORMFORM:def 14;

Lemma42: for A being set
for u, v being Element of (NormForm A) holds mi ((@ u) \/ (@ v)) = the L_join of (NormForm A) . u,v
by NORMFORM:def 14;

theorem Th4: :: HEYTING1:4
canceled;

theorem Th5: :: HEYTING1:5
canceled;

theorem Th6: :: HEYTING1:6
canceled;

theorem Th7: :: HEYTING1:7
for A being set
for K being Element of Normal_forms_on A holds mi (K ^ K) = K
proof end;

theorem Th8: :: HEYTING1:8
for A being set
for K being Element of Normal_forms_on A
for X being set st X c= K holds
X in Normal_forms_on A
proof end;

theorem Th9: :: HEYTING1:9
canceled;

theorem Th10: :: HEYTING1:10
for A being set
for u being Element of (NormForm A)
for X being set st X c= u holds
X is Element of (NormForm A)
proof end;

definition
let A be set ;
func Atom c1 -> Function of DISJOINT_PAIRS a1,the carrier of (NormForm a1) means :Def4: :: HEYTING1:def 4
for a being Element of DISJOINT_PAIRS A holds it . a = {a};
existence
ex b1 being Function of DISJOINT_PAIRS A,the carrier of (NormForm A) st
for a being Element of DISJOINT_PAIRS A holds b1 . a = {a}
proof end;
uniqueness
for b1, b2 being Function of DISJOINT_PAIRS A,the carrier of (NormForm A) st ( for a being Element of DISJOINT_PAIRS A holds b1 . a = {a} ) & ( for a being Element of DISJOINT_PAIRS A holds b2 . a = {a} ) holds
b1 = b2
proof end;
end;

:: deftheorem Def4 defines Atom HEYTING1:def 4 :
for A being set
for b2 being Function of DISJOINT_PAIRS A,the carrier of (NormForm A) holds
( b2 = Atom A iff for a being Element of DISJOINT_PAIRS A holds b2 . a = {a} );

theorem Th11: :: HEYTING1:11
for A being set
for c, a being Element of DISJOINT_PAIRS A st c in (Atom A) . a holds
c = a
proof end;

theorem Th12: :: HEYTING1:12
for A being set
for a being Element of DISJOINT_PAIRS A holds a in (Atom A) . a
proof end;

theorem Th13: :: HEYTING1:13
for A being set
for a being Element of DISJOINT_PAIRS A holds (Atom A) . a = (singleton (DISJOINT_PAIRS A)) . a
proof end;

theorem Th14: :: HEYTING1:14
for A being set
for K being Element of Normal_forms_on A holds FinJoin K,(Atom A) = FinUnion K,(singleton (DISJOINT_PAIRS A))
proof end;

theorem Th15: :: HEYTING1:15
for A being set
for u being Element of (NormForm A) holds u = FinJoin (@ u),(Atom A)
proof end;

Lemma78: for A being set
for u, v being Element of (NormForm A) st u [= v holds
for x being Element of [:(Fin A),(Fin A):] st x in u holds
ex b being Element of DISJOINT_PAIRS A st
( b in v & b c= x )
proof end;

Lemma79: for A being set
for u, v being Element of (NormForm A) st ( for a being Element of DISJOINT_PAIRS A st a in u holds
ex b being Element of DISJOINT_PAIRS A st
( b in v & b c= a ) ) holds
u [= v
proof end;

definition
let A be set ;
func pair_diff c1 -> BinOp of [:(Fin a1),(Fin a1):] means :Def5: :: HEYTING1:def 5
for a, b being Element of [:(Fin A),(Fin A):] holds it . a,b = a \ b;
existence
ex b1 being BinOp of [:(Fin A),(Fin A):] st
for a, b being Element of [:(Fin A),(Fin A):] holds b1 . a,b = a \ b
proof end;
correctness
uniqueness
for b1, b2 being BinOp of [:(Fin A),(Fin A):] st ( for a, b being Element of [:(Fin A),(Fin A):] holds b1 . a,b = a \ b ) & ( for a, b being Element of [:(Fin A),(Fin A):] holds b2 . a,b = a \ b ) holds
b1 = b2
;
proof end;
end;

:: deftheorem Def5 defines pair_diff HEYTING1:def 5 :
for A being set
for b2 being BinOp of [:(Fin A),(Fin A):] holds
( b2 = pair_diff A iff for a, b being Element of [:(Fin A),(Fin A):] holds b2 . a,b = a \ b );

definition
let A be set ;
let B be Element of Fin (DISJOINT_PAIRS A);
func - c2 -> Element of Fin (DISJOINT_PAIRS a1) equals :: HEYTING1:def 6
(DISJOINT_PAIRS A) /\ { [{ (g . t1) where t1 is Element of DISJOINT_PAIRS A : ( g . t1 in t1 `2 & t1 in B ) } ,{ (g . t2) where t2 is Element of DISJOINT_PAIRS A : ( g . t2 in t2 `1 & t2 in B ) } ] where g is Element of Funcs (DISJOINT_PAIRS A),[A] : for s being Element of DISJOINT_PAIRS A st s in B holds
g . s in (s `1 ) \/ (s `2 )
}
;
coherence
(DISJOINT_PAIRS A) /\ { [{ (g . t1) where t1 is Element of DISJOINT_PAIRS A : ( g . t1 in t1 `2 & t1 in B ) } ,{ (g . t2) where t2 is Element of DISJOINT_PAIRS A : ( g . t2 in t2 `1 & t2 in B ) } ] where g is Element of Funcs (DISJOINT_PAIRS A),[A] : for s being Element of DISJOINT_PAIRS A st s in B holds
g . s in (s `1 ) \/ (s `2 )
}
is Element of Fin (DISJOINT_PAIRS A)
proof end;
correctness
;
let C be Element of Fin (DISJOINT_PAIRS A);
func c2 =>> c3 -> Element of Fin (DISJOINT_PAIRS a1) equals :: HEYTING1:def 7
(DISJOINT_PAIRS A) /\ { (FinPairUnion B,((pair_diff A) .: f,(incl (DISJOINT_PAIRS A)))) where f is Element of Funcs (DISJOINT_PAIRS A),[:(Fin A),(Fin A):] : f .: B c= C } ;
coherence
(DISJOINT_PAIRS A) /\ { (FinPairUnion B,((pair_diff A) .: f,(incl (DISJOINT_PAIRS A)))) where f is Element of Funcs (DISJOINT_PAIRS A),[:(Fin A),(Fin A):] : f .: B c= C } is Element of Fin (DISJOINT_PAIRS A)
proof end;
correctness
;
end;

:: deftheorem Def6 defines - HEYTING1:def 6 :
for A being set
for B being Element of Fin (DISJOINT_PAIRS A) holds - B = (DISJOINT_PAIRS A) /\ { [{ (g . t1) where t1 is Element of DISJOINT_PAIRS A : ( g . t1 in t1 `2 & t1 in B ) } ,{ (g . t2) where t2 is Element of DISJOINT_PAIRS A : ( g . t2 in t2 `1 & t2 in B ) } ] where g is Element of Funcs (DISJOINT_PAIRS A),[A] : for s being Element of DISJOINT_PAIRS A st s in B holds
g . s in (s `1 ) \/ (s `2 )
}
;

:: deftheorem Def7 defines =>> HEYTING1:def 7 :
for A being set
for B, C being Element of Fin (DISJOINT_PAIRS A) holds B =>> C = (DISJOINT_PAIRS A) /\ { (FinPairUnion B,((pair_diff A) .: f,(incl (DISJOINT_PAIRS A)))) where f is Element of Funcs (DISJOINT_PAIRS A),[:(Fin A),(Fin A):] : f .: B c= C } ;

theorem Th16: :: HEYTING1:16
for A being set
for B being Element of Fin (DISJOINT_PAIRS A)
for c being Element of DISJOINT_PAIRS A st c in - B holds
ex g being Element of Funcs (DISJOINT_PAIRS A),[A] st
( ( for s being Element of DISJOINT_PAIRS A st s in B holds
g . s in (s `1 ) \/ (s `2 ) ) & c = [{ (g . t1) where t1 is Element of DISJOINT_PAIRS A : ( g . t1 in t1 `2 & t1 in B ) } ,{ (g . t2) where t2 is Element of DISJOINT_PAIRS A : ( g . t2 in t2 `1 & t2 in B ) } ] )
proof end;

theorem Th17: :: HEYTING1:17
for A being set holds [{} ,{} ] is Element of DISJOINT_PAIRS A
proof end;

theorem Th18: :: HEYTING1:18
for A being set
for K being Element of Normal_forms_on A st K = {} holds
- K = {[{} ,{} ]}
proof end;

theorem Th19: :: HEYTING1:19
for A being set
for K, L being Element of Normal_forms_on A st K = {} & L = {} holds
K =>> L = {[{} ,{} ]}
proof end;

theorem Th20: :: HEYTING1:20
for a being Element of DISJOINT_PAIRS {} holds a = [{} ,{} ]
proof end;

theorem Th21: :: HEYTING1:21
DISJOINT_PAIRS {} = {[{} ,{} ]}
proof end;

Lemma102: Fin (DISJOINT_PAIRS {} ) = {{} ,{[{} ,{} ]}}
proof end;

theorem Th22: :: HEYTING1:22
for A being set holds {[{} ,{} ]} is Element of Normal_forms_on A
proof end;

theorem Th23: :: HEYTING1:23
for A being set
for B, C being Element of Fin (DISJOINT_PAIRS A)
for c being Element of DISJOINT_PAIRS A st c in B =>> C holds
ex f being Element of Funcs (DISJOINT_PAIRS A),[:(Fin A),(Fin A):] st
( f .: B c= C & c = FinPairUnion B,((pair_diff A) .: f,(incl (DISJOINT_PAIRS A))) )
proof end;

theorem Th24: :: HEYTING1:24
for A being set
for a being Element of DISJOINT_PAIRS A
for K being Element of Normal_forms_on A st K ^ {a} = {} holds
ex b being Element of DISJOINT_PAIRS A st
( b in - K & b c= a )
proof end;

E109: now
let A be set ;
let K be Element of Normal_forms_on A;
let b be Element of DISJOINT_PAIRS A;
let f be Element of Funcs (DISJOINT_PAIRS A),[:(Fin A),(Fin A):];
thus ((pair_diff A) .: f,(incl (DISJOINT_PAIRS A))) . b = (pair_diff A) . (f . b),((incl (DISJOINT_PAIRS A)) . b) by FUNCOP_1:48
.= (pair_diff A) . (f . b),b by FUNCT_1:35
.= (f . b) \ b by ;
end;

theorem Th25: :: HEYTING1:25
for A being set
for a being Element of DISJOINT_PAIRS A
for u, v being Element of (NormForm A) st ( for b being Element of DISJOINT_PAIRS A st b in u holds
b \/ a in DISJOINT_PAIRS A ) & ( for c being Element of DISJOINT_PAIRS A st c in u holds
ex b being Element of DISJOINT_PAIRS A st
( b in v & b c= c \/ a ) ) holds
ex b being Element of DISJOINT_PAIRS A st
( b in (@ u) =>> (@ v) & b c= a )
proof end;

Lemma111: for A being set
for a being Element of DISJOINT_PAIRS A
for u being Element of (NormForm A)
for K being Element of Normal_forms_on A st a in K ^ (K =>> (@ u)) holds
ex b being Element of DISJOINT_PAIRS A st
( b in u & b c= a )
proof end;

theorem Th26: :: HEYTING1:26
for A being set
for K being Element of Normal_forms_on A holds K ^ (- K) = {}
proof end;

definition
let A be set ;
func pseudo_compl c1 -> UnOp of the carrier of (NormForm a1) means :Def8: :: HEYTING1:def 8
for u being Element of (NormForm A) holds it . u = mi (- (@ u));
existence
ex b1 being UnOp of the carrier of (NormForm A) st
for u being Element of (NormForm A) holds b1 . u = mi (- (@ u))
proof end;
correctness
uniqueness
for b1, b2 being UnOp of the carrier of (NormForm A) st ( for u being Element of (NormForm A) holds b1 . u = mi (- (@ u)) ) & ( for u being Element of (NormForm A) holds b2 . u = mi (- (@ u)) ) holds
b1 = b2
;
proof end;
func StrongImpl c1 -> BinOp of the carrier of (NormForm a1) means :Def9: :: HEYTING1:def 9
for u, v being Element of (NormForm A) holds it . u,v = mi ((@ u) =>> (@ v));
existence
ex b1 being BinOp of the carrier of (NormForm A) st
for u, v being Element of (NormForm A) holds b1 . u,v = mi ((@ u) =>> (@ v))
proof end;
correctness
uniqueness
for b1, b2 being BinOp of the carrier of (NormForm A) st ( for u, v being Element of (NormForm A) holds b1 . u,v = mi ((@ u) =>> (@ v)) ) & ( for u, v being Element of (NormForm A) holds b2 . u,v = mi ((@ u) =>> (@ v)) ) holds
b1 = b2
;
proof end;
let u be Element of (NormForm A);
func SUB c2 -> Element of Fin the carrier of (NormForm a1) equals :: HEYTING1:def 10
bool u;
coherence
bool u is Element of Fin the carrier of (NormForm A)
proof end;
correctness
;
func diff c2 -> UnOp of the carrier of (NormForm a1) means :Def11: :: HEYTING1:def 11
for v being Element of (NormForm A) holds it . v = u \ v;
existence
ex b1 being UnOp of the carrier of (NormForm A) st
for v being Element of (NormForm A) holds b1 . v = u \ v
proof end;
correctness
uniqueness
for b1, b2 being UnOp of the carrier of (NormForm A) st ( for v being Element of (NormForm A) holds b1 . v = u \ v ) & ( for v being Element of (NormForm A) holds b2 . v = u \ v ) holds
b1 = b2
;
proof end;
end;

:: deftheorem Def8 defines pseudo_compl HEYTING1:def 8 :
for A being set
for b2 being UnOp of the carrier of (NormForm A) holds
( b2 = pseudo_compl A iff for u being Element of (NormForm A) holds b2 . u = mi (- (@ u)) );

:: deftheorem Def9 defines StrongImpl HEYTING1:def 9 :
for A being set
for b2 being BinOp of the carrier of (NormForm A) holds
( b2 = StrongImpl A iff for u, v being Element of (NormForm A) holds b2 . u,v = mi ((@ u) =>> (@ v)) );

:: deftheorem Def10 defines SUB HEYTING1:def 10 :
for A being set
for u being Element of (NormForm A) holds SUB u = bool u;

:: deftheorem Def11 defines diff HEYTING1:def 11 :
for A being set
for u being Element of (NormForm A)
for b3 being UnOp of the carrier of (NormForm A) holds
( b3 = diff u iff for v being Element of (NormForm A) holds b3 . v = u \ v );

deffunc H1( set ) -> Relation of [:the carrier of (NormForm a1),the carrier of (NormForm a1):],the carrier of (NormForm a1) = the L_join of (NormForm a1);

deffunc H2( set ) -> Relation of [:the carrier of (NormForm a1),the carrier of (NormForm a1):],the carrier of (NormForm a1) = the L_meet of (NormForm a1);

Lemma117: for A being set
for u, v being Element of (NormForm A) st v in SUB u holds
v "\/" ((diff u) . v) = u
proof end;

theorem Th27: :: HEYTING1:27
for A being set
for u, v being Element of (NormForm A) holds (diff u) . v [= u
proof end;

Lemma119: for A being set
for a being Element of DISJOINT_PAIRS A
for u being Element of (NormForm A) ex v being Element of (NormForm A) st
( v in SUB u & (@ v) ^ {a} = {} & ( for b being Element of DISJOINT_PAIRS A st b in (diff u) . v holds
b \/ a in DISJOINT_PAIRS A ) )
proof end;

theorem Th28: :: HEYTING1:28
for A being set
for u being Element of (NormForm A) holds u "/\" ((pseudo_compl A) . u) = Bottom (NormForm A)
proof end;

theorem Th29: :: HEYTING1:29
for A being set
for u, v being Element of (NormForm A) holds u "/\" ((StrongImpl A) . u,v) [= v
proof end;

theorem Th30: :: HEYTING1:30
for A being set
for a being Element of DISJOINT_PAIRS A
for u being Element of (NormForm A) st (@ u) ^ {a} = {} holds
(Atom A) . a [= (pseudo_compl A) . u
proof end;

theorem Th31: :: HEYTING1:31
for A being set
for a being Element of DISJOINT_PAIRS A
for u, w being Element of (NormForm A) st ( for b being Element of DISJOINT_PAIRS A st b in u holds
b \/ a in DISJOINT_PAIRS A ) & u "/\" ((Atom A) . a) [= w holds
(Atom A) . a [= (StrongImpl A) . u,w
proof end;

E129: now
let A be set ;
let u be Element of (NormForm A);
let v be Element of (NormForm A);
deffunc H3( Element of (NormForm A), Element of (NormForm A)) -> Element of the carrier of (NormForm A) = FinJoin (SUB a1),(H2(A) .: (pseudo_compl A),((StrongImpl A) [:] (diff a1),a2));
set Psi = H2(A) .: (pseudo_compl A),((StrongImpl A) [:] (diff u),v);
E21: now
let w be Element of (NormForm A);
set u2 = (diff u) . w;
set pc = (pseudo_compl A) . w;
set si = (StrongImpl A) . ((diff u) . w),v;
assume w in SUB u ;
then E23: w "\/" ((diff u) . w) = u by ;
E52: w "/\" (((pseudo_compl A) . w) "/\" ((StrongImpl A) . ((diff u) . w),v)) = (w "/\" ((pseudo_compl A) . w)) "/\" ((StrongImpl A) . ((diff u) . w),v) by LATTICES:def 7
.= (Bottom (NormForm A)) "/\" ((StrongImpl A) . ((diff u) . w),v) by
.= Bottom (NormForm A) by LATTICES:40 ;
E56: ((diff u) . w) "/\" ((StrongImpl A) . ((diff u) . w),v) [= v by Th3;
(H2(A) [;] u,(H2(A) .: (pseudo_compl A),((StrongImpl A) [:] (diff u),v))) . w = H2(A) . u,((H2(A) .: (pseudo_compl A),((StrongImpl A) [:] (diff u),v)) . w) by FUNCOP_1:66
.= u "/\" ((H2(A) .: (pseudo_compl A),((StrongImpl A) [:] (diff u),v)) . w) by LATTICES:def 2
.= u "/\" (H2(A) . ((pseudo_compl A) . w),(((StrongImpl A) [:] (diff u),v) . w)) by FUNCOP_1:48
.= u "/\" (((pseudo_compl A) . w) "/\" (((StrongImpl A) [:] (diff u),v) . w)) by LATTICES:def 2
.= u "/\" (((pseudo_compl A) . w) "/\" ((StrongImpl A) . ((diff u) . w),v)) by FUNCOP_1:60
.= (w "/\" (((pseudo_compl A) . w) "/\" ((StrongImpl A) . ((diff u) . w),v))) "\/" (((diff u) . w) "/\" (((pseudo_compl A) . w) "/\" ((StrongImpl A) . ((diff u) . w),v))) by , LATTICES:def 11
.= ((diff u) . w) "/\" (((StrongImpl A) . ((diff u) . w),v) "/\" ((pseudo_compl A) . w)) by Th7, LATTICES:39
.= (((diff u) . w) "/\" ((StrongImpl A) . ((diff u) . w),v)) "/\" ((pseudo_compl A) . w) by LATTICES:def 7 ;
then (H2(A) [;] u,(H2(A) .: (pseudo_compl A),((StrongImpl A) [:] (diff u),v))) . w [= ((diff u) . w) "/\" ((StrongImpl A) . ((diff u) . w),v) by LATTICES:23;
hence (H2(A) [;] u,(H2(A) .: (pseudo_compl A),((StrongImpl A) [:] (diff u),v))) . w [= v by Th8, LATTICES:25;
end;
u "/\" H3(u,v) = FinJoin (SUB u),(H2(A) [;] u,(H2(A) .: (pseudo_compl A),((StrongImpl A) [:] (diff u),v))) by LATTICE2:83;
hence u "/\" H3(u,v) [= v by , LATTICE2:70;
let w be Element of (NormForm A);
E59: v = FinJoin (@ v),(Atom A) by ;
then E60: u "/\" v = FinJoin (@ v),(H2(A) [;] u,(Atom A)) by LATTICE2:83;
assume E61: u "/\" v [= w ;
now
let a be Element of DISJOINT_PAIRS A;
assume a in @ v ;
then (H2(A) [;] u,(Atom A)) . a [= w by Th7, Th8, LATTICE2:46;
then H2(A) . u,((Atom A) . a) [= w by FUNCOP_1:66;
then E62: u "/\" ((Atom A) . a) [= w by LATTICES:def 2;
consider v being Element of (NormForm A) such that
E64: v in SUB u and
E65: (@ v) ^ {a} = {} and
E66: for b being Element of DISJOINT_PAIRS A st b in (diff u) . v holds
b \/ a in DISJOINT_PAIRS A by ;
(diff u) . v [= u by ;
then ((diff u) . v) "/\" ((Atom A) . a) [= u "/\" ((Atom A) . a) by LATTICES:27;
then E67: ((diff u) . v) "/\" ((Atom A) . a) [= w by , LATTICES:25;
set pf = pseudo_compl A;
set sf = (StrongImpl A) [:] (diff u),w;
E68: (Atom A) . a [= (pseudo_compl A) . v by Def4, Lemma40;
(Atom A) . a [= (StrongImpl A) . ((diff u) . v),w by , , Lemma41;
then E69: (Atom A) . a [= ((StrongImpl A) [:] (diff u),w) . v by FUNCOP_1:60;
((pseudo_compl A) . v) "/\" (((StrongImpl A) [:] (diff u),w) . v) = H2(A) . ((pseudo_compl A) . v),(((StrongImpl A) [:] (diff u),w) . v) by LATTICES:def 2
.= (H2(A) .: (pseudo_compl A),((StrongImpl A) [:] (diff u),w)) . v by FUNCOP_1:48 ;
then (Atom A) . a [= (H2(A) .: (pseudo_compl A),((StrongImpl A) [:] (diff u),w)) . v by , Th11, FILTER_0:7;
hence (Atom A) . a [= H3(u,w) by Th10, LATTICE2:44;
end;
hence v [= H3(u,w) by , LATTICE2:70;
end;

Lemma137: for A being set holds NormForm A is implicative
proof end;

registration
let A be set ;
cluster NormForm a1 -> implicative ;
coherence
NormForm A is implicative
by ;
end;

theorem Th32: :: HEYTING1:32
canceled;

theorem Th33: :: HEYTING1:33
for A being set
for u, v being Element of (NormForm A) holds u => v = FinJoin (SUB u),(the L_meet of (NormForm A) .: (pseudo_compl A),((StrongImpl A) [:] (diff u),v))
proof end;

theorem Th34: :: HEYTING1:34
for A being set holds Top (NormForm A) = {[{} ,{} ]}
proof end;