:: FUZZY_4 semantic presentation

registration
let c1 be non empty set ;
let c2 be Membership_Func of c1;
cluster K2(a2) -> non empty ;
coherence
not rng c2 is empty
proof end;
end;

theorem Th1: :: FUZZY_4:1
for b1 being non empty set
for b2 being Membership_Func of b1 holds
( rng b2 is bounded & ( for b3 being set holds
( b3 in dom b2 implies b2 . b3 <= sup (rng b2) ) ) & ( for b3 being set holds
( b3 in dom b2 implies b2 . b3 >= inf (rng b2) ) ) )
proof end;

theorem Th2: :: FUZZY_4:2
for b1 being non empty set
for b2, b3 being Membership_Func of b1 holds
( ( for b4 being set holds
( b4 in b1 implies b2 . b4 <= b3 . b4 ) ) implies sup (rng b2) <= sup (rng b3) )
proof end;

theorem Th3: :: FUZZY_4:3
for b1, b2 being non empty set
for b3 being RMembership_Func of b1,b2
for b4 being Element of [:b1,b2:] holds
( 0 <= b3 . b4 & b3 . b4 <= 1 )
proof end;

theorem Th4: :: FUZZY_4:4
for b1, b2 being non empty set
for b3 being RMembership_Func of b1,b2
for b4, b5 being set holds
( [b4,b5] in [:b1,b2:] implies ( 0 <= b3 . [b4,b5] & b3 . [b4,b5] <= 1 ) ) by Th3;

definition
let c1, c2 be non empty set ;
let c3 be RMembership_Func of c2,c1;
func converse c3 -> RMembership_Func of a1,a2 means :Def1: :: FUZZY_4:def 1
for b1, b2 being set holds
( [b1,b2] in [:a1,a2:] implies a4 . [b1,b2] = a3 . [b2,b1] );
existence
ex b1 being RMembership_Func of c1,c2 st
for b2, b3 being set holds
( [b2,b3] in [:c1,c2:] implies b1 . [b2,b3] = c3 . [b3,b2] )
proof end;
uniqueness
for b1, b2 being RMembership_Func of c1,c2 holds
( ( for b3, b4 being set holds
( [b3,b4] in [:c1,c2:] implies b1 . [b3,b4] = c3 . [b4,b3] ) ) & ( for b3, b4 being set holds
( [b3,b4] in [:c1,c2:] implies b2 . [b3,b4] = c3 . [b4,b3] ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def1 defines converse FUZZY_4:def 1 :
for b1, b2 being non empty set
for b3 being RMembership_Func of b2,b1
for b4 being RMembership_Func of b1,b2 holds
( b4 = converse b3 iff for b5, b6 being set holds
( [b5,b6] in [:b1,b2:] implies b4 . [b5,b6] = b3 . [b6,b5] ) );

theorem Th5: :: FUZZY_4:5
for b1, b2 being non empty set
for b3 being RMembership_Func of b1,b2 holds converse (converse b3) = b3
proof end;

theorem Th6: :: FUZZY_4:6
for b1, b2 being non empty set
for b3 being RMembership_Func of b1,b2 holds 1_minus (converse b3) = converse (1_minus b3)
proof end;

theorem Th7: :: FUZZY_4:7
for b1, b2 being non empty set
for b3, b4 being RMembership_Func of b1,b2 holds converse (max b3,b4) = max (converse b3),(converse b4)
proof end;

theorem Th8: :: FUZZY_4:8
for b1, b2 being non empty set
for b3, b4 being RMembership_Func of b1,b2 holds converse (min b3,b4) = min (converse b3),(converse b4)
proof end;

theorem Th9: :: FUZZY_4:9
for b1, b2 being non empty set
for b3, b4 being RMembership_Func of b1,b2
for b5, b6 being set holds
( b5 in b1 & b6 in b2 & b3 . [b5,b6] <= b4 . [b5,b6] implies (converse b3) . [b6,b5] <= (converse b4) . [b6,b5] )
proof end;

theorem Th10: :: FUZZY_4:10
for b1, b2 being non empty set
for b3, b4 being RMembership_Func of b1,b2 holds
( b4 c= implies converse b4 c= )
proof end;

theorem Th11: :: FUZZY_4:11
for b1, b2 being non empty set
for b3, b4 being RMembership_Func of b1,b2 holds converse (b3 \ b4) = (converse b3) \ (converse b4)
proof end;

theorem Th12: :: FUZZY_4:12
for b1, b2 being non empty set
for b3, b4 being RMembership_Func of b1,b2 holds converse (b3 \+\ b4) = (converse b3) \+\ (converse b4)
proof end;

definition
let c1, c2, c3 be non empty set ;
let c4 be RMembership_Func of c1,c2;
let c5 be RMembership_Func of c2,c3;
let c6, c7 be set ;
assume E8: ( c6 in c1 & c7 in c3 ) ;
func min c4,c5,c6,c7 -> Membership_Func of a2 means :Def2: :: FUZZY_4:def 2
for b1 being Element of a2 holds a8 . b1 = min (a4 . [a6,b1]),(a5 . [b1,a7]);
existence
ex b1 being Membership_Func of c2 st
for b2 being Element of c2 holds b1 . b2 = min (c4 . [c6,b2]),(c5 . [b2,c7])
proof end;
uniqueness
for b1, b2 being Membership_Func of c2 holds
( ( for b3 being Element of c2 holds b1 . b3 = min (c4 . [c6,b3]),(c5 . [b3,c7]) ) & ( for b3 being Element of c2 holds b2 . b3 = min (c4 . [c6,b3]),(c5 . [b3,c7]) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def2 defines min FUZZY_4:def 2 :
for b1, b2, b3 being non empty set
for b4 being RMembership_Func of b1,b2
for b5 being RMembership_Func of b2,b3
for b6, b7 being set holds
( b6 in b1 & b7 in b3 implies for b8 being Membership_Func of b2 holds
( b8 = min b4,b5,b6,b7 iff for b9 being Element of b2 holds b8 . b9 = min (b4 . [b6,b9]),(b5 . [b9,b7]) ) );

definition
let c1, c2, c3 be non empty set ;
let c4 be RMembership_Func of c1,c2;
let c5 be RMembership_Func of c2,c3;
func c4 (#) c5 -> RMembership_Func of a1,a3 means :Def3: :: FUZZY_4:def 3
for b1, b2 being set holds
( [b1,b2] in [:a1,a3:] implies a6 . [b1,b2] = sup (rng (min a4,a5,b1,b2)) );
existence
ex b1 being RMembership_Func of c1,c3 st
for b2, b3 being set holds
( [b2,b3] in [:c1,c3:] implies b1 . [b2,b3] = sup (rng (min c4,c5,b2,b3)) )
proof end;
uniqueness
for b1, b2 being RMembership_Func of c1,c3 holds
( ( for b3, b4 being set holds
( [b3,b4] in [:c1,c3:] implies b1 . [b3,b4] = sup (rng (min c4,c5,b3,b4)) ) ) & ( for b3, b4 being set holds
( [b3,b4] in [:c1,c3:] implies b2 . [b3,b4] = sup (rng (min c4,c5,b3,b4)) ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def3 defines (#) FUZZY_4:def 3 :
for b1, b2, b3 being non empty set
for b4 being RMembership_Func of b1,b2
for b5 being RMembership_Func of b2,b3
for b6 being RMembership_Func of b1,b3 holds
( b6 = b4 (#) b5 iff for b7, b8 being set holds
( [b7,b8] in [:b1,b3:] implies b6 . [b7,b8] = sup (rng (min b4,b5,b7,b8)) ) );

Lemma10: for b1, b2, b3 being non empty set
for b4 being RMembership_Func of b1,b2
for b5, b6 being RMembership_Func of b2,b3
for b7, b8 being set holds
( b7 in b1 & b8 in b3 implies sup (rng (min b4,(max b5,b6),b7,b8)) = max (sup (rng (min b4,b5,b7,b8))),(sup (rng (min b4,b6,b7,b8))) )
proof end;

theorem Th13: :: FUZZY_4:13
for b1, b2, b3 being non empty set
for b4 being RMembership_Func of b1,b2
for b5, b6 being RMembership_Func of b2,b3 holds b4 (#) (max b5,b6) = max (b4 (#) b5),(b4 (#) b6)
proof end;

Lemma11: for b1, b2, b3 being non empty set
for b4, b5 being RMembership_Func of b1,b2
for b6 being RMembership_Func of b2,b3
for b7, b8 being set holds
( b7 in b1 & b8 in b3 implies sup (rng (min (max b4,b5),b6,b7,b8)) = max (sup (rng (min b4,b6,b7,b8))),(sup (rng (min b5,b6,b7,b8))) )
proof end;

theorem Th14: :: FUZZY_4:14
for b1, b2, b3 being non empty set
for b4, b5 being RMembership_Func of b1,b2
for b6 being RMembership_Func of b2,b3 holds (max b4,b5) (#) b6 = max (b4 (#) b6),(b5 (#) b6)
proof end;

Lemma12: for b1, b2, b3 being non empty set
for b4 being RMembership_Func of b1,b2
for b5, b6 being RMembership_Func of b2,b3
for b7, b8 being set holds
( b7 in b1 & b8 in b3 implies sup (rng (min b4,(min b5,b6),b7,b8)) <= min (sup (rng (min b4,b5,b7,b8))),(sup (rng (min b4,b6,b7,b8))) )
proof end;

theorem Th15: :: FUZZY_4:15
for b1, b2, b3 being non empty set
for b4 being RMembership_Func of b1,b2
for b5, b6 being RMembership_Func of b2,b3 holds min (b4 (#) b5),(b4 (#) b6) c=
proof end;

Lemma13: for b1, b2, b3 being non empty set
for b4, b5 being RMembership_Func of b1,b2
for b6 being RMembership_Func of b2,b3
for b7, b8 being set holds
( b7 in b1 & b8 in b3 implies sup (rng (min (min b4,b5),b6,b7,b8)) <= min (sup (rng (min b4,b6,b7,b8))),(sup (rng (min b5,b6,b7,b8))) )
proof end;

theorem Th16: :: FUZZY_4:16
for b1, b2, b3 being non empty set
for b4, b5 being RMembership_Func of b1,b2
for b6 being RMembership_Func of b2,b3 holds min (b4 (#) b6),(b5 (#) b6) c=
proof end;

Lemma14: for b1, b2, b3 being non empty set
for b4 being RMembership_Func of b1,b2
for b5 being RMembership_Func of b2,b3
for b6, b7 being set holds
( b6 in b1 & b7 in b3 implies sup (rng (min (converse b5),(converse b4),b7,b6)) = sup (rng (min b4,b5,b6,b7)) )
proof end;

theorem Th17: :: FUZZY_4:17
for b1, b2, b3 being non empty set
for b4 being RMembership_Func of b1,b2
for b5 being RMembership_Func of b2,b3 holds converse (b4 (#) b5) = (converse b5) (#) (converse b4)
proof end;

theorem Th18: :: FUZZY_4:18
for b1, b2, b3 being non empty set
for b4, b5 being RMembership_Func of b1,b2
for b6, b7 being RMembership_Func of b2,b3
for b8, b9 being set holds
( b8 in b1 & b9 in b3 & ( for b10 being set holds
( b10 in b2 implies ( b4 . [b8,b10] <= b5 . [b8,b10] & b6 . [b10,b9] <= b7 . [b10,b9] ) ) ) implies (b4 (#) b6) . [b8,b9] <= (b5 (#) b7) . [b8,b9] )
proof end;

theorem Th19: :: FUZZY_4:19
for b1, b2, b3 being non empty set
for b4, b5 being RMembership_Func of b1,b2
for b6, b7 being RMembership_Func of b2,b3 holds
( b5 c= & b7 c= implies b5 (#) b7 c= )
proof end;

definition
let c1, c2 be non empty set ;
func Imf c1,c2 -> RMembership_Func of a1,a2 means :: FUZZY_4:def 4
for b1, b2 being set holds
( [b1,b2] in [:a1,a2:] implies ( ( b1 = b2 implies a3 . [b1,b2] = 1 ) & ( not b1 = b2 implies a3 . [b1,b2] = 0 ) ) );
existence
ex b1 being RMembership_Func of c1,c2 st
for b2, b3 being set holds
( [b2,b3] in [:c1,c2:] implies ( ( b2 = b3 implies b1 . [b2,b3] = 1 ) & ( not b2 = b3 implies b1 . [b2,b3] = 0 ) ) )
proof end;
uniqueness
for b1, b2 being RMembership_Func of c1,c2 holds
( ( for b3, b4 being set holds
( [b3,b4] in [:c1,c2:] implies ( ( b3 = b4 implies b1 . [b3,b4] = 1 ) & ( not b3 = b4 implies b1 . [b3,b4] = 0 ) ) ) ) & ( for b3, b4 being set holds
( [b3,b4] in [:c1,c2:] implies ( ( b3 = b4 implies b2 . [b3,b4] = 1 ) & ( not b3 = b4 implies b2 . [b3,b4] = 0 ) ) ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def4 defines Imf FUZZY_4:def 4 :
for b1, b2 being non empty set
for b3 being RMembership_Func of b1,b2 holds
( b3 = Imf b1,b2 iff for b4, b5 being set holds
( [b4,b5] in [:b1,b2:] implies ( ( b4 = b5 implies b3 . [b4,b5] = 1 ) & ( not b4 = b5 implies b3 . [b4,b5] = 0 ) ) ) );

theorem Th20: :: FUZZY_4:20
for b1, b2 being non empty set
for b3 being Element of [:b1,b2:] holds
( (Zmf b1,b2) . b3 = 0 & (Umf b1,b2) . b3 = 1 )
proof end;

theorem Th21: :: FUZZY_4:21
for b1, b2 being non empty set
for b3, b4 being set holds
( [b3,b4] in [:b1,b2:] implies ( (Zmf b1,b2) . [b3,b4] = 0 & (Umf b1,b2) . [b3,b4] = 1 ) ) by Th20;

Lemma17: for b1, b2, b3 being non empty set
for b4 being RMembership_Func of b1,b2
for b5, b6 being set holds
( b5 in b3 & b6 in b2 implies sup (rng (min (Zmf b3,b1),b4,b5,b6)) = (Zmf b3,b2) . [b5,b6] )
proof end;

theorem Th22: :: FUZZY_4:22
for b1, b2, b3 being non empty set
for b4 being RMembership_Func of b1,b2 holds (Zmf b3,b1) (#) b4 = Zmf b3,b2
proof end;

Lemma19: for b1, b2, b3 being non empty set
for b4 being RMembership_Func of b1,b2
for b5, b6 being set holds
( b5 in b1 & b6 in b3 implies sup (rng (min b4,(Zmf b2,b3),b5,b6)) = (Zmf b1,b3) . [b5,b6] )
proof end;

theorem Th23: :: FUZZY_4:23
for b1, b2, b3 being non empty set
for b4 being RMembership_Func of b1,b2 holds b4 (#) (Zmf b2,b3) = Zmf b1,b3
proof end;

theorem Th24: :: FUZZY_4:24
for b1 being non empty set
for b2 being RMembership_Func of b1,b1 holds b2 (#) (Zmf b1,b1) = (Zmf b1,b1) (#) b2
proof end;