begin
Lm1:
for x being set holds not x in x
;
theorem Th1:
definition
let IIG be non
empty non
void Circuit-like monotonic ManySortedSign ;
let A be
non-empty Circuit of
IIG;
let iv be
InputValues of
A;
func Fix_inp iv -> ManySortedFunction of
FreeGen the
Sorts of
A, the
Sorts of
(FreeEnv A) means :
Def1:
for
v being
Vertex of
IIG holds
( (
v in InputVertices IIG implies
it . v = (FreeGen (v, the Sorts of A)) --> (root-tree [(iv . v),v]) ) & (
v in SortsWithConstants IIG implies
it . v = (FreeGen (v, the Sorts of A)) --> (root-tree [(action_at v), the carrier of IIG]) ) & (
v in (InnerVertices IIG) \ (SortsWithConstants IIG) implies
it . v = id (FreeGen (v, the Sorts of A)) ) );
existence
ex b1 being ManySortedFunction of FreeGen the Sorts of A, the Sorts of (FreeEnv A) st
for v being Vertex of IIG holds
( ( v in InputVertices IIG implies b1 . v = (FreeGen (v, the Sorts of A)) --> (root-tree [(iv . v),v]) ) & ( v in SortsWithConstants IIG implies b1 . v = (FreeGen (v, the Sorts of A)) --> (root-tree [(action_at v), the carrier of IIG]) ) & ( v in (InnerVertices IIG) \ (SortsWithConstants IIG) implies b1 . v = id (FreeGen (v, the Sorts of A)) ) )
uniqueness
for b1, b2 being ManySortedFunction of FreeGen the Sorts of A, the Sorts of (FreeEnv A) st ( for v being Vertex of IIG holds
( ( v in InputVertices IIG implies b1 . v = (FreeGen (v, the Sorts of A)) --> (root-tree [(iv . v),v]) ) & ( v in SortsWithConstants IIG implies b1 . v = (FreeGen (v, the Sorts of A)) --> (root-tree [(action_at v), the carrier of IIG]) ) & ( v in (InnerVertices IIG) \ (SortsWithConstants IIG) implies b1 . v = id (FreeGen (v, the Sorts of A)) ) ) ) & ( for v being Vertex of IIG holds
( ( v in InputVertices IIG implies b2 . v = (FreeGen (v, the Sorts of A)) --> (root-tree [(iv . v),v]) ) & ( v in SortsWithConstants IIG implies b2 . v = (FreeGen (v, the Sorts of A)) --> (root-tree [(action_at v), the carrier of IIG]) ) & ( v in (InnerVertices IIG) \ (SortsWithConstants IIG) implies b2 . v = id (FreeGen (v, the Sorts of A)) ) ) ) holds
b1 = b2
end;
:: deftheorem Def1 defines Fix_inp CIRCUIT2:def 1 :
for IIG being non empty non void Circuit-like monotonic ManySortedSign
for A being non-empty Circuit of IIG
for iv being InputValues of A
for b4 being ManySortedFunction of FreeGen the Sorts of A, the Sorts of (FreeEnv A) holds
( b4 = Fix_inp iv iff for v being Vertex of IIG holds
( ( v in InputVertices IIG implies b4 . v = (FreeGen (v, the Sorts of A)) --> (root-tree [(iv . v),v]) ) & ( v in SortsWithConstants IIG implies b4 . v = (FreeGen (v, the Sorts of A)) --> (root-tree [(action_at v), the carrier of IIG]) ) & ( v in (InnerVertices IIG) \ (SortsWithConstants IIG) implies b4 . v = id (FreeGen (v, the Sorts of A)) ) ) );
:: deftheorem Def2 defines Fix_inp_ext CIRCUIT2:def 2 :
for IIG being non empty non void Circuit-like monotonic ManySortedSign
for A being non-empty Circuit of IIG
for iv being InputValues of A
for b4 being ManySortedFunction of (FreeEnv A),(FreeEnv A) holds
( b4 = Fix_inp_ext iv iff ( b4 is_homomorphism FreeEnv A, FreeEnv A & Fix_inp iv c= b4 ) );
theorem Th2:
theorem Th3:
theorem Th4:
theorem Th5:
theorem Th6:
theorem Th7:
:: deftheorem Def3 defines IGTree CIRCUIT2:def 3 :
for IIG being non empty non void Circuit-like monotonic ManySortedSign
for SCS being non-empty Circuit of IIG
for v being Vertex of IIG
for iv being InputValues of SCS
for b5 being Element of the Sorts of (FreeEnv SCS) . v holds
( b5 = IGTree (v,iv) iff ex e being Element of the Sorts of (FreeEnv SCS) . v st
( card e = size (v,SCS) & b5 = ((Fix_inp_ext iv) . v) . e ) );
theorem Th8:
theorem Th9:
:: deftheorem defines IGValue CIRCUIT2:def 4 :
for IIG being non empty non void Circuit-like monotonic ManySortedSign
for SCS being non-empty Circuit of IIG
for v being Vertex of IIG
for iv being InputValues of SCS holds IGValue (v,iv) = ((Eval SCS) . v) . (IGTree (v,iv));
theorem Th10:
theorem Th11:
begin
:: deftheorem Def5 defines Following CIRCUIT2:def 5 :
for IIG being non empty non void Circuit-like ManySortedSign
for SCS being non-empty Circuit of IIG
for s, b4 being State of SCS holds
( b4 = Following s iff for v being Vertex of IIG holds
( ( v in InputVertices IIG implies b4 . v = s . v ) & ( v in InnerVertices IIG implies b4 . v = (Den ((action_at v),SCS)) . ((action_at v) depends_on_in s) ) ) );
theorem Th12:
:: deftheorem Def6 defines stable CIRCUIT2:def 6 :
for IIG being non empty non void Circuit-like ManySortedSign
for SCS being non-empty Circuit of IIG
for IT being State of SCS holds
( IT is stable iff IT = Following IT );
:: deftheorem defines Following CIRCUIT2:def 7 :
for IIG being non empty non void Circuit-like monotonic ManySortedSign
for SCS being non-empty Circuit of IIG
for s being State of SCS
for iv being InputValues of SCS holds Following (s,iv) = Following (s +* iv);
:: deftheorem defines InitialComp CIRCUIT2:def 8 :
for IIG being non empty non void Circuit-like monotonic ManySortedSign
for SCS being non-empty Circuit of IIG
for InpFs being InputFuncs of SCS
for s being State of SCS holds InitialComp (s,InpFs) = (s +* (0 -th_InputValues InpFs)) +* (Set-Constants SCS);
definition
let IIG be non
empty non
void Circuit-like monotonic ManySortedSign ;
let SCS be
non-empty Circuit of
IIG;
let InpFs be
InputFuncs of
SCS;
let s be
State of
SCS;
func Computation (
s,
InpFs)
-> Function of
NAT,
(product the Sorts of SCS) means :
Def9:
(
it . 0 = InitialComp (
s,
InpFs) & ( for
i being
Nat holds
it . (i + 1) = Following (
(it . i),
((i + 1) -th_InputValues InpFs)) ) );
correctness
existence
ex b1 being Function of NAT,(product the Sorts of SCS) st
( b1 . 0 = InitialComp (s,InpFs) & ( for i being Nat holds b1 . (i + 1) = Following ((b1 . i),((i + 1) -th_InputValues InpFs)) ) );
uniqueness
for b1, b2 being Function of NAT,(product the Sorts of SCS) st b1 . 0 = InitialComp (s,InpFs) & ( for i being Nat holds b1 . (i + 1) = Following ((b1 . i),((i + 1) -th_InputValues InpFs)) ) & b2 . 0 = InitialComp (s,InpFs) & ( for i being Nat holds b2 . (i + 1) = Following ((b2 . i),((i + 1) -th_InputValues InpFs)) ) holds
b1 = b2;
end;
:: deftheorem Def9 defines Computation CIRCUIT2:def 9 :
for IIG being non empty non void Circuit-like monotonic ManySortedSign
for SCS being non-empty Circuit of IIG
for InpFs being InputFuncs of SCS
for s being State of SCS
for b5 being Function of NAT,(product the Sorts of SCS) holds
( b5 = Computation (s,InpFs) iff ( b5 . 0 = InitialComp (s,InpFs) & ( for i being Nat holds b5 . (i + 1) = Following ((b5 . i),((i + 1) -th_InputValues InpFs)) ) ) );
theorem Th13:
theorem Th14:
theorem
theorem Th16:
theorem Th17:
theorem
theorem