begin
:: deftheorem Def1 defines Set-Constants CIRCUIT1:def 1 :
for IIG being non empty non void Circuit-like ManySortedSign
for SCS being non-empty Circuit of IIG
for b3 being ManySortedSet of SortsWithConstants IIG holds
( b3 = Set-Constants SCS iff for x being Vertex of IIG st x in dom b3 holds
b3 . x in Constants (SCS,x) );
theorem
theorem Th2:
:: deftheorem defines -th_InputValues CIRCUIT1:def 2 :
for IIG being non empty non void Circuit-like ManySortedSign st IIG is with_input_V holds
for SCS being non-empty Circuit of IIG
for InpFs being InputFuncs of SCS
for n being Element of NAT holds n -th_InputValues InpFs = (commute InpFs) . n;
theorem
canceled;
theorem
theorem
:: deftheorem defines depends_on_in CIRCUIT1:def 3 :
for IIG being non empty non void Circuit-like ManySortedSign
for SCS being non-empty Circuit of IIG
for s being State of SCS
for o being OperSymbol of IIG holds o depends_on_in s = s * (the_arity_of o);
theorem Th6:
theorem Th7:
theorem Th8:
theorem
theorem Th10:
begin
defpred S1[ set ] means verum;
:: deftheorem Def4 defines size CIRCUIT1:def 4 :
for IIG being non empty non void Circuit-like monotonic ManySortedSign
for A being non-empty finite-yielding MSAlgebra of IIG
for v being SortSymbol of IIG
for b4 being natural number holds
( b4 = size (v,A) iff ex s being non empty finite Subset of NAT st
( s = { (card t) where t is Element of the Sorts of (FreeEnv A) . v : verum } & b4 = max s ) );
theorem Th11:
theorem
theorem Th13:
theorem
:: deftheorem Def5 defines depth CIRCUIT1:def 5 :
for S being non empty non void ManySortedSign
for A being non-empty finite-yielding MSAlgebra of S
for v being SortSymbol of S
for e being Element of the Sorts of (FreeEnv A) . v
for b5 being Element of NAT holds
( b5 = depth e iff ex e9 being Element of the Sorts of (FreeMSA the Sorts of A) . v st
( e = e9 & b5 = depth e9 ) );
theorem Th15:
theorem Th16:
theorem
begin
:: deftheorem Def6 defines depth CIRCUIT1:def 6 :
for S being non empty non void monotonic ManySortedSign
for A being non-empty finite-yielding MSAlgebra of S
for v being SortSymbol of S
for b4 being natural number holds
( b4 = depth (v,A) iff ex s being non empty finite Subset of NAT st
( s = { (depth t) where t is Element of the Sorts of (FreeEnv A) . v : verum } & b4 = max s ) );
:: deftheorem Def7 defines depth CIRCUIT1:def 7 :
for IIG being non empty finite non void Circuit-like monotonic ManySortedSign
for A being non-empty Circuit of IIG
for b3 being natural number holds
( b3 = depth A iff ex Ds being non empty finite Subset of NAT st
( Ds = { (depth (v,A)) where v is Element of IIG : v in the carrier of IIG } & b3 = max Ds ) );
theorem
theorem Th19:
theorem